add W to internal language

This commit is contained in:
rhiannon morris 2023-08-06 10:46:55 +02:00
parent 00d92d3f25
commit 1da902d13a
14 changed files with 544 additions and 94 deletions

View file

@ -102,6 +102,16 @@ mutual
||| pair value
Pair : (fst, snd : Term d n) -> (loc : Loc) -> Term d n
||| inductive (w) type `(x : A) ⊲ B`
W : (shape : Term d n) ->
(body : ScopeTerm d n) -> (loc : Loc) -> Term d n
||| subterms for `(x : A) ⊲ B` are:
||| 1. `x : A`
||| (the "constructor" and non-recursive fields)
||| 2. `f : 1.(B x) → (x : A) ⊲ B`
||| (the recursive fields, one for each element of B x)
Sup : (root, sub : Term d n) -> (loc : Loc) -> Term d n
||| enumeration type
Enum : (cases : SortedSet TagVal) -> (loc : Loc) -> Term d n
||| enumeration value
@ -155,6 +165,13 @@ mutual
(loc : Loc) ->
Elim d n
||| recursion
CaseW : (qty, qtyIH : Qty) -> (tree : Elim d n) ->
(ret : ScopeTerm d n) ->
(body : ScopeTermN 3 d n) ->
(loc : Loc) ->
Elim d n
||| enum matching
CaseEnum : (qty : Qty) -> (tag : Elim d n) ->
(ret : ScopeTerm d n) ->
@ -364,6 +381,7 @@ Located (Elim d n) where
(B _ loc).loc = loc
(App _ _ loc).loc = loc
(CasePair _ _ _ _ loc).loc = loc
(CaseW _ _ _ _ _ loc).loc = loc
(CaseEnum _ _ _ _ loc).loc = loc
(CaseNat _ _ _ _ _ _ loc).loc = loc
(CaseBox _ _ _ _ loc).loc = loc
@ -382,6 +400,8 @@ Located (Term d n) where
(Lam _ loc).loc = loc
(Sig _ _ loc).loc = loc
(Pair _ _ loc).loc = loc
(W _ _ loc).loc = loc
(Sup _ _ loc).loc = loc
(Enum _ loc).loc = loc
(Tag _ loc).loc = loc
(Eq _ _ _ loc).loc = loc
@ -412,6 +432,8 @@ Relocatable (Elim d n) where
setLoc loc (App fun arg _) = App fun arg loc
setLoc loc (CasePair qty pair ret body _) =
CasePair qty pair ret body loc
setLoc loc (CaseW qty qtyIH tree ret body _) =
CaseW qty qtyIH tree ret body loc
setLoc loc (CaseEnum qty tag ret arms _) =
CaseEnum qty tag ret arms loc
setLoc loc (CaseNat qty qtyIH nat ret zero succ _) =
@ -440,6 +462,8 @@ Relocatable (Term d n) where
setLoc loc (Lam body _) = Lam body loc
setLoc loc (Sig fst snd _) = Sig fst snd loc
setLoc loc (Pair fst snd _) = Pair fst snd loc
setLoc loc (W shape body _) = W shape body loc
setLoc loc (Sup root sub _) = Sup root sub loc
setLoc loc (Enum cases _) = Enum cases loc
setLoc loc (Tag tag _) = Tag tag loc
setLoc loc (Eq ty l r _) = Eq ty l r loc

View file

