wip make qtys into polynomials
This commit is contained in:
parent
1d8a6bb325
commit
4c008577b4
22 changed files with 1650 additions and 1254 deletions
|
@ -5,6 +5,7 @@ import Quox.Syntax.Term.Subst
|
|||
import Quox.Context
|
||||
import Quox.Pretty
|
||||
|
||||
import Quox.SingletonExtra
|
||||
import Data.Vect
|
||||
import Derive.Prelude
|
||||
|
||||
|
@ -13,21 +14,23 @@ import Derive.Prelude
|
|||
|
||||
|
||||
export
|
||||
prettyUniverse : {opts : _} -> Universe -> Eff Pretty (Doc opts)
|
||||
prettyUniverse : {opts : LayoutOpts} -> Universe -> Eff Pretty (Doc opts)
|
||||
prettyUniverse = hl Universe . text . show
|
||||
|
||||
|
||||
export
|
||||
prettyTerm : {opts : _} -> BContext d -> BContext n -> Term d n ->
|
||||
Eff Pretty (Doc opts)
|
||||
prettyTerm : {opts : LayoutOpts} ->
|
||||
NameContexts q d n -> Term q d n -> Eff Pretty (Doc opts)
|
||||
|
||||
export
|
||||
prettyElim : {opts : _} -> BContext d -> BContext n -> Elim d n ->
|
||||
Eff Pretty (Doc opts)
|
||||
prettyElim : {opts : LayoutOpts} ->
|
||||
NameContexts q d n -> Elim q d n -> Eff Pretty (Doc opts)
|
||||
|
||||
private
|
||||
BTelescope : Nat -> Nat -> Type
|
||||
BTelescope = Telescope' BindName
|
||||
|
||||
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
|
||||
|
||||
|
||||
private
|
||||
|
@ -40,109 +43,111 @@ superscript = pack . map sup . unpack where
|
|||
|
||||
|
||||
private
|
||||
PiBind : Nat -> Nat -> Type
|
||||
PiBind d n = (Qty, BindName, Term d n)
|
||||
PiBind : TermLike
|
||||
PiBind q d n = (Qty q, BindName, Term q d n)
|
||||
|
||||
private
|
||||
pbname : PiBind d n -> BindName
|
||||
pbname : PiBind q d n -> BindName
|
||||
pbname (_, x, _) = x
|
||||
|
||||
private
|
||||
record SplitPi d n where
|
||||
record SplitPi q d n where
|
||||
constructor MkSplitPi
|
||||
{0 inner : Nat}
|
||||
binds : Telescope (PiBind d) n inner
|
||||
cod : Term d inner
|
||||
binds : Telescope (PiBind q d) n inner
|
||||
cod : Term q d inner
|
||||
|
||||
private
|
||||
splitPi : Telescope (PiBind d) n n' -> Term d n' -> SplitPi d n
|
||||
splitPi binds (Pi qty arg res _) =
|
||||
splitPi : {q : Nat} ->
|
||||
Telescope (PiBind q d) n n' -> Term q d n' -> SplitPi q d n
|
||||
splitPi binds cod@(Pi qty arg res _) =
|
||||
splitPi (binds :< (qty, res.name, arg)) $
|
||||
assert_smaller res $ pushSubsts' res.term
|
||||
assert_smaller cod $ pushSubsts' res.term
|
||||
splitPi binds cod = MkSplitPi {binds, cod}
|
||||
|
||||
private
|
||||
prettyPiBind1 : {opts : _} ->
|
||||
Qty -> BindName -> BContext d -> BContext n -> Term d n ->
|
||||
prettyPiBind1 : {opts : LayoutOpts} ->
|
||||
NameContexts q d n -> Qty q -> BindName -> Term q d n ->
|
||||
Eff Pretty (Doc opts)
|
||||
prettyPiBind1 pi (BN Unused _) dnames tnames s =
|
||||
prettyPiBind1 names pi (BN Unused _) s =
|
||||
hcat <$> sequence
|
||||
[prettyQty pi, dotD,
|
||||
withPrec Arg $ assert_total prettyTerm dnames tnames s]
|
||||
prettyPiBind1 pi x dnames tnames s = hcat <$> sequence
|
||||
[prettyQty pi, dotD,
|
||||
hl Delim $ text "(",
|
||||
hsep <$> sequence
|
||||
[prettyTBind x, hl Delim $ text ":",
|
||||
withPrec Outer $ assert_total prettyTerm dnames tnames s],
|
||||
hl Delim $ text ")"]
|
||||
[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 ")"]
|
||||
|
||||
private
|
||||
prettyPiBinds : {opts : _} ->
|
||||
BContext d -> BContext n ->
|
||||
Telescope (PiBind d) n n' ->
|
||||
prettyPiBinds : {opts : LayoutOpts} ->
|
||||
NameContexts q d n ->
|
||||
Telescope (PiBind q d) n n' ->
|
||||
Eff Pretty (SnocList (Doc opts))
|
||||
prettyPiBinds _ _ [<] = pure [<]
|
||||
prettyPiBinds dnames tnames (binds :< (q, x, t)) =
|
||||
let tnames' = tnames . map pbname binds in
|
||||
[|prettyPiBinds dnames tnames binds :<
|
||||
prettyPiBind1 q x dnames tnames' t|]
|
||||
prettyPiBinds _ [<] = pure [<]
|
||||
prettyPiBinds names (binds :< (q, x, t)) =
|
||||
let names' = extendTermN' (map pbname binds) names in
|
||||
[|prettyPiBinds names binds :<
|
||||
prettyPiBind1 names' q x t|]
|
||||
|
||||
|
||||
private
|
||||
SigBind : Nat -> Nat -> Type
|
||||
SigBind d n = (BindName, Term d n)
|
||||
SigBind : TermLike
|
||||
SigBind q d n = (BindName, Term q d n)
|
||||
|
||||
private
|
||||
record SplitSig d n where
|
||||
record SplitSig q d n where
|
||||
constructor MkSplitSig
|
||||
{0 inner : Nat}
|
||||
binds : Telescope (SigBind d) n inner
|
||||
last : Term d inner
|
||||
binds : Telescope (SigBind q d) n inner
|
||||
last : Term q d inner
|
||||
|
||||
private
|
||||
splitSig : Telescope (SigBind d) n n' -> Term d n' -> SplitSig d n
|
||||
splitSig : {q : Nat} ->
|
||||
Telescope (SigBind q d) n n' -> Term q d n' -> SplitSig q d n
|
||||
splitSig binds (Sig fst snd _) =
|
||||
splitSig (binds :< (snd.name, fst)) $
|
||||
assert_smaller snd $ pushSubsts' snd.term
|
||||
splitSig binds last = MkSplitSig {binds, last}
|
||||
|
||||
private
|
||||
prettySigBind1 : {opts : _} ->
|
||||
BindName -> BContext d -> BContext n -> Term d n ->
|
||||
prettySigBind1 : {opts : LayoutOpts} ->
|
||||
NameContexts q d n -> BindName -> Term q d n ->
|
||||
Eff Pretty (Doc opts)
|
||||
prettySigBind1 (BN Unused _) dnames tnames s =
|
||||
withPrec InTimes $ assert_total prettyTerm dnames tnames s
|
||||
prettySigBind1 x dnames tnames s = hcat <$> sequence
|
||||
prettySigBind1 names (BN Unused _) s =
|
||||
withPrec InTimes $ assert_total prettyTerm names s
|
||||
prettySigBind1 names x s = hcat <$> sequence
|
||||
[hl Delim $ text "(",
|
||||
hsep <$> sequence
|
||||
[prettyTBind x, hl Delim $ text ":",
|
||||
withPrec Outer $ assert_total prettyTerm dnames tnames s],
|
||||
withPrec Outer $ assert_total prettyTerm names s],
|
||||
hl Delim $ text ")"]
|
||||
|
||||
private
|
||||
prettySigBinds : {opts : _} ->
|
||||
BContext d -> BContext n ->
|
||||
Telescope (SigBind d) n n' ->
|
||||
prettySigBinds : {opts : LayoutOpts} ->
|
||||
NameContexts q d n -> Telescope (SigBind q d) n n' ->
|
||||
Eff Pretty (SnocList (Doc opts))
|
||||
prettySigBinds _ _ [<] = pure [<]
|
||||
prettySigBinds dnames tnames (binds :< (x, t)) =
|
||||
let tnames' = tnames . map fst binds in
|
||||
[|prettySigBinds dnames tnames binds :<
|
||||
prettySigBind1 x dnames tnames' t|]
|
||||
prettySigBinds _ [<] = pure [<]
|
||||
prettySigBinds names (binds :< (x, t)) =
|
||||
let names' = extendTermN' (map fst binds) names in
|
||||
[|prettySigBinds names binds :<
|
||||
prettySigBind1 names' x t|]
|
||||
|
||||
|
||||
private
|
||||
prettyTypeLine : {opts : _} ->
|
||||
BContext d -> BContext n ->
|
||||
DScopeTerm d n ->
|
||||
prettyTypeLine : {opts : LayoutOpts} ->
|
||||
NameContexts q d n -> DScopeTerm q d n ->
|
||||
Eff Pretty (Doc opts)
|
||||
prettyTypeLine dnames tnames (S _ (N body)) =
|
||||
withPrec Arg (prettyTerm dnames tnames body)
|
||||
prettyTypeLine dnames tnames (S [< i] (Y body)) =
|
||||
prettyTypeLine names (S _ (N body)) =
|
||||
withPrec Arg (prettyTerm names body)
|
||||
prettyTypeLine names (S [< i] (Y body)) =
|
||||
parens =<< do
|
||||
let names' = extendDim i names
|
||||
i' <- prettyDBind i
|
||||
ty' <- withPrec Outer $ prettyTerm (dnames :< i) tnames body
|
||||
ty' <- withPrec Outer $ prettyTerm names' body
|
||||
pure $ sep [hsep [i', !darrowD], ty']
|
||||
|
||||
|
||||
|
@ -155,16 +160,16 @@ NameChunks : Type
|
|||
NameChunks = SnocList (NameSort, SnocList BindName)
|
||||
|
||||
private
|
||||
record SplitLams d n where
|
||||
record SplitLams q d n where
|
||||
constructor MkSplitLams
|
||||
{0 dinner, ninner : Nat}
|
||||
dnames : BTelescope d dinner
|
||||
tnames : BTelescope n ninner
|
||||
chunks : NameChunks
|
||||
body : Term dinner ninner
|
||||
body : Term q dinner ninner
|
||||
|
||||
private
|
||||
splitLams : Term d n -> SplitLams d n
|
||||
splitLams : {q : Nat} -> Term q d n -> SplitLams q d n
|
||||
splitLams s = go [<] [<] [<] (pushSubsts' s)
|
||||
where
|
||||
push : NameSort -> BindName -> NameChunks -> NameChunks
|
||||
|
@ -175,7 +180,7 @@ where
|
|||
|
||||
go : BTelescope d d' -> BTelescope n n' ->
|
||||
SnocList (NameSort, SnocList BindName) ->
|
||||
Term d' n' -> SplitLams d n
|
||||
Term q d' n' -> SplitLams q d n
|
||||
go dnames tnames chunks (Lam b _) =
|
||||
go dnames (tnames :< b.name) (push T b.name chunks) $
|
||||
assert_smaller b $ pushSubsts' b.term
|
||||
|
@ -187,28 +192,31 @@ where
|
|||
|
||||
|
||||
private
|
||||
splitTuple : SnocList (Term d n) -> Term d n -> SnocList (Term d n)
|
||||
splitTuple : {q : Nat} ->
|
||||
SnocList (Term q d n) -> Term q d n -> SnocList (Term q d n)
|
||||
splitTuple ss p@(Pair t1 t2 _) =
|
||||
splitTuple (ss :< t1) $ assert_smaller p $ pushSubsts' t2
|
||||
splitTuple ss t = ss :< t
|
||||
|
||||
|
||||
private
|
||||
prettyTArg : {opts : _} -> BContext d -> BContext n ->
|
||||
Term d n -> Eff Pretty (Doc opts)
|
||||
prettyTArg dnames tnames s =
|
||||
withPrec Arg $ assert_total prettyTerm dnames tnames s
|
||||
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
|
||||
|
||||
private
|
||||
prettyDArg : {opts : _} -> BContext d -> Dim d -> Eff Pretty (Doc opts)
|
||||
prettyDArg dnames p = [|atD <+> withPrec Arg (prettyDim dnames p)|]
|
||||
prettyDArg : {opts : LayoutOpts} ->
|
||||
NameContexts _ d _ -> Dim d -> Eff Pretty (Doc opts)
|
||||
prettyDArg names p = [|atD <+> withPrec Arg (prettyDim names.dnames p)|]
|
||||
|
||||
private
|
||||
splitApps : Elim d n -> (Elim d n, List (Either (Dim d) (Term d n)))
|
||||
splitApps : {q : Nat} ->
|
||||
Elim q d n -> (Elim q d n, List (Either (Dim d) (Term q d n)))
|
||||
splitApps e = go [] (pushSubsts' e)
|
||||
where
|
||||
go : List (Either (Dim d) (Term d n)) -> Elim d n ->
|
||||
(Elim d n, List (Either (Dim d) (Term d n)))
|
||||
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)))
|
||||
go xs e@(App f s _) =
|
||||
go (Right s :: xs) $ assert_smaller e $ pushSubsts' f
|
||||
go xs e@(DApp f p _) =
|
||||
|
@ -216,49 +224,53 @@ where
|
|||
go xs s = (s, xs)
|
||||
|
||||
private
|
||||
prettyDTApps : {opts : _} ->
|
||||
BContext d -> BContext n ->
|
||||
Elim d n -> List (Either (Dim d) (Term d n)) ->
|
||||
prettyDTApps : {opts : LayoutOpts} ->
|
||||
NameContexts q d n ->
|
||||
Elim q d n -> List (Either (Dim d) (Term q d n)) ->
|
||||
Eff Pretty (Doc opts)
|
||||
prettyDTApps dnames tnames f xs = do
|
||||
f <- withPrec Arg $ assert_total prettyElim dnames tnames f
|
||||
xs <- for xs $ either (prettyDArg dnames) (prettyTArg dnames tnames)
|
||||
parensIfM App =<< prettyAppD f xs
|
||||
prettyDTApps names f xs = do
|
||||
f <- withPrec Arg $ assert_total prettyElim names f
|
||||
xs <- for xs $ either (prettyDArg names) (prettyTArg names)
|
||||
prettyAppD f xs
|
||||
|
||||
|
||||
private
|
||||
record CaseArm opts d n where
|
||||
record CaseArm opts q d n where
|
||||
constructor MkCaseArm
|
||||
pat : Doc opts
|
||||
dbinds : BTelescope d dinner -- 🍴
|
||||
tbinds : BTelescope n ninner
|
||||
body : Term dinner ninner
|
||||
|
||||
parameters {opts : LayoutOpts} (dnames : BContext d) (tnames : BContext n)
|
||||
private
|
||||
prettyCaseArm : CaseArm opts d n -> Eff Pretty (Doc opts)
|
||||
prettyCaseArm (MkCaseArm pat dbinds tbinds body) = do
|
||||
body <- withPrec Outer $ assert_total
|
||||
prettyTerm (dnames . dbinds) (tnames . tbinds) body
|
||||
header <- (pat <++>) <$> darrowD
|
||||
pure $ ifMultiline (header <++> body) (vsep [header, !(indentD body)])
|
||||
|
||||
private
|
||||
prettyCaseBody : List (CaseArm opts d n) -> Eff Pretty (List (Doc opts))
|
||||
prettyCaseBody xs = traverse prettyCaseArm xs
|
||||
body : Term q dinner ninner
|
||||
|
||||
private
|
||||
prettyCompPat : {opts : _} -> DimConst -> BindName -> Eff Pretty (Doc opts)
|
||||
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)])
|
||||
|
||||
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
|
||||
|
||||
private
|
||||
prettyCompPat : {opts : LayoutOpts} ->
|
||||
DimConst -> BindName -> Eff Pretty (Doc opts)
|
||||
prettyCompPat e x = [|prettyDimConst e <++> prettyDBind x|]
|
||||
|
||||
private
|
||||
prettyCompArm : {opts : _} -> BContext d -> BContext n ->
|
||||
DimConst -> DScopeTerm d n -> Eff Pretty (Doc opts)
|
||||
prettyCompArm dnames tnames e s = prettyCaseArm dnames tnames $
|
||||
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
|
||||
layoutComp : {opts : _} ->
|
||||
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
|
||||
|
@ -273,32 +285,33 @@ layoutComp typq val r arms = do
|
|||
|
||||
|
||||
export
|
||||
prettyEnum : {opts : _} -> List String -> Eff Pretty (Doc opts)
|
||||
prettyEnum : {opts : LayoutOpts} -> List String -> Eff Pretty (Doc opts)
|
||||
prettyEnum cases =
|
||||
tightBraces =<<
|
||||
fillSeparateTight !commaD <$>
|
||||
traverse (hl Constant . Doc.text . quoteTag) cases
|
||||
|
||||
private
|
||||
prettyCaseRet : {opts : _} ->
|
||||
BContext d -> BContext n ->
|
||||
ScopeTerm d n -> Eff Pretty (Doc opts)
|
||||
prettyCaseRet dnames tnames body = withPrec Outer $ case body of
|
||||
S _ (N tm) => assert_total prettyTerm dnames tnames tm
|
||||
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
|
||||
S [< x] (Y tm) => do
|
||||
let names' = extendTerm x names
|
||||
header <- [|prettyTBind x <++> darrowD|]
|
||||
body <- assert_total prettyTerm dnames (tnames :< x) tm
|
||||
body <- assert_total prettyTerm names' tm
|
||||
hangDSingle header body
|
||||
|
||||
private
|
||||
prettyCase_ : {opts : _} ->
|
||||
BContext d -> BContext n ->
|
||||
Doc opts -> Elim d n -> ScopeTerm d n -> List (CaseArm opts d n) ->
|
||||
prettyCase_ : {opts : LayoutOpts} ->
|
||||
NameContexts q d n ->
|
||||
Doc opts -> Elim q d n -> ScopeTerm q d n ->
|
||||
List (CaseArm opts q d n) ->
|
||||
Eff Pretty (Doc opts)
|
||||
prettyCase_ dnames tnames intro head ret body = do
|
||||
head <- withPrec Outer $ assert_total prettyElim dnames tnames head
|
||||
ret <- prettyCaseRet dnames tnames ret
|
||||
bodys <- prettyCaseBody dnames tnames body
|
||||
prettyCase_ names intro head ret body = do
|
||||
head <- withPrec Outer $ assert_total prettyElim names head
|
||||
ret <- prettyCaseRet names ret
|
||||
bodys <- prettyCaseBody names body
|
||||
return <- returnD; of_ <- ofD
|
||||
lb <- hl Delim "{"; rb <- hl Delim "}"; semi <- semiD
|
||||
ind <- askAt INDENT
|
||||
|
@ -308,25 +321,27 @@ prettyCase_ dnames tnames intro head ret body = do
|
|||
indent ind $ vseparateTight semi bodys, rb])
|
||||
|
||||
private
|
||||
prettyCase : {opts : _} ->
|
||||
BContext d -> BContext n ->
|
||||
Qty -> Elim d n -> ScopeTerm d n -> List (CaseArm opts d n) ->
|
||||
prettyCase : {opts : LayoutOpts} ->
|
||||
NameContexts q d n ->
|
||||
Qty q -> Elim q d n -> ScopeTerm q d n ->
|
||||
List (CaseArm opts q d n) ->
|
||||
Eff Pretty (Doc opts)
|
||||
prettyCase dnames tnames qty head ret body =
|
||||
prettyCase_ dnames tnames ![|caseD <+> prettyQty qty|] head ret body
|
||||
prettyCase names qty head ret body =
|
||||
prettyCase_ names ![|caseD <+> prettyQty names qty|] head ret body
|
||||
|
||||
|
||||
private
|
||||
LetBinder : Nat -> Nat -> Type
|
||||
LetBinder d n = (Qty, BindName, Elim d n)
|
||||
LetBinder : TermLike
|
||||
LetBinder q d n = (Qty q, BindName, Elim q d n)
|
||||
|
||||
private
|
||||
LetExpr : Nat -> Nat -> Nat -> Type
|
||||
LetExpr d n n' = (Telescope (LetBinder d) n n', Term d n')
|
||||
LetExpr : (q, d, n, n' : Nat) -> Type
|
||||
LetExpr q d n n' = (Telescope (LetBinder q d) n n', Term q d n')
|
||||
|
||||
-- [todo] factor out this and the untyped version somehow
|
||||
export
|
||||
splitLet : Telescope (LetBinder d) n n' -> Term d n' -> Exists (LetExpr d n)
|
||||
splitLet : Telescope (LetBinder q d) n n' -> Term q d n' ->
|
||||
Exists (LetExpr q d n)
|
||||
splitLet ys t@(Let qty rhs body _) =
|
||||
splitLet (ys :< (qty, body.name, rhs)) (assert_smaller t body.term)
|
||||
splitLet ys t =
|
||||
|
@ -334,38 +349,40 @@ splitLet ys t =
|
|||
|
||||
private covering
|
||||
prettyLets : {opts : LayoutOpts} ->
|
||||
BContext d -> BContext a -> Telescope (LetBinder d) a b ->
|
||||
NameContexts q d a -> Telescope (LetBinder q d) a b ->
|
||||
Eff Pretty (SnocList (Doc opts))
|
||||
prettyLets dnames xs lets = snd <$> go lets where
|
||||
peelAnn : forall d, n. Elim d n -> Maybe (Term d n, Term d n)
|
||||
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)
|
||||
peelAnn (Ann tm ty _) = Just (tm, ty)
|
||||
peelAnn e = Nothing
|
||||
|
||||
letHeader : Qty -> BindName -> Eff Pretty (Doc opts)
|
||||
letHeader : Qty q -> BindName -> Eff Pretty (Doc opts)
|
||||
letHeader qty x = do
|
||||
lett <- [|letD <+> prettyQty qty|]
|
||||
lett <- [|letD <+> prettyQty qnames qty|]
|
||||
x <- prettyTBind x
|
||||
pure $ lett <++> x
|
||||
|
||||
letBody : forall n. BContext n ->
|
||||
Doc opts -> Elim d n -> Eff Pretty (Doc opts)
|
||||
letBody tnames hdr e = case peelAnn e of
|
||||
Just (tm, ty) => do
|
||||
ty <- withPrec Outer $ assert_total prettyTerm dnames tnames ty
|
||||
tm <- withPrec Outer $ assert_total prettyTerm dnames tnames 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 dnames tnames e
|
||||
eq <- cstD; d <- askAt INDENT
|
||||
inn <- inD
|
||||
pure $ ifMultiline
|
||||
(hsep [hdr, eq, e, inn])
|
||||
(vsep [hdr, indent d $ hsep [eq, e, inn]])
|
||||
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 d) a b ->
|
||||
go : forall b. Telescope (LetBinder q d) a b ->
|
||||
Eff Pretty (BContext b, SnocList (Doc opts))
|
||||
go [<] = pure (xs, [<])
|
||||
go [<] = pure (tnames, [<])
|
||||
go (lets :< (qty, x, rhs)) = do
|
||||
(ys, docs) <- go lets
|
||||
doc <- letBody ys !(letHeader qty x) rhs
|
||||
|
@ -385,7 +402,7 @@ toTel [<] = [<]
|
|||
toTel (ctx :< x) = toTel ctx :< x
|
||||
|
||||
private
|
||||
prettyTyCasePat : {opts : _} ->
|
||||
prettyTyCasePat : {opts : LayoutOpts} ->
|
||||
(k : TyConKind) -> BContext (arity k) ->
|
||||
Eff Pretty (Doc opts)
|
||||
prettyTyCasePat KTYPE [<] = typeD
|
||||
|
@ -402,13 +419,14 @@ prettyTyCasePat KString [<] = stringD
|
|||
prettyTyCasePat KBOX [< a] = bracks =<< prettyTBind a
|
||||
|
||||
|
||||
prettyLambda : {opts : _} -> BContext d -> BContext n ->
|
||||
Term d n -> Eff Pretty (Doc opts)
|
||||
prettyLambda dnames tnames s =
|
||||
parensIfM Outer =<< do
|
||||
let MkSplitLams {dnames = ds, tnames = ts, chunks, body} = splitLams s
|
||||
hangDSingle !(header chunks)
|
||||
!(assert_total prettyTerm (dnames . ds) (tnames . ts) body)
|
||||
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)
|
||||
where
|
||||
introChar : NameSort -> Eff Pretty (Doc opts)
|
||||
introChar T = lamD
|
||||
|
@ -426,198 +444,207 @@ where
|
|||
header cs = sep <$> traverse (\(s, xs) => header1 s (toList xs)) (toList cs)
|
||||
|
||||
|
||||
prettyDisp : {opts : _} -> Universe -> Eff Pretty (Maybe (Doc opts))
|
||||
prettyDisp : {opts : LayoutOpts} -> Universe -> Eff Pretty (Maybe (Doc opts))
|
||||
prettyDisp 0 = pure Nothing
|
||||
prettyDisp u = map Just $ hl Universe =<<
|
||||
ifUnicode (text $ superscript $ show u) (text $ "^" ++ show u)
|
||||
|
||||
|
||||
prettyTerm dnames tnames (TYPE l _) = do
|
||||
prettyTerm names (TYPE l _) = do
|
||||
type <- hl Syntax . text =<< ifUnicode "★" "Type"
|
||||
level <- prettyDisp l
|
||||
pure $ maybe type (type <+>) level
|
||||
|
||||
prettyTerm dnames tnames (IOState _) =
|
||||
prettyTerm _ (IOState _) =
|
||||
ioStateD
|
||||
|
||||
prettyTerm dnames tnames (Pi qty arg res _) =
|
||||
parensIfM Outer =<< do
|
||||
let MkSplitPi {binds, cod} = splitPi [< (qty, res.name, arg)] res.term
|
||||
arr <- arrowD
|
||||
lines <- map (<++> arr) <$> prettyPiBinds dnames tnames binds
|
||||
let tnames = tnames . map pbname binds
|
||||
cod <- withPrec Outer $ prettyTerm dnames tnames (assert_smaller res cod)
|
||||
pure $ sepSingle $ toList $ lines :< cod
|
||||
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 dnames tnames s@(Lam {}) =
|
||||
prettyLambda dnames tnames s
|
||||
prettyTerm names s@(Lam {}) =
|
||||
prettyLambda names s
|
||||
|
||||
prettyTerm dnames tnames (Sig fst snd _) =
|
||||
parensIfM Times =<< do
|
||||
let MkSplitSig {binds, last} = splitSig [< (snd.name, fst)] snd.term
|
||||
times <- timesD
|
||||
lines <- map (<++> times) <$> prettySigBinds dnames tnames binds
|
||||
let tnames = tnames . map Builtin.fst binds
|
||||
last <- withPrec InTimes $
|
||||
prettyTerm dnames tnames (assert_smaller snd last)
|
||||
pure $ sepSingle $ toList $ lines :< last
|
||||
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 dnames tnames p@(Pair fst snd _) =
|
||||
parens =<< do
|
||||
let elems = splitTuple [< fst] snd
|
||||
lines <- for elems $ \t =>
|
||||
withPrec Outer $ prettyTerm dnames tnames $ assert_smaller p t
|
||||
pure $ separateTight !commaD lines
|
||||
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 dnames tnames (Enum cases _) =
|
||||
prettyTerm _ (Enum cases _) =
|
||||
prettyEnum $ SortedSet.toList cases
|
||||
|
||||
prettyTerm dnames tnames (Tag tag _) =
|
||||
prettyTerm _ (Tag tag _) =
|
||||
prettyTag tag
|
||||
|
||||
prettyTerm dnames tnames (Eq (S _ (N ty)) l r _) =
|
||||
parensIfM Eq =<< do
|
||||
l <- withPrec InEq $ prettyTerm dnames tnames l
|
||||
r <- withPrec InEq $ prettyTerm dnames tnames r
|
||||
ty <- withPrec InEq $ prettyTerm dnames tnames ty
|
||||
pure $ sep [l <++> !eqndD, r <++> !colonD, ty]
|
||||
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 dnames tnames (Eq ty l r _) =
|
||||
parensIfM App =<< do
|
||||
ty <- prettyTypeLine dnames tnames ty
|
||||
l <- withPrec Arg $ prettyTerm dnames tnames l
|
||||
r <- withPrec Arg $ prettyTerm dnames tnames r
|
||||
prettyAppD !eqD [ty, l, r]
|
||||
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 dnames tnames s@(DLam {}) =
|
||||
prettyLambda dnames tnames s
|
||||
prettyTerm names s@(DLam {}) =
|
||||
prettyLambda names s
|
||||
|
||||
prettyTerm dnames tnames (NAT _) = natD
|
||||
prettyTerm dnames tnames (Nat n _) = hl Syntax $ pshow n
|
||||
prettyTerm dnames tnames (Succ p _) =
|
||||
parensIfM App =<<
|
||||
prettyAppD !succD [!(withPrec Arg $ prettyTerm dnames tnames p)]
|
||||
prettyTerm _ (NAT _) = natD
|
||||
prettyTerm _ (Nat n _) = hl Syntax $ pshow n
|
||||
prettyTerm names (Succ p _) =
|
||||
prettyAppD !succD [!(withPrec Arg $ prettyTerm names p)]
|
||||
|
||||
prettyTerm dnames tnames (STRING _) = stringD
|
||||
prettyTerm dnames tnames (Str s _) = prettyStrLit s
|
||||
prettyTerm _ (STRING _) = stringD
|
||||
prettyTerm _ (Str s _) = prettyStrLit s
|
||||
|
||||
prettyTerm dnames tnames (BOX qty ty _) =
|
||||
prettyTerm names (BOX qty ty _) =
|
||||
bracks . hcat =<<
|
||||
sequence [prettyQty qty, dotD,
|
||||
withPrec Outer $ prettyTerm dnames tnames ty]
|
||||
sequence [prettyQty names qty, dotD,
|
||||
withPrec Outer $ prettyTerm names ty]
|
||||
|
||||
prettyTerm dnames tnames (Box val _) =
|
||||
bracks =<< withPrec Outer (prettyTerm dnames tnames val)
|
||||
prettyTerm names (Box val _) =
|
||||
bracks =<< withPrec Outer (prettyTerm names val)
|
||||
|
||||
prettyTerm dnames tnames (Let qty rhs body _) = do
|
||||
prettyTerm names (Let qty rhs body _) = do
|
||||
let Evidence _ (lets, body) = splitLet [< (qty, body.name, rhs)] body.term
|
||||
heads <- prettyLets dnames tnames lets
|
||||
let tnames = tnames . map (\(_, x, _) => x) lets
|
||||
body <- withPrec Outer $ assert_total prettyTerm dnames tnames body
|
||||
heads <- prettyLets names lets
|
||||
let names' = extendTermN' (map (fst . snd) lets) names
|
||||
body <- withPrec Outer $ assert_total prettyTerm names' body
|
||||
let lines = toList $ heads :< body
|
||||
pure $ ifMultiline (hsep lines) (vsep lines)
|
||||
|
||||
prettyTerm dnames tnames (E e) =
|
||||
case the (Elim d n) (pushSubsts' e) of
|
||||
Ann tm _ _ => assert_total prettyTerm dnames tnames tm
|
||||
_ => assert_total prettyElim dnames tnames e
|
||||
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
|
||||
|
||||
prettyTerm dnames tnames t0@(CloT (Sub t ph)) =
|
||||
prettyTerm dnames tnames $ assert_smaller t0 $ pushSubstsWith' id ph t
|
||||
prettyTerm names t0@(CloT (Sub t ph)) = do
|
||||
let Val q = names.qtyLen
|
||||
prettyTerm names $ assert_smaller t0 $ pushSubstsWith' id id ph t
|
||||
|
||||
prettyTerm dnames tnames t0@(DCloT (Sub t ph)) =
|
||||
prettyTerm dnames tnames $ assert_smaller t0 $ pushSubstsWith' ph id t
|
||||
prettyTerm names t0@(DCloT (Sub t ph)) = do
|
||||
let Val q = names.qtyLen
|
||||
prettyTerm names $ assert_smaller t0 $ pushSubstsWith' id ph id t
|
||||
|
||||
prettyElim dnames tnames (F x u _) = do
|
||||
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
|
||||
x <- prettyFree x
|
||||
u <- prettyDisp u
|
||||
pure $ maybe x (x <+>) u
|
||||
|
||||
prettyElim dnames tnames (B i _) =
|
||||
prettyTBind $ tnames !!! i
|
||||
prettyElim names (B i _) =
|
||||
prettyTBind $ names.tnames !!! i
|
||||
|
||||
prettyElim dnames tnames e@(App {}) =
|
||||
let (f, xs) = splitApps e in
|
||||
prettyDTApps dnames tnames f xs
|
||||
prettyElim names e@(App {}) = do
|
||||
let Val q = names.qtyLen
|
||||
(f, xs) = splitApps e
|
||||
prettyDTApps names f xs
|
||||
|
||||
prettyElim dnames tnames (CasePair qty pair ret body _) = do
|
||||
prettyElim names (CasePair qty pair ret body _) = do
|
||||
let [< x, y] = body.names
|
||||
pat <- parens . hsep =<< sequence
|
||||
[[|prettyTBind x <+> commaD|], prettyTBind y]
|
||||
prettyCase dnames tnames qty pair ret
|
||||
pat <- parens $ !(prettyTBind x) <+> !commaD <++> !(prettyTBind y)
|
||||
prettyCase names qty pair ret
|
||||
[MkCaseArm pat [<] [< x, y] body.term]
|
||||
|
||||
prettyElim dnames tnames (Fst pair _) =
|
||||
parensIfM App =<< do
|
||||
pair <- prettyTArg dnames tnames (E pair)
|
||||
prettyAppD !fstD [pair]
|
||||
prettyElim names (Fst pair _) = do
|
||||
pair <- prettyTArg names (E pair)
|
||||
prettyAppD !fstD [pair]
|
||||
|
||||
prettyElim dnames tnames (Snd pair _) =
|
||||
parensIfM App =<< do
|
||||
pair <- prettyTArg dnames tnames (E pair)
|
||||
prettyAppD !sndD [pair]
|
||||
prettyElim names (Snd pair _) = do
|
||||
pair <- prettyTArg names (E pair)
|
||||
prettyAppD !sndD [pair]
|
||||
|
||||
prettyElim dnames tnames (CaseEnum qty tag ret arms _) = do
|
||||
prettyElim names (CaseEnum qty tag ret arms _) = do
|
||||
arms <- for (SortedMap.toList arms) $ \(tag, body) =>
|
||||
pure $ MkCaseArm !(prettyTag tag) [<] [<] body
|
||||
prettyCase dnames tnames qty tag ret arms
|
||||
prettyCase names qty tag ret arms
|
||||
|
||||
prettyElim dnames tnames (CaseNat qty qtyIH nat ret zero succ _) = do
|
||||
prettyElim names (CaseNat qty qtyIH nat ret zero succ _) = do
|
||||
let zarm = MkCaseArm !zeroD [<] [<] zero
|
||||
[< p, ih] = succ.names
|
||||
spat0 <- [|succD <++> prettyTBind p|]
|
||||
ihpat0 <- map hcat $ sequence [prettyQty qtyIH, dotD, prettyTBind ih]
|
||||
ihpat0 <- map hcat $ sequence [prettyQty names qtyIH, dotD, prettyTBind ih]
|
||||
spat <- if ih.val == Unused
|
||||
then pure spat0
|
||||
else pure $ hsep [spat0 <+> !commaD, ihpat0]
|
||||
else pure $ spat0 <+> !commaD <++> ihpat0
|
||||
let sarm = MkCaseArm spat [<] [< p, ih] succ.term
|
||||
prettyCase dnames tnames qty nat ret [zarm, sarm]
|
||||
prettyCase names qty nat ret [zarm, sarm]
|
||||
|
||||
prettyElim dnames tnames (CaseBox qty box ret body _) = do
|
||||
prettyElim names (CaseBox qty box ret body _) = do
|
||||
pat <- bracks =<< prettyTBind body.name
|
||||
let arm = MkCaseArm pat [<] [< body.name] body.term
|
||||
prettyCase dnames tnames qty box ret [arm]
|
||||
prettyCase names qty box ret [arm]
|
||||
|
||||
prettyElim dnames tnames e@(DApp {}) =
|
||||
let (f, xs) = splitApps e in
|
||||
prettyDTApps dnames tnames f xs
|
||||
prettyElim names e@(DApp {}) = do
|
||||
let Val q = names.qtyLen
|
||||
(f, xs) = splitApps e
|
||||
prettyDTApps names f xs
|
||||
|
||||
prettyElim dnames tnames (Ann tm ty _) =
|
||||
case the (Term d n) (pushSubsts' tm) of
|
||||
E e => assert_total prettyElim dnames tnames e
|
||||
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
|
||||
_ => do
|
||||
tm <- withPrec AnnL $ assert_total prettyTerm dnames tnames tm
|
||||
ty <- withPrec Outer $ assert_total prettyTerm dnames tnames ty
|
||||
tm <- withPrec AnnL $ assert_total prettyTerm names tm
|
||||
ty <- withPrec Outer $ assert_total prettyTerm names ty
|
||||
parensIfM Outer =<< hangDSingle (tm <++> !annD) ty
|
||||
|
||||
prettyElim dnames tnames (Coe ty p q val _) =
|
||||
parensIfM App =<< do
|
||||
ty <- prettyTypeLine dnames tnames ty
|
||||
p <- prettyDArg dnames p
|
||||
q <- prettyDArg dnames q
|
||||
val <- prettyTArg dnames tnames val
|
||||
prettyAppD !coeD [ty, sep [p, q], val]
|
||||
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 dnames tnames e@(Comp ty p q val r zero one _) =
|
||||
parensIfM App =<< do
|
||||
ty <- assert_total $ prettyTypeLine dnames tnames $ SN ty
|
||||
pq <- sep <$> sequence [prettyDArg dnames p, prettyDArg dnames q]
|
||||
val <- prettyTArg dnames tnames val
|
||||
r <- prettyDArg dnames r
|
||||
arm0 <- [|prettyCompArm dnames tnames Zero zero <+> semiD|]
|
||||
arm1 <- prettyCompArm dnames tnames One one
|
||||
ind <- askAt INDENT
|
||||
layoutComp [ty, pq] val r [arm0, arm1]
|
||||
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 dnames tnames (TypeCase ty ret arms def _) = do
|
||||
prettyElim names (TypeCase ty ret arms def _) = do
|
||||
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
|
||||
prettyCase_ dnames tnames !typecaseD ty (SN ret) $ arms ++ [darm]
|
||||
prettyCase_ names !typecaseD ty (SN ret) $ arms ++ [darm]
|
||||
|
||||
prettyElim dnames tnames e0@(CloE (Sub e ph)) =
|
||||
prettyElim dnames tnames $ assert_smaller e0 $ pushSubstsWith' id ph e
|
||||
prettyElim names e0@(CloE (Sub e ph)) = do
|
||||
let Val q = names.qtyLen
|
||||
prettyElim names $ assert_smaller e0 $ pushSubstsWith' id id ph e
|
||||
|
||||
prettyElim dnames tnames e0@(DCloE (Sub e ph)) =
|
||||
prettyElim dnames tnames $ assert_smaller e0 $ pushSubstsWith' ph id e
|
||||
prettyElim names e0@(DCloE (Sub e ph)) = do
|
||||
let Val q = names.qtyLen
|
||||
prettyElim names $ assert_smaller e0 $ pushSubstsWith' id ph id e
|
||||
|
||||
prettyElim names e0@(QCloE (SubR e ph)) = do
|
||||
let Val q = names.qtyLen
|
||||
prettyElim names $ assert_smaller e0 $ pushSubstsWith' ph id id e
|
||||
|
|
Loading…
Add table
Add a link
Reference in a new issue