Files
newt/port/Lib/Syntax.newt

297 lines
11 KiB
Agda
Raw Blame History

This file contains ambiguous Unicode characters
This file contains Unicode characters that might be confused with other characters. If you think that this is intentional, you can safely ignore this warning. Use the Escape button to reveal them.
module Lib.Syntax
import Data.String
import Data.Maybe
import Lib.Parser.Impl
import Lib.Prettier
import Lib.Types
Raw : U
data Pattern
= PatVar FC Icit Name
| PatCon FC Icit QName (List Pattern) (Maybe Name)
| PatWild FC Icit
-- Not handling this yet, but we need to be able to work with numbers and strings...
| PatLit FC Literal
getIcit : Pattern -> Icit
getIcit (PatVar x icit str) = icit
getIcit (PatCon x icit str xs as) = icit
getIcit (PatWild x icit) = icit
getIcit (PatLit fc lit) = Explicit
instance HasFC Pattern where
getFC (PatVar fc _ _) = fc
getFC (PatCon fc _ _ _ _) = fc
getFC (PatWild fc _) = fc
getFC (PatLit fc lit) = fc
-- %runElab deriveShow `{Pattern}
Constraint : U
Constraint = (String × Pattern)
record Clause where
constructor MkClause
clauseFC : FC
-- I'm including the type of the left, so we can check pats and get the list of possibilities
-- But maybe rethink what happens on the left.
-- It's a VVar k or possibly a pattern.
-- a pattern either is zipped out, dropped (non-match) or is assigned to rhs
-- if we can do all three then we can have a VVar here.
cons : List Constraint
pats : List Pattern
-- We'll need some context to typecheck this
-- it has names from Pats, which will need types in the env
expr : Raw
-- could be a pair, but I suspect stuff will be added?
data RCaseAlt = MkAlt Raw Raw
data DoStmt : U where
DoExpr : (fc : FC) -> Raw -> DoStmt
DoLet : (fc : FC) -> String -> Raw -> DoStmt
DoArrow : (fc : FC) -> Raw -> Raw -> List RCaseAlt -> DoStmt
Decl : U
data Raw : U where
RVar : (fc : FC) -> (nm : Name) -> Raw
RLam : (fc : FC) -> BindInfo -> (ty : Raw) -> Raw
RApp : (fc : FC) -> (t : Raw) -> (u : Raw) -> (icit : Icit) -> Raw
RU : (fc : FC) -> Raw
RPi : (fc : FC) -> BindInfo -> (ty : Raw) -> (sc : Raw) -> Raw
RLet : (fc : FC) -> (nm : Name) -> (ty : Raw) -> (v : Raw) -> (sc : Raw) -> Raw
RAnn : (fc : FC) -> (tm : Raw) -> (ty : Raw) -> Raw
RLit : (fc : FC) -> Literal -> Raw
RCase : (fc : FC) -> (scrut : Raw) -> (alts : List RCaseAlt) -> Raw
RImplicit : (fc : FC) -> Raw
RHole : (fc : FC) -> Raw
RDo : (fc : FC) -> List DoStmt -> Raw
RIf : (fc : FC) -> Raw -> Raw -> Raw -> Raw
RWhere : (fc : FC) -> (List Decl) -> Raw -> Raw
RAs : (fc : FC) -> Name -> Raw -> Raw
instance HasFC Raw where
getFC (RVar fc nm) = fc
getFC (RLam fc _ ty) = fc
getFC (RApp fc t u icit) = fc
getFC (RU fc) = fc
getFC (RPi fc _ ty sc) = fc
getFC (RLet fc nm ty v sc) = fc
getFC (RAnn fc tm ty) = fc
getFC (RLit fc y) = fc
getFC (RCase fc scrut alts) = fc
getFC (RImplicit fc) = fc
getFC (RHole fc) = fc
getFC (RDo fc stmts) = fc
getFC (RIf fc _ _ _) = fc
getFC (RWhere fc _ _) = fc
getFC (RAs fc _ _) = fc
data Import = MkImport FC Name
Telescope : U
Telescope = List (BindInfo × Raw)
data Decl
= TypeSig FC (List Name) Raw
| Def FC Name (List (Raw × Raw)) -- (List Clause)
| DCheck FC Raw Raw
| Data FC Name Raw (List Decl)
| ShortData FC Raw (List Raw)
| PType FC Name (Maybe Raw)
| PFunc FC Name (List String) Raw String
| PMixFix FC (List Name) Int Fixity
| Class FC Name Telescope (List Decl)
| Instance FC Raw (Maybe (List Decl))
| Record FC Name Telescope (Maybe Name) (List Decl)
instance HasFC Decl where
getFC (TypeSig x strs tm) = x
getFC (Def x str xs) = x
getFC (DCheck x tm tm1) = x
getFC (Data x str tm xs) = x
getFC (ShortData x _ _) = x
getFC (PType x str mtm) = x
getFC (PFunc x str _ tm str1) = x
getFC (PMixFix x strs k y) = x
getFC (Class x str xs ys) = x
getFC (Instance x tm xs) = x
getFC (Record x str tm str1 xs) = x
record Module where
constructor MkModule
modname : Name
imports : List Import
decls : List Decl
foo : List String -> String
foo ts = "(" ++ unwords ts ++ ")"
instance Show Raw
instance Show Pattern
instance Show Clause where
show (MkClause fc cons pats expr) = show (fc, cons, pats, expr)
instance Show Import where
show (MkImport _ str) = foo ("MkImport" :: show str :: Nil)
instance Show BindInfo where
show (BI _ nm icit quant) = foo ("BI" :: show nm :: show icit :: show quant :: Nil)
-- this is for debugging, use pretty when possible
instance Show Decl where
show (TypeSig _ str x) = foo ("TypeSig" :: show str :: show x :: Nil)
show (Def _ str clauses) = foo ("Def" :: show str :: show clauses :: Nil)
show (Data _ str xs ys) = foo ("Data" :: show str :: show xs :: show ys :: Nil)
show (DCheck _ x y) = foo ("DCheck" :: show x :: show y :: Nil)
show (PType _ name ty) = foo ("PType" :: name :: show ty :: Nil)
show (ShortData _ lhs sigs) = foo ("ShortData" :: show lhs :: show sigs :: Nil)
show (PFunc _ nm used ty src) = foo ("PFunc" :: nm :: show used :: show ty :: show src :: Nil)
show (PMixFix _ nms prec fix) = foo ("PMixFix" :: show nms :: show prec :: show fix :: Nil)
show (Class _ nm tele decls) = foo ("Class" :: nm :: "..." :: (show $ map show decls) :: Nil)
show (Instance _ nm decls) = foo ("Instance" :: show nm :: (show $ map show decls) :: Nil)
show (Record _ nm tele nm1 decls) = foo ("Record" :: show nm :: show tele :: show nm1 :: show decls :: Nil)
instance Show Module where
show (MkModule name imports decls) = foo ("MkModule" :: show name :: show imports :: show decls :: Nil)
instance Show Pattern where
show (PatVar _ icit str) = foo ("PatVar" :: show icit :: show str :: Nil)
show (PatCon _ icit str xs as) = foo ("PatCon" :: show icit :: show str :: show xs :: show as :: Nil)
show (PatWild _ icit) = foo ("PatWild" :: show icit :: Nil)
show (PatLit _ lit) = foo ("PatLit" :: show lit :: Nil)
instance Show RCaseAlt where
show (MkAlt x y)= foo ("MkAlt" :: show x :: show y :: Nil)
instance Show Raw where
show (RImplicit _) = "_"
show (RHole _) = "?"
show (RVar _ name) = foo ("RVar" :: show name :: Nil)
show (RAnn _ t ty) = foo ( "RAnn" :: show t :: show ty :: Nil)
show (RLit _ x) = foo ( "RLit" :: show x :: Nil)
show (RLet _ x ty v scope) = foo ( "Let" :: show x :: " : " :: show ty :: " = " :: show v :: " in " :: show scope :: Nil)
show (RPi _ bi y z) = foo ( "Pi" :: show bi :: show y :: show z :: Nil)
show (RApp _ x y z) = foo ( "App" :: show x :: show y :: show z :: Nil)
show (RLam _ bi y) = foo ( "Lam" :: show bi :: show y :: Nil)
show (RCase _ x xs) = foo ( "Case" :: show x :: show xs :: Nil)
show (RDo _ stmts) = foo ( "DO" :: "FIXME" :: Nil)
show (RU _) = "U"
show (RIf _ x y z) = foo ( "If" :: show x :: show y :: show z :: Nil)
show (RWhere _ _ _) = foo ( "Where" :: "FIXME" :: Nil)
show (RAs _ nm x) = foo ( "RAs" :: nm :: show x :: Nil)
instance Pretty Literal where
pretty (LString t) = text t
pretty (LInt i) = text $ show i
pretty (LChar c) = text $ show c
instance Pretty Pattern where
-- FIXME - wrap Implicit with {}
pretty (PatVar _ icit str) = text str
pretty (PatCon _ icit nm args Nothing) = text (show nm) <+> spread (map pretty args)
pretty (PatCon _ icit nm args (Just as)) = text as ++ text "@(" ++ text (show nm) <+> spread (map pretty args) ++ text ")"
pretty (PatWild _ icit) = text "_"
pretty (PatLit _ lit) = pretty lit
wrap : Icit -> Doc -> Doc
wrap Explicit x = text "(" ++ x ++ text ")"
wrap Implicit x = text "{" ++ x ++ text "}"
wrap Auto x = text "{{" ++ x ++ text "}}"
instance Pretty Raw where
pretty = asDoc 0
where
bindDoc : BindInfo -> Doc
bindDoc (BI _ nm icit quant) = text "BINDDOC"
par : Int -> Int -> Doc -> Doc
par p p' d = if p' < p then text "(" ++ d ++ text ")" else d
asDoc : Int -> Raw -> Doc
asDoc p (RVar _ str) = text str
asDoc p (RLam _ (BI _ nm icit q) x) = par p 0 $ text "\\" ++ wrap icit (text nm) <+> text "=>" <+> asDoc 0 x
-- This needs precedence and operators...
asDoc p (RApp _ x y Explicit) = par p 2 $ asDoc 2 x <+> asDoc 3 y
asDoc p (RApp _ x y Implicit) = par p 2 $ asDoc 2 x <+> text "{" ++ asDoc 0 y ++ text "}"
asDoc p (RApp _ x y Auto) = par p 2 $ asDoc 2 x <+> text "{{" ++ asDoc 0 y ++ text "}}"
asDoc p (RU _) = text "U"
asDoc p (RPi _ (BI _ "_" Explicit Many) ty scope) = par p 1 $ asDoc p ty <+> text "->" <+/> asDoc p scope
asDoc p (RPi _ (BI _ nm icit quant) ty scope) =
par p 1 $ wrap icit (text (show quant ++ nm) <+> text ":" <+> asDoc p ty ) <+> text "->" <+/> asDoc 1 scope
asDoc p (RLet _ x v ty scope) =
par p 0 $ text "let" <+> text x <+> text ":" <+> asDoc p ty
<+> text "=" <+> asDoc p v
<+/> text "in" <+> asDoc p scope
-- does this exist?
asDoc p (RAnn _ x y) = text "TODO - RAnn"
asDoc p (RLit _ lit) = pretty lit
asDoc p (RCase _ x xs) = text "TODO - RCase"
asDoc p (RImplicit _) = text "_"
asDoc p (RHole _) = text "?"
asDoc p (RDo _ stmts) = text "TODO - RDo"
asDoc p (RIf _ x y z) = par p 0 $ text "if" <+> asDoc 0 x <+/> text "then" <+> asDoc 0 y <+/> text "else" <+> asDoc 0 z
asDoc p (RWhere _ dd b) = text "TODO pretty where"
asDoc p (RAs _ nm x) = text nm ++ text "@(" ++ asDoc 0 x ++ text ")"
prettyBind : (BindInfo × Raw) -> Doc
prettyBind (BI _ nm icit quant, ty) = wrap icit (text (show quant ++ nm) <+> text ":" <+> pretty ty)
pipeSep : List Doc -> Doc
pipeSep = folddoc (\a b => a <+/> text "|" <+> b)
instance Pretty Decl where
pretty (TypeSig _ nm ty) = spread (map text nm) <+> text ":" <+> nest 2 (pretty ty)
pretty (Def _ nm clauses) = stack $ map prettyPair clauses
where
prettyPair : Raw × Raw Doc
prettyPair (a, b) = pretty a <+> text "=" <+> pretty b
pretty (Data _ nm x xs) = text "data" <+> text nm <+> text ":" <+> pretty x <+> (nest 2 $ text "where" </> stack (map pretty xs))
pretty (DCheck _ x y) = text "#check" <+> pretty x <+> text ":" <+> pretty y
pretty (PType _ nm ty) = text "ptype" <+> text nm <+> (maybe empty (\ty => text ":" <+> pretty ty) ty)
pretty (PFunc _ nm Nil ty src) = text "pfunc" <+> text nm <+> text ":" <+> nest 2 (pretty ty <+> text ":=" <+/> text (show src))
pretty (PFunc _ nm used ty src) = text "pfunc" <+> text nm <+> text "uses" <+> spread (map text used) <+> text ":" <+> nest 2 (pretty ty <+> text ":=" <+/> text (show src))
pretty (PMixFix _ names prec fix) = text (show fix) <+> text (show prec) <+> spread (map text names)
pretty (Record _ nm tele cname decls) = text "record" <+> text nm <+> text ":" <+> spread (map prettyBind tele)
<+> (nest 2 $ text "where" </> stack (maybe empty (\ nm' => text "constructor" <+> text nm') cname :: map pretty decls))
pretty (Class _ nm tele decls) = text "class" <+> text nm <+> text ":" <+> spread (map prettyBind tele)
<+> (nest 2 $ text "where" </> stack (map pretty decls))
pretty (Instance _ _ _) = text "TODO pretty Instance"
pretty (ShortData _ lhs sigs) = text "data" <+> pretty lhs <+> text "=" <+> pipeSep (map pretty sigs)
instance Pretty Module where
pretty (MkModule name imports decls) =
text "module" <+> text name
</> stack (map doImport imports)
</> stack (map pretty decls)
where
doImport : Import -> Doc
doImport (MkImport _ nm) = text "import" <+> text nm ++ line