always vsep scheme lets, otherwise they are unreadable
This commit is contained in:
parent
4cc50c6bcd
commit
e0ed37720f
29 changed files with 474 additions and 301 deletions
|
@ -61,6 +61,10 @@ mutual
|
|||
||| type of types
|
||||
TYPE : (l : Universe) -> (loc : Loc) -> Term d n
|
||||
|
||||
||| IO state token. this is a builtin because otherwise #[main] being a
|
||||
||| builtin makes no sense
|
||||
IOState : (loc : Loc) -> Term d n
|
||||
|
||||
||| function type
|
||||
Pi : (qty : Qty) -> (arg : Term d n) ->
|
||||
(res : ScopeTerm d n) -> (loc : Loc) -> Term d n
|
||||
|
@ -88,6 +92,10 @@ mutual
|
|||
Zero : (loc : Loc) -> Term d n
|
||||
Succ : (p : Term d n) -> (loc : Loc) -> Term d n
|
||||
|
||||
||| strings
|
||||
STRING : (loc : Loc) -> Term d n
|
||||
Str : (str : String) -> (loc : Loc) -> Term d n
|
||||
|
||||
||| "box" (package a value up with a certain quantity)
|
||||
BOX : (qty : Qty) -> (ty : Term d n) -> (loc : Loc) -> Term d n
|
||||
Box : (val : Term d n) -> (loc : Loc) -> Term d n
|
||||
|
@ -361,6 +369,7 @@ Located (Elim d n) where
|
|||
export
|
||||
Located (Term d n) where
|
||||
(TYPE _ loc).loc = loc
|
||||
(IOState loc).loc = loc
|
||||
(Pi _ _ _ loc).loc = loc
|
||||
(Lam _ loc).loc = loc
|
||||
(Sig _ _ loc).loc = loc
|
||||
|
@ -371,6 +380,8 @@ Located (Term d n) where
|
|||
(DLam _ loc).loc = loc
|
||||
(Nat loc).loc = loc
|
||||
(Zero loc).loc = loc
|
||||
(STRING loc).loc = loc
|
||||
(Str _ loc).loc = loc
|
||||
(Succ _ loc).loc = loc
|
||||
(BOX _ _ loc).loc = loc
|
||||
(Box _ loc).loc = loc
|
||||
|
@ -421,6 +432,7 @@ Relocatable (Elim d n) where
|
|||
export
|
||||
Relocatable (Term d n) where
|
||||
setLoc loc (TYPE l _) = TYPE l loc
|
||||
setLoc loc (IOState _) = IOState loc
|
||||
setLoc loc (Pi qty arg res _) = Pi qty arg res loc
|
||||
setLoc loc (Lam body _) = Lam body loc
|
||||
setLoc loc (Sig fst snd _) = Sig fst snd loc
|
||||
|
@ -432,6 +444,8 @@ Relocatable (Term d n) where
|
|||
setLoc loc (Nat _) = Nat loc
|
||||
setLoc loc (Zero _) = Zero loc
|
||||
setLoc loc (Succ p _) = Succ p loc
|
||||
setLoc loc (STRING _) = STRING loc
|
||||
setLoc loc (Str s _) = Str s loc
|
||||
setLoc loc (BOX qty ty _) = BOX qty ty loc
|
||||
setLoc loc (Box val _) = Box val loc
|
||||
setLoc loc (E e) = E $ setLoc loc e
|
||||
|
|
|
@ -343,6 +343,7 @@ prettyTyCasePat : {opts : _} ->
|
|||
(k : TyConKind) -> BContext (arity k) ->
|
||||
Eff Pretty (Doc opts)
|
||||
prettyTyCasePat KTYPE [<] = typeD
|
||||
prettyTyCasePat KIOState [<] = ioStateD
|
||||
prettyTyCasePat KPi [< a, b] =
|
||||
parens . hsep =<< sequence [prettyTBind a, arrowD, prettyTBind b]
|
||||
prettyTyCasePat KSig [< a, b] =
|
||||
|
@ -351,6 +352,7 @@ 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
|
||||
prettyTyCasePat KBOX [< a] = bracks =<< prettyTBind a
|
||||
|
||||
|
||||
|
@ -392,6 +394,9 @@ prettyTerm dnames tnames (TYPE l _) =
|
|||
pure $ hcat [star, level]
|
||||
Ascii => [|hl Syntax "Type" <++> hl Universe (text $ show l)|]
|
||||
|
||||
prettyTerm dnames tnames (IOState _) =
|
||||
ioStateD
|
||||
|
||||
prettyTerm dnames tnames (Pi qty arg res _) =
|
||||
parensIfM Outer =<< do
|
||||
let MkSplitPi {binds, cod} = splitPi [< (qty, res.name, arg)] res.term
|
||||
|
@ -459,6 +464,9 @@ prettyTerm dnames tnames (Succ p _) = do
|
|||
prettyTerm dnames tnames $ assert_smaller s s'
|
||||
either succ (hl Syntax . text . show . S) =<< toNat p
|
||||
|
||||
prettyTerm dnames tnames (STRING _) = stringD
|
||||
prettyTerm dnames tnames (Str s _) = prettyStrLit s
|
||||
|
||||
prettyTerm dnames tnames (BOX qty ty _) =
|
||||
bracks . hcat =<<
|
||||
sequence [prettyQty qty, dotD,
|
||||
|
|
|
@ -254,6 +254,8 @@ mutual
|
|||
PushSubsts Term Subst.isCloT where
|
||||
pushSubstsWith th ph (TYPE l loc) =
|
||||
nclo $ TYPE l loc
|
||||
pushSubstsWith th ph (IOState loc) =
|
||||
nclo $ IOState loc
|
||||
pushSubstsWith th ph (Pi qty a body loc) =
|
||||
nclo $ Pi qty (a // th // ph) (body // th // ph) loc
|
||||
pushSubstsWith th ph (Lam body loc) =
|
||||
|
@ -276,6 +278,10 @@ mutual
|
|||
nclo $ Zero loc
|
||||
pushSubstsWith th ph (Succ n loc) =
|
||||
nclo $ Succ (n // th // ph) loc
|
||||
pushSubstsWith _ _ (STRING loc) =
|
||||
nclo $ STRING loc
|
||||
pushSubstsWith _ _ (Str s loc) =
|
||||
nclo $ Str s loc
|
||||
pushSubstsWith th ph (BOX pi ty loc) =
|
||||
nclo $ BOX pi (ty // th // ph) loc
|
||||
pushSubstsWith th ph (Box val loc) =
|
||||
|
|
|
@ -44,6 +44,7 @@ mutual
|
|||
tightenT' : OPE n1 n2 -> (t : Term d n2) -> (0 nt : NotClo t) =>
|
||||
Maybe (Term d n1)
|
||||
tightenT' p (TYPE l loc) = pure $ TYPE l loc
|
||||
tightenT' p (IOState loc) = pure $ IOState loc
|
||||
tightenT' p (Pi qty arg res loc) =
|
||||
Pi qty <$> tightenT p arg <*> tightenS p res <*> pure loc
|
||||
tightenT' p (Lam body loc) =
|
||||
|
@ -66,6 +67,10 @@ mutual
|
|||
pure $ Zero loc
|
||||
tightenT' p (Succ s loc) =
|
||||
Succ <$> tightenT p s <*> pure loc
|
||||
tightenT' p (STRING loc) =
|
||||
pure $ STRING loc
|
||||
tightenT' p (Str s loc) =
|
||||
pure $ Str s loc
|
||||
tightenT' p (BOX qty ty loc) =
|
||||
BOX qty <$> tightenT p ty <*> pure loc
|
||||
tightenT' p (Box val loc) =
|
||||
|
@ -163,6 +168,8 @@ mutual
|
|||
Maybe (Term d1 n)
|
||||
dtightenT' p (TYPE l loc) =
|
||||
pure $ TYPE l loc
|
||||
dtightenT' p (IOState loc) =
|
||||
pure $ IOState loc
|
||||
dtightenT' p (Pi qty arg res loc) =
|
||||
Pi qty <$> dtightenT p arg <*> dtightenS p res <*> pure loc
|
||||
dtightenT' p (Lam body loc) =
|
||||
|
@ -185,6 +192,10 @@ mutual
|
|||
pure $ Zero loc
|
||||
dtightenT' p (Succ s loc) =
|
||||
Succ <$> dtightenT p s <*> pure loc
|
||||
dtightenT' p (STRING loc) =
|
||||
pure $ STRING loc
|
||||
dtightenT' p (Str s loc) =
|
||||
pure $ Str s loc
|
||||
dtightenT' p (BOX qty ty loc) =
|
||||
BOX qty <$> dtightenT p ty <*> pure loc
|
||||
dtightenT' p (Box val loc) =
|
||||
|
|
|
@ -9,7 +9,8 @@ import Generics.Derive
|
|||
|
||||
|
||||
public export
|
||||
data TyConKind = KTYPE | KPi | KSig | KEnum | KEq | KNat | KBOX
|
||||
data TyConKind =
|
||||
KTYPE | KIOState | KPi | KSig | KEnum | KEq | KNat | KString | KBOX
|
||||
%name TyConKind k
|
||||
%runElab derive "TyConKind" [Eq.Eq, Ord.Ord, Show.Show, Generic, Meta, DecEq]
|
||||
|
||||
|
@ -25,10 +26,12 @@ allKinds = %runElab do
|
|||
||| in `type-case`, how many variables are bound in this branch
|
||||
public export %inline
|
||||
arity : TyConKind -> Nat
|
||||
arity KTYPE = 0
|
||||
arity KPi = 2
|
||||
arity KSig = 2
|
||||
arity KEnum = 0
|
||||
arity KEq = 5
|
||||
arity KNat = 0
|
||||
arity KBOX = 1
|
||||
arity KTYPE = 0
|
||||
arity KIOState = 0
|
||||
arity KPi = 2
|
||||
arity KSig = 2
|
||||
arity KEnum = 0
|
||||
arity KEq = 5
|
||||
arity KNat = 0
|
||||
arity KString = 0
|
||||
arity KBOX = 1
|
||||
|
|
Loading…
Add table
Add a link
Reference in a new issue