pi/src/Parser/Parse.idr

471 lines
13 KiB
Idris
Raw Normal View History

2022-07-23 03:38:15 +02:00
module Parser.Parse
import Text.Lexer
import Text.Parser
import Data.Vect
import Data.List
import Data.List1
import Data.Maybe
-- eventually start using elaborator representation here, or just a parser
-- exclusive one
import Core.Term
import Core.Misc
%default total
data PiTokenKind
= PTLambda
| PTIdentifier
| PTDot
| PTArrow
| PTPi
| PTSigma
| PTComma
| PTLParen
| PTRParen
| PTColon
| PTLet
| PTDefEq
| PTIn
| PTIgnore
| PTType
| PTTop
| PTStar
| PTTopInd
| PTBot
| PTBotInd
| PTNat
2022-07-26 07:57:44 +02:00
| PTNum
2022-07-23 03:38:15 +02:00
| PTSuc
| PTNatInd
| PTSigInd
| PTId
| PTRefl
| PTJ
2022-07-23 03:38:15 +02:00
Eq PiTokenKind where
(==) PTLambda PTLambda = True
(==) PTIdentifier PTIdentifier = True
(==) PTDot PTDot = True
(==) PTArrow PTArrow = True
(==) PTPi PTPi = True
(==) PTSigma PTSigma = True
(==) PTComma PTComma = True
(==) PTLParen PTLParen = True
(==) PTRParen PTRParen = True
(==) PTColon PTColon = True
(==) PTLet PTLet = True
(==) PTDefEq PTDefEq = True
(==) PTIn PTIn = True
(==) PTIgnore PTIgnore = True
(==) PTType PTType = True
(==) PTTop PTTop = True
(==) PTStar PTStar = True
(==) PTTopInd PTTopInd = True
(==) PTBot PTBot = True
(==) PTBotInd PTBotInd = True
(==) PTNat PTNat = True
2022-07-26 07:57:44 +02:00
(==) PTNum PTNum = True
2022-07-23 03:38:15 +02:00
(==) PTSuc PTSuc = True
(==) PTNatInd PTNatInd = True
(==) PTSigInd PTSigInd = True
(==) PTId PTId = True
(==) PTRefl PTRefl = True
(==) PTJ PTJ = True
2022-07-23 03:38:15 +02:00
(==) _ _ = False
Show PiTokenKind where
show PTLambda = "PTLambda"
show PTIdentifier = "PTIdentifier"
show PTDot = "PTDot"
show PTArrow = "PTArrow"
show PTPi = "PTPi"
show PTSigma = "PTSigma"
show PTComma = "PTComma"
show PTLParen = "PTLParen"
show PTRParen = "PTRParen"
show PTColon = "PTColon"
show PTLet = "PTLet"
show PTDefEq = "PTDefEq"
show PTIn = "PTIn"
show PTIgnore = "PTIgnore"
show PTType = "PTType"
show PTTop = "PTTop"
show PTStar = "PTSTar"
show PTTopInd = "PTTopInd"
show PTBot = "PTBot"
show PTBotInd = "PTBotInd"
show PTNat = "PTNat"
2022-07-26 07:57:44 +02:00
show PTNum = "PTNum "
2022-07-23 03:38:15 +02:00
show PTSuc = "PTSuc"
show PTNatInd = "PTNatInd"
show PTSigInd = "PTSigInd"
show PTId = "PTId"
show PTRefl = "PTRefl"
show PTJ = "PTJ"
2022-07-23 03:38:15 +02:00
PiToken : Type
PiToken = Token PiTokenKind
Show PiToken where
show (Tok kind text) = "Tok kind: " ++ show kind ++ " text: " ++ text
TokenKind PiTokenKind where
TokType PTIdentifier = String
2022-07-26 07:57:44 +02:00
TokType PTNum = Nat
2022-07-23 03:38:15 +02:00
TokType _ = ()
tokValue PTLambda _ = ()
tokValue PTIdentifier s = s
tokValue PTDot _ = ()
tokValue PTArrow _ = ()
tokValue PTPi _ = ()
tokValue PTSigma _ = ()
tokValue PTComma _ = ()
tokValue PTLParen _ = ()
tokValue PTRParen _ = ()
tokValue PTColon _ = ()
tokValue PTLet _ = ()
tokValue PTDefEq _ = ()
tokValue PTIn _ = ()
tokValue PTIgnore _ = ()
tokValue PTType _ = ()
tokValue PTTop _ = ()
tokValue PTStar _ = ()
tokValue PTTopInd _ = ()
tokValue PTBot _ = ()
tokValue PTBotInd _ = ()
tokValue PTNat _ = ()
2022-07-26 07:57:44 +02:00
tokValue PTNum s = cast s
2022-07-23 03:38:15 +02:00
tokValue PTSuc _ = ()
tokValue PTNatInd _ = ()
tokValue PTSigInd _ = ()
tokValue PTId _ = ()
tokValue PTRefl _ = ()
tokValue PTJ _ = ()
2022-07-23 03:38:15 +02:00
ignored : WithBounds PiToken -> Bool
ignored (MkBounded (Tok PTIgnore _) _ _) = True
ignored _ = False
identifier : Lexer
identifier = some (alphaNum <|> oneOf "-_")
2022-07-23 03:38:15 +02:00
keywords : List (String, PiTokenKind)
keywords = [
("in", PTIn),
("let", PTLet),
("suc", PTSuc),
("Type", PTType),
("J", PTJ),
("Id", PTId),
("refl", PTRefl)
]
tokenmap : List (Lexer, PiTokenKind)
tokenmap = [
(exact "-ind", PTTopInd),
(exact "⊥-ind", PTBotInd),
(exact "-ind", PTNatInd),
(exact "Σ-ind", PTSigInd),
(is 'λ', PTLambda),
(is '.', PTDot),
(is '', PTArrow),
(is 'Π', PTPi),
(is '', PTPi),
(is 'Σ', PTSigma),
(is '', PTSigma),
(is ',', PTComma),
(is '(', PTLParen),
(is ')', PTRParen),
(is ':', PTColon),
(is '', PTDefEq),
(is '', PTTop),
(is '', PTBot),
(is '', PTNat),
(is '', PTStar),
2022-07-26 07:57:44 +02:00
(digits, PTNum)
2022-07-23 03:38:15 +02:00
]
piTokenMap : TokenMap PiToken
2022-07-26 07:57:44 +02:00
piTokenMap =
(spaces, Tok PTIgnore) ::
(lineComment (exact "--"), Tok PTIgnore) ::
(blockComment (exact "{-") (exact "-}"), Tok PTIgnore) ::
toTokenMap tokenmap ++
2022-07-23 03:38:15 +02:00
[(identifier, \s =>
case lookup s keywords of
(Just kind) => Tok kind s
Nothing => Tok PTIdentifier s
)
]
lexPi : String -> Maybe (List (WithBounds PiToken))
lexPi str =
case lex piTokenMap str of
(tokens, _, _, "") => Just tokens
_ => Nothing
{- Basic idea, parsing has a list of the identifiers to convert to -}
{- de bruijn indecies, and a Nat to keep track of context size -}
mutual
2022-07-26 23:07:13 +02:00
expr : List String -> (n : Nat) -> Vect n String -> Grammar () PiToken True (Term n)
expr defs n env = tpi defs n env
<|> tsigma defs n env
<|> tarr defs n env
<|> (do
e <- expr1 defs n env
tapp defs n env e <|> pure e)
expr1 : List String -> (n : Nat) -> Vect n String -> Grammar () PiToken True (Term n)
expr1 defs n env = ttopind defs n env
<|> tbotind defs n env
<|> tsuc defs n env
<|> tnatind defs n env
<|> tsigind defs n env
<|> tid defs n env
<|> trefl defs n env
<|> tj defs n env
<|> (do
t <- term defs n env
tapp defs n env t <|> pure t)
term : List String -> (n : Nat) -> Vect n String -> Grammar () PiToken True (Term n)
term defs n env = ttype
<|> ttop
<|> tstar
<|> tbot
<|> tnat
<|> tnum
<|> tpair defs n env
<|> tlet defs n env
<|> tlam defs n env
<|> tvar defs n env
<|> paren defs n env
2022-07-23 03:38:15 +02:00
ttype : Grammar () PiToken True (Term n)
2022-07-23 03:38:15 +02:00
ttype = match PTType >> pure TType
ttop : Grammar () PiToken True (Term n)
2022-07-23 03:38:15 +02:00
ttop = match PTTop >> pure TTop
tstar : Grammar () PiToken True (Term n)
2022-07-23 03:38:15 +02:00
tstar = match PTStar >> pure TStar
2022-07-26 23:07:13 +02:00
ttopind : List String -> (n : Nat) -> Vect n String -> Grammar () PiToken True (Term n)
ttopind defs n env = do
2022-07-23 03:38:15 +02:00
match PTTopInd
commit
2022-07-26 23:07:13 +02:00
c <- term defs n env
st <- term defs n env
2022-07-23 03:38:15 +02:00
pure (TTopInd c st)
tbot : Grammar () PiToken True (Term n)
2022-07-23 03:38:15 +02:00
tbot = match PTBot >> pure TBot
2022-07-26 23:07:13 +02:00
tbotind : List String -> (n : Nat) -> Vect n String -> Grammar () PiToken True (Term n)
tbotind defs n env = do
2022-07-23 03:38:15 +02:00
match PTBotInd
commit
2022-07-26 23:07:13 +02:00
c <- term defs n env
2022-07-23 03:38:15 +02:00
pure (TBotInd c)
tnat : Grammar () PiToken True (Term n)
2022-07-23 03:38:15 +02:00
tnat = match PTNat >> pure TNat
2022-07-26 07:57:44 +02:00
tnum : Grammar () PiToken True (Term n)
tnum = do
n <- match PTNum
pure (conv n)
where
conv : Nat -> Term n
conv 0 = TZero
conv (S n) = TSuc (conv n)
2022-07-23 03:38:15 +02:00
2022-07-26 23:07:13 +02:00
tsuc : List String -> (n : Nat) -> Vect n String -> Grammar () PiToken True (Term n)
tsuc defs n env = do
2022-07-23 03:38:15 +02:00
match PTSuc
commit
2022-07-26 23:07:13 +02:00
TSuc <$> term defs n env
2022-07-23 03:38:15 +02:00
2022-07-26 23:07:13 +02:00
tnatind : List String -> (n : Nat) -> Vect n String -> Grammar () PiToken True (Term n)
tnatind defs n env = do
2022-07-23 03:38:15 +02:00
match PTNatInd
commit
2022-07-26 23:07:13 +02:00
c <- term defs n env
z <- term defs n env
s <- term defs n env
2022-07-23 03:38:15 +02:00
pure (TNatInd c z s)
2022-07-26 23:07:13 +02:00
tsigma : List String -> (n : Nat) -> Vect n String -> Grammar () PiToken True (Term n)
tsigma defs n env = do
2022-07-23 03:38:15 +02:00
match PTSigma
commit
match PTLParen
arg <- match PTIdentifier
match PTColon
2022-07-26 23:07:13 +02:00
a <- expr defs n env
2022-07-23 03:38:15 +02:00
match PTRParen
2022-07-26 23:07:13 +02:00
b <- expr defs (S n) (arg :: env)
2022-07-23 03:38:15 +02:00
pure (TSigma a (TLam b))
2022-07-26 23:07:13 +02:00
tpair : List String -> (n : Nat) -> Vect n String -> Grammar () PiToken True (Term n)
tpair defs n env = do
2022-07-23 03:38:15 +02:00
match PTLParen
2022-07-26 23:07:13 +02:00
x <- expr defs n env
2022-07-23 03:38:15 +02:00
match PTComma
commit
2022-07-26 23:07:13 +02:00
y <- expr defs n env
2022-07-23 03:38:15 +02:00
match PTRParen
pure (TPair x y)
2022-07-26 23:07:13 +02:00
tsigind : List String -> (n : Nat) -> Vect n String -> Grammar () PiToken True (Term n)
tsigind defs n env = do
2022-07-23 03:38:15 +02:00
match PTSigInd
commit
2022-07-26 23:07:13 +02:00
a <- term defs n env
b <- term defs n env
c <- term defs n env
f <- term defs n env
2022-07-23 03:38:15 +02:00
pure (TSigInd a b c f)
2022-07-26 23:07:13 +02:00
tid : List String -> (n : Nat) -> Vect n String -> Grammar () PiToken True (Term n)
tid defs n env = do
match PTId
commit
2022-07-26 23:07:13 +02:00
ty <- term defs n env
a <- term defs n env
b <- term defs n env
pure (TId ty a b)
2022-07-26 23:07:13 +02:00
trefl : List String -> (n : Nat) -> Vect n String -> Grammar () PiToken True (Term n)
trefl defs n env = do
match PTRefl
commit
2022-07-26 23:07:13 +02:00
ty <- term defs n env
tr <- term defs n env
pure (TRefl ty tr)
2022-07-26 23:07:13 +02:00
tj : List String -> (n : Nat) -> Vect n String -> Grammar () PiToken True (Term n)
tj defs n env = do
match PTJ
commit
2022-07-26 23:07:13 +02:00
ty <- term defs n env
a <- term defs n env
b <- term defs n env
c <- term defs n env
d <- term defs n env
pure (TJ ty a b c d)
2022-07-26 23:07:13 +02:00
tlet : List String -> (n : Nat) -> Vect n String -> Grammar () PiToken True (Term n)
tlet defs n env = do
2022-07-23 03:38:15 +02:00
match PTLet
commit
arg <- match PTIdentifier
match PTColon
2022-07-26 23:07:13 +02:00
ty <- expr defs n env
2022-07-23 03:38:15 +02:00
match PTDefEq
2022-07-26 23:07:13 +02:00
tr <- expr defs n env
2022-07-23 03:38:15 +02:00
match PTIn
2022-07-26 23:07:13 +02:00
tri <- expr defs (S n) (arg :: env)
2022-07-23 03:38:15 +02:00
pure (TLet ty tr tri)
2022-07-26 23:07:13 +02:00
tlam : List String -> (n : Nat) -> Vect n String -> Grammar () PiToken True (Term n)
tlam defs n env = do
2022-07-23 03:38:15 +02:00
match PTLambda
commit
arg <- match PTIdentifier
match PTDot
2022-07-26 23:07:13 +02:00
e <- expr defs (S n) (arg :: env)
2022-07-23 03:38:15 +02:00
pure (TLam e)
2022-07-26 23:07:13 +02:00
tpi : List String -> (n : Nat) -> Vect n String -> Grammar () PiToken True (Term n)
tpi defs n env = do
2022-07-23 03:38:15 +02:00
match PTPi
commit
match PTLParen
arg <- match PTIdentifier
match PTColon
2022-07-26 23:07:13 +02:00
a <- expr defs n env
2022-07-23 03:38:15 +02:00
match PTRParen
2022-07-26 23:07:13 +02:00
b <- expr defs (S n) (arg :: env)
2022-07-23 03:38:15 +02:00
pure (TPi a b)
2022-07-26 23:07:13 +02:00
tarr : List String -> (n : Nat) -> Vect n String -> Grammar () PiToken True (Term n)
tarr defs n env = do
l <- expr1 defs n env
match PTArrow
commit
2022-07-26 23:07:13 +02:00
r <- expr defs (S n) ("" :: env)
pure (TPi l r)
2022-07-26 23:07:13 +02:00
tapp : List String -> (n : Nat) -> Vect n String -> Term n -> Grammar () PiToken True (Term n)
tapp defs n env e1 = do
e2 <- term defs n env
tapp1 defs n env (TApp e1 e2)
2022-07-23 03:38:15 +02:00
2022-07-26 23:07:13 +02:00
tapp1 : List String -> (n : Nat) -> Vect n String -> Term n -> Grammar () PiToken False (Term n)
tapp1 defs n env e = tapp defs n env e <|> pure e
2022-07-23 03:38:15 +02:00
2022-07-26 23:07:13 +02:00
tvar : List String -> (n : Nat) -> Vect n String -> Grammar () PiToken True (Term n)
tvar defs n env = do
2022-07-23 03:38:15 +02:00
str <- match PTIdentifier
2022-07-26 07:57:44 +02:00
commit
2022-07-26 23:07:13 +02:00
fromMaybe (fromMaybe (fail ("'" ++ str ++ "' not idefs n env"))
(pure . TDef . cast . finToInteger <$> findIndex (== str) defs))
(pure . TVar <$> findIndex (== str) env)
2022-07-23 03:38:15 +02:00
2022-07-26 23:07:13 +02:00
paren : List String -> (n : Nat) -> Vect n String -> Grammar () PiToken True (Term n)
paren defs n env = do
2022-07-23 03:38:15 +02:00
match PTLParen
commit
2022-07-26 23:07:13 +02:00
e <- expr defs n env
2022-07-23 03:38:15 +02:00
match PTRParen
pure e
2022-07-26 23:07:13 +02:00
definitions : List String -> Grammar () PiToken True (List String, List (Term 0, Term 0))
definitions defs = do
2022-07-25 01:21:38 +02:00
match PTLet
commit
arg <- match PTIdentifier
match PTColon
2022-07-26 23:07:13 +02:00
ty <- expr defs 0 []
2022-07-25 01:21:38 +02:00
match PTDefEq
2022-07-26 23:07:13 +02:00
tr <- expr defs 0 []
next <- definitions (arg :: defs) <|> pure ([], [])
pure (arg :: fst next, (ty, tr) :: snd next)
2022-07-25 01:21:38 +02:00
2022-07-26 23:07:13 +02:00
parsePi : List String -> Grammar () PiToken True a -> List (WithBounds PiToken) -> Either String a
parsePi defs parseEntry toks =
case parse parseEntry $ filter (not . ignored) toks of
2022-07-23 03:38:15 +02:00
Right (l, []) => Right l
Right (_, l) => Left ("contains tokens that were not consumed: " ++ show l)
2022-07-23 03:38:15 +02:00
Left e => (Left . show . map getErr) e
where
getErr : ParsingError tok -> String
2022-07-26 07:57:44 +02:00
getErr (Error "Unrecognised input" _) = ""
2022-07-23 03:38:15 +02:00
getErr (Error s _) = s
public export
2022-07-26 23:07:13 +02:00
parse : List String -> Grammar () PiToken True a -> String -> Either String a
parse defs parseEntry x =
2022-07-23 03:38:15 +02:00
case lexPi x of
2022-07-26 23:07:13 +02:00
Just toks => parsePi defs parseEntry toks
2022-07-23 03:38:15 +02:00
Nothing => Left "Failed to lex."
public export
parse0 : String -> Either String (Term 0)
2022-07-26 23:07:13 +02:00
parse0 = parse [] (expr [] 0 [])
public export
2022-07-26 23:07:13 +02:00
parseEnv : List String -> String -> Either String (Term 0)
parseEnv defs = parse defs (expr defs 0 [])
public export
2022-07-26 23:07:13 +02:00
toplevel : String -> Either String (List String, List (Term 0, Term 0))
toplevel = parse [] (definitions [])