@ -346,6 +346,8 @@ 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 KW [< a, b] =
parens . hsep =<< sequence [prettyTBind a, triD, prettyTBind b]
prettyTyCasePat KEnum [<] = hl Syntax $ text "{}"
prettyTyCasePat KEq [< a0, a1, a, l, r] =
hsep <$> sequence (eqD :: map prettyTBind [a0, a1, a, l, r])
@ -420,6 +422,21 @@ prettyTerm dnames tnames p@(Pair fst snd _) =
withPrec Outer $ prettyTerm dnames tnames $ assert_smaller p t
pure $ separateTight !commaD lines
prettyTerm dnames tnames (W a b _) = do
parensIfM W =<< do
left <- prettySigBind1 b.name dnames tnames a
right <- withPrec InW $
prettyTerm dnames (tnames :< b.name) (assert_smaller b b.term)
pure $ sep [hsep [left, !triD], right]
prettyTerm dnames tnames (Sup root sub _) = do
parensIfM Sup =<< do
left <- withPrec InSup $ prettyTerm dnames tnames root
right <- withPrec InSup $ prettyTerm dnames tnames sub
pure $
hsep [left, !diamondD, right] <|>
vsep [hsep [left, !diamondD], !(indentD right)]
prettyTerm dnames tnames (Enum cases _) =
prettyEnum $ SortedSet.toList cases
@ -446,17 +463,22 @@ prettyTerm dnames tnames s@(DLam {}) =
prettyTerm dnames tnames (Nat _) = natD
prettyTerm dnames tnames (Zero _) = hl Syntax "0"
prettyTerm dnames tnames (Succ p _) = do
succD <- succD
let succ : Doc opts -> Eff Pretty (Doc opts)
succ t = prettyAppD succD [t]
toNat : Term d n -> Eff Pretty (Either (Doc opts) Nat)
toNat s with (pushSubsts' s)
_ | Zero _ = pure $ Right 0
_ | Succ d _ = bitraverse succ (pure . S) =<<
toNat (assert_smaller s d)
_ | s' = map Left . withPrec Arg $
prettyTerm dnames tnames $ assert_smaller s s'
either succ (hl Syntax . text . show . S) =<< toNat p
s <- succD
either (succ s) prettyNat =<< tryToNat s p
where
succ : Doc opts -> Doc opts -> Eff Pretty (Doc opts)
succ s t = prettyAppD s [t]
tryToNat : Doc opts -> Term d n -> Eff Pretty (Either (Doc opts) Nat)
tryToNat s t with (pushSubsts' t)
_ | Zero _ = pure $ Right 0
_ | Succ d _ = bitraverse (succ s) (pure . S) =<<
tryToNat s (assert_smaller t d)
_ | t' = map Left . withPrec Arg $
prettyTerm dnames tnames $ assert_smaller t t'
prettyNat : Nat -> Eff Pretty (Doc opts)
prettyNat = hl Syntax . text . show . S
prettyTerm dnames tnames (BOX qty ty _) =
bracks . hcat =<<
@ -493,6 +515,16 @@ prettyElim dnames tnames (CasePair qty pair ret body _) = do
prettyCase dnames tnames qty pair ret
[MkCaseArm pat [<] [< x, y] body.term]
prettyElim dnames tnames (CaseW qty qtyIH tree ret body _) = do
let [< t, r, ih] = body.names
pat0 <- hsep <$> sequence [prettyTBind t, diamondD, prettyTBind r]
ihpat <- map hcat $ sequence [prettyQty qtyIH, dotD, prettyTBind ih]
pat <- if ih.name == Unused
then pure pat0
else pure $ hsep [pat0 <+> !commaD, ihpat]
let arm = MkCaseArm pat [<] [< t, r, ih] body.term
prettyCase dnames tnames qty tree ret [arm]
prettyElim dnames tnames (CaseEnum qty tag ret arms _) = do
arms <- for (SortedMap.toList arms) $ \(tag, body) =>
pure $ MkCaseArm !(prettyTag tag) [<] [<] body

View file

@ -261,6 +261,10 @@ mutual
nclo $ Sig (a // th // ph) (b // th // ph) loc
pushSubstsWith th ph (Pair s t loc) =
nclo $ Pair (s // th // ph) (t // th // ph) loc
pushSubstsWith th ph (W a b loc) =
nclo $ W (a // th // ph) (b // th // ph) loc
pushSubstsWith th ph (Sup s t loc) =
nclo $ Sup (s // th // ph) (t // th // ph) loc
pushSubstsWith th ph (Enum tags loc) =
nclo $ Enum tags loc
pushSubstsWith th ph (Tag tag loc) =
@ -299,6 +303,8 @@ mutual
nclo $ App (f // th // ph) (s // th // ph) loc
pushSubstsWith th ph (CasePair pi p r b loc) =
nclo $ CasePair pi (p // th // ph) (r // th // ph) (b // th // ph) loc
pushSubstsWith th ph (CaseW pi pi' e r b loc) =
nclo $ CaseW pi pi' (e // th // ph) (r // th // ph) (b // th // ph) loc
pushSubstsWith th ph (CaseEnum pi t r arms loc) =
nclo $ CaseEnum pi (t // th // ph) (r // th // ph)
(map (\b => b // th // ph) arms) loc

View file

@ -52,6 +52,10 @@ mutual
Sig <$> tightenT p fst <*> tightenS p snd <*> pure loc
tightenT' p (Pair fst snd loc) =
Pair <$> tightenT p fst <*> tightenT p snd <*> pure loc
tightenT' p (W shape body loc) =
W <$> tightenT p shape <*> tightenS p body <*> pure loc
tightenT' p (Sup root sub loc) =
Sup <$> tightenT p root <*> tightenT p sub <*> pure loc
tightenT' p (Enum cases loc) =
pure $ Enum cases loc
tightenT' p (Tag tag loc) =
@ -87,6 +91,12 @@ mutual
<*> tightenS p ret
<*> tightenS p body
<*> pure loc
tightenE' p (CaseW qty qtyIH tree ret body loc) =
CaseW qty qtyIH
<$> tightenE p tree
<*> tightenS p ret
<*> tightenS p body
<*> pure loc
tightenE' p (CaseEnum qty tag ret arms loc) =
CaseEnum qty <$> tightenE p tag
<*> tightenS p ret
@ -167,6 +177,10 @@ mutual
Sig <$> dtightenT p fst <*> dtightenS p snd <*> pure loc
dtightenT' p (Pair fst snd loc) =
Pair <$> dtightenT p fst <*> dtightenT p snd <*> pure loc
dtightenT' p (W shape body loc) =
W <$> dtightenT p shape <*> dtightenS p body <*> pure loc
dtightenT' p (Sup root sub loc) =
Sup <$> dtightenT p root <*> dtightenT p sub <*> pure loc
dtightenT' p (Enum cases loc) =
pure $ Enum cases loc
dtightenT' p (Tag tag loc) =
@ -202,6 +216,12 @@ mutual
<*> dtightenS p ret
<*> dtightenS p body
<*> pure loc
dtightenE' p (CaseW qty qtyIH tree ret body loc) =
CaseW qty qtyIH
<$> dtightenE p tree
<*> dtightenS p ret
<*> dtightenS p body
<*> pure loc
dtightenE' p (CaseEnum qty tag ret arms loc) =
CaseEnum qty <$> dtightenE p tag
<*> dtightenS p ret

View file

@ -9,7 +9,7 @@ import Generics.Derive
public export
data TyConKind = KTYPE | KPi | KSig | KEnum | KEq | KNat | KBOX
data TyConKind = KTYPE | KPi | KSig | KW | KEnum | KEq | KNat | KBOX
%name TyConKind k
%runElab derive "TyConKind" [Eq.Eq, Ord.Ord, Show.Show, Generic, Meta, DecEq]
@ -28,6 +28,7 @@ arity : TyConKind -> Nat
arity KTYPE = 0
arity KPi = 2
arity KSig = 2
arity KW = 2
arity KEnum = 0
arity KEq = 5
arity KNat = 0