quox/lib/Quox/Parser/FromParser.idr

374 lines
10 KiB
Idris
Raw Normal View History

2023-05-01 21:06:25 -04:00
||| take freshly-parsed input, scope check, type check, add to env
module Quox.Parser.FromParser
import Quox.Parser.Syntax
2023-03-13 14:33:09 -04:00
import Quox.Parser.Parser
import Quox.Typechecker
import Data.List
2023-05-14 13:58:46 -04:00
import Data.Maybe
import Data.SnocVect
2023-03-31 13:23:30 -04:00
import Quox.EffExtra
2023-03-13 14:33:09 -04:00
import System.File
import System.Path
2023-03-31 13:23:30 -04:00
import Data.IORef
2023-03-13 14:33:09 -04:00
2023-03-31 17:43:25 -04:00
import public Quox.Parser.FromParser.Error as Quox.Parser.FromParser
2023-03-31 13:23:30 -04:00
%default total
%hide Typing.Error
%hide Lexer.Error
%hide Parser.Error
public export
NDefinition : Type
2023-04-01 13:16:43 -04:00
NDefinition = (Name, Definition)
public export
IncludePath : Type
2023-03-31 13:23:30 -04:00
IncludePath = List String
2023-03-13 14:33:09 -04:00
public export
SeenFiles : Type
2023-03-31 13:23:30 -04:00
SeenFiles = SortedSet String
2023-03-13 14:33:09 -04:00
public export
data StateTag = NS | SEEN
public export
FromParserPure : List (Type -> Type)
FromParserPure =
2023-05-01 21:06:25 -04:00
[Except Error, DefsState, StateL NS Mods, NameGen]
public export
2023-05-01 21:06:25 -04:00
LoadFile' : List (Type -> Type)
LoadFile' = [IO, StateL SEEN SeenFiles, Reader IncludePath]
2023-03-31 13:23:30 -04:00
public export
2023-05-01 21:06:25 -04:00
LoadFile : List (Type -> Type)
LoadFile = LoadFile' ++ [Except Error]
2023-05-01 21:06:25 -04:00
public export
FromParserIO : List (Type -> Type)
FromParserIO = FromParserPure ++ LoadFile'
parameters {auto _ : Functor m} (b : Var n -> m a) (f : PName -> m a)
(xs : Context' PatVar n)
private
fromBaseName : PBaseName -> m a
fromBaseName x = maybe (f $ MakePName [<] x) b $
Context.find (\y => y.name == Just x) xs
private
fromName : PName -> m a
fromName x = if null x.mods then fromBaseName x.base else f x
export
2023-03-31 13:23:30 -04:00
fromPDimWith : Has (Except Error) fs =>
Context' PatVar d -> PDim -> Eff fs (Dim d)
2023-05-01 21:06:25 -04:00
fromPDimWith ds (K e loc) = pure $ K e loc
fromPDimWith ds (V i loc) =
fromBaseName (\i => pure $ B i loc)
(const $ throw $ DimNotInScope loc i) ds i
private
2023-03-31 13:23:30 -04:00
avoidDim : Has (Except Error) fs =>
2023-05-01 21:06:25 -04:00
Context' PatVar d -> Loc -> PName -> Eff fs Name
avoidDim ds loc x =
fromName (const $ throw $ DimNameInTerm loc x.base) (pure . fromPName) ds x
private
2023-05-01 21:06:25 -04:00
resolveName : Mods -> Loc -> Name -> Eff FromParserPure (Term d n)
resolveName ns loc x =
let here = addMods ns x in
if isJust $ lookup here !(getAt DEFS) then
2023-05-01 21:06:25 -04:00
pure $ FT here loc
else do
let ns :< _ = ns
2023-05-01 21:06:25 -04:00
| _ => throw $ TermNotInScope loc x
resolveName ns loc x
export
2023-05-01 21:06:25 -04:00
fromPatVar : PatVar -> BindName
fromPatVar (Unused loc) = BN Unused loc
fromPatVar (PV x loc) = BN (UN x) loc
export
fromPQty : PQty -> Qty
fromPQty (PQ q _) = q
export
fromPTagVal : PTagVal -> TagVal
fromPTagVal (PT t _) = t
mutual
export
fromPTermWith : Context' PatVar d -> Context' PatVar n ->
PTerm -> Eff FromParserPure (Term d n)
fromPTermWith ds ns t0 = case t0 of
2023-05-01 21:06:25 -04:00
TYPE k loc =>
pure $ TYPE k loc
2023-05-01 21:06:25 -04:00
Pi pi x s t loc =>
Pi (fromPQty pi)
<$> fromPTermWith ds ns s
<*> fromPTermTScope ds ns [< x] t
2023-05-01 21:06:25 -04:00
<*> pure loc
2023-05-01 21:06:25 -04:00
Lam x s loc =>
Lam <$> fromPTermTScope ds ns [< x] s <*> pure loc
2023-05-01 21:06:25 -04:00
App s t loc =>
map E $ App
<$> fromPTermElim ds ns s
<*> fromPTermWith ds ns t
<*> pure loc
2023-05-01 21:06:25 -04:00
Sig x s t loc =>
Sig <$> fromPTermWith ds ns s
<*> fromPTermTScope ds ns [< x] t
<*> pure loc
2023-05-01 21:06:25 -04:00
Pair s t loc =>
Pair <$> fromPTermWith ds ns s <*> fromPTermWith ds ns t <*> pure loc
2023-05-01 21:06:25 -04:00
Case pi pair (r, ret) (CasePair (x, y) body _) loc =>
map E $ CasePair (fromPQty pi)
<$> fromPTermElim ds ns pair
<*> fromPTermTScope ds ns [< r] ret
<*> fromPTermTScope ds ns [< x, y] body
2023-05-01 21:06:25 -04:00
<*> pure loc
2023-05-01 21:06:25 -04:00
Case pi tag (r, ret) (CaseEnum arms _) loc =>
map E $ CaseEnum (fromPQty pi)
<$> fromPTermElim ds ns tag
<*> fromPTermTScope ds ns [< r] ret
<*> assert_total fromPTermEnumArms ds ns arms
2023-05-01 21:06:25 -04:00
<*> pure loc
2023-05-01 21:06:25 -04:00
Nat loc => pure $ Nat loc
Zero loc => pure $ Zero loc
Succ n loc => [|Succ (fromPTermWith ds ns n) (pure loc)|]
2023-03-26 08:40:54 -04:00
2023-05-01 21:06:25 -04:00
Case pi nat (r, ret) (CaseNat zer (s, pi', ih, suc) _) loc =>
map E $ CaseNat (fromPQty pi) (fromPQty pi')
2023-03-26 08:40:54 -04:00
<$> fromPTermElim ds ns nat
<*> fromPTermTScope ds ns [< r] ret
<*> fromPTermWith ds ns zer
<*> fromPTermTScope ds ns [< s, ih] suc
2023-05-01 21:06:25 -04:00
<*> pure loc
2023-03-26 08:40:54 -04:00
2023-05-01 21:06:25 -04:00
Enum strs loc =>
let set = SortedSet.fromList strs in
if length strs == length (SortedSet.toList set) then
2023-05-01 21:06:25 -04:00
pure $ Enum set loc
else
2023-05-01 21:06:25 -04:00
throw $ DuplicatesInEnum loc strs
2023-05-01 21:06:25 -04:00
Tag str loc => pure $ Tag str loc
2023-05-01 21:06:25 -04:00
Eq (i, ty) s t loc =>
Eq <$> fromPTermDScope ds ns [< i] ty
<*> fromPTermWith ds ns s
<*> fromPTermWith ds ns t
2023-05-01 21:06:25 -04:00
<*> pure loc
2023-05-01 21:06:25 -04:00
DLam i s loc =>
DLam <$> fromPTermDScope ds ns [< i] s <*> pure loc
2023-05-01 21:06:25 -04:00
DApp s p loc =>
map E $ DApp
<$> fromPTermElim ds ns s
<*> fromPDimWith ds p
<*> pure loc
2023-05-01 21:06:25 -04:00
BOX q ty loc => BOX (fromPQty q) <$> fromPTermWith ds ns ty <*> pure loc
2023-03-31 13:11:35 -04:00
2023-05-01 21:06:25 -04:00
Box val loc => Box <$> fromPTermWith ds ns val <*> pure loc
2023-03-31 13:11:35 -04:00
2023-05-01 21:06:25 -04:00
Case pi box (r, ret) (CaseBox b body _) loc =>
map E $ CaseBox (fromPQty pi)
2023-03-31 13:11:35 -04:00
<$> fromPTermElim ds ns box
<*> fromPTermTScope ds ns [< r] ret
<*> fromPTermTScope ds ns [< b] body
2023-05-01 21:06:25 -04:00
<*> pure loc
2023-03-31 13:11:35 -04:00
2023-05-01 21:06:25 -04:00
V x loc =>
fromName (\i => pure $ E $ B i loc)
(resolveName !(getAt NS) loc <=< avoidDim ds loc) ns x
2023-05-01 21:06:25 -04:00
Ann s a loc =>
map E $ Ann
<$> fromPTermWith ds ns s
<*> fromPTermWith ds ns a
<*> pure loc
2023-05-01 21:06:25 -04:00
Coe (i, ty) p q val loc =>
2023-04-15 09:13:01 -04:00
map E $ Coe
<$> fromPTermDScope ds ns [< i] ty
<*> fromPDimWith ds p
<*> fromPDimWith ds q
<*> fromPTermWith ds ns val
2023-05-01 21:06:25 -04:00
<*> pure loc
2023-04-15 09:13:01 -04:00
2023-05-01 21:06:25 -04:00
Comp (i, ty) p q val r (j0, val0) (j1, val1) loc =>
map E $ CompH'
2023-04-15 09:13:01 -04:00
<$> fromPTermDScope ds ns [< i] ty
<*> fromPDimWith ds p
<*> fromPDimWith ds q
<*> fromPTermWith ds ns val
<*> fromPDimWith ds r
<*> fromPTermDScope ds ns [< j0] val0
<*> fromPTermDScope ds ns [< j1] val1
2023-05-01 21:06:25 -04:00
<*> pure loc
2023-04-15 09:13:01 -04:00
private
fromPTermEnumArms : Context' PatVar d -> Context' PatVar n ->
List (PTagVal, PTerm) ->
Eff FromParserPure (CaseEnumArms d n)
fromPTermEnumArms ds ns =
map SortedMap.fromList .
traverse (bitraverse (pure . fromPTagVal) (fromPTermWith ds ns))
private
fromPTermElim : Context' PatVar d -> Context' PatVar n ->
PTerm -> Eff FromParserPure (Elim d n)
fromPTermElim ds ns e =
case !(fromPTermWith ds ns e) of
E e => pure e
t => let ctx = MkNameContexts (map fromPatVar ds) (map fromPatVar ns) in
2023-05-01 21:06:25 -04:00
throw $ AnnotationNeeded t.loc ctx t
-- [todo] use SN if the var is named but still unused
private
fromPTermTScope : {s : Nat} -> Context' PatVar d -> Context' PatVar n ->
SnocVect s PatVar -> PTerm ->
Eff FromParserPure (ScopeTermN s d n)
fromPTermTScope ds ns xs t =
if all isUnused xs then
SN <$> fromPTermWith ds ns t
else
ST (fromSnocVect $ map fromPatVar xs) <$> fromPTermWith ds (ns ++ xs) t
private
fromPTermDScope : {s : Nat} -> Context' PatVar d -> Context' PatVar n ->
SnocVect s PatVar -> PTerm ->
Eff FromParserPure (DScopeTermN s d n)
fromPTermDScope ds ns xs t =
if all isUnused xs then
SN <$> fromPTermWith ds ns t
else
DST (fromSnocVect $ map fromPatVar xs) <$> fromPTermWith (ds ++ xs) ns t
export %inline
fromPTerm : PTerm -> Eff FromParserPure (Term 0 0)
fromPTerm = fromPTermWith [<] [<]
export
2023-05-01 21:06:25 -04:00
globalPQty : Loc -> (q : Qty) -> Eff [Except Error] (So $ isGlobal q)
globalPQty loc pi = case choose $ isGlobal pi of
2023-04-01 13:16:43 -04:00
Left y => pure y
2023-05-01 21:06:25 -04:00
Right _ => throw $ QtyNotGlobal loc pi
2023-03-13 14:33:09 -04:00
export
fromPBaseNameNS : PBaseName -> Eff [StateL NS Mods] Name
fromPBaseNameNS name = pure $ addMods !(getAt NS) $ fromPBaseName name
2023-03-13 14:33:09 -04:00
private
2023-05-01 21:06:25 -04:00
liftTC : TC a -> Eff FromParserPure a
liftTC act = do
res <- lift $ runExcept $ runReaderAt DEFS !(getAt DEFS) act
rethrow $ mapFst WrapTypeError res
2023-03-31 13:23:30 -04:00
export covering
fromPDef : PDefinition -> Eff FromParserPure NDefinition
2023-05-01 21:06:25 -04:00
fromPDef (MkPDef qty pname ptype pterm defLoc) = do
name <- lift $ fromPBaseNameNS pname
2023-05-01 21:06:25 -04:00
qtyGlobal <- lift $ globalPQty qty.loc qty.val
let gqty = Element qty.val qtyGlobal
sqty = globalToSubj gqty
type <- lift $ traverse fromPTerm ptype
term <- lift $ fromPTerm pterm
2023-03-13 14:33:09 -04:00
case type of
Just type => do
2023-05-01 21:06:25 -04:00
liftTC $ checkTypeC empty type Nothing
liftTC $ ignore $ checkC empty sqty term type
let def = mkDef gqty type term defLoc
2023-03-31 13:23:30 -04:00
modifyAt DEFS $ insert name def
pure (name, def)
2023-03-13 14:33:09 -04:00
Nothing => do
2023-05-01 21:06:25 -04:00
let E elim = term | _ => throw $ AnnotationNeeded term.loc empty term
res <- liftTC $ inferC empty sqty elim
let def = mkDef gqty res.type term defLoc
2023-03-31 13:23:30 -04:00
modifyAt DEFS $ insert name def
pure (name, def)
export covering
fromPDecl : PDecl -> Eff FromParserPure (List NDefinition)
2023-03-31 13:23:30 -04:00
fromPDecl (PDef def) = singleton <$> fromPDef def
2023-03-13 14:33:09 -04:00
fromPDecl (PNs ns) =
2023-03-31 13:23:30 -04:00
localAt NS (<+> ns.name) $ concat <$> traverse fromPDecl ns.decls
2023-03-13 14:33:09 -04:00
2023-03-31 13:23:30 -04:00
export covering
2023-05-01 21:06:25 -04:00
loadFile : Loc -> String -> Eff LoadFile (Maybe String)
loadFile loc file =
2023-03-31 13:23:30 -04:00
if contains file !(getAt SEEN) then
2023-03-13 14:33:09 -04:00
pure Nothing
else do
2023-04-17 17:59:02 -04:00
Just ifile <- firstExists (map (</> file) !ask)
2023-05-01 21:06:25 -04:00
| Nothing => throw $ LoadError loc file FileNotFound
2023-04-17 17:59:02 -04:00
case !(readFile ifile) of
2023-03-31 13:23:30 -04:00
Right res => modifyAt SEEN (insert file) $> Just res
2023-05-01 21:06:25 -04:00
Left err => throw $ LoadError loc ifile err
2023-03-31 13:23:30 -04:00
mutual
2023-03-31 13:23:30 -04:00
export covering
2023-05-01 21:06:25 -04:00
loadProcessFile : Loc -> String -> Eff FromParserIO (List NDefinition)
loadProcessFile loc file =
case !(lift $ loadFile loc file) of
2023-03-31 13:23:30 -04:00
Just inp => do
tl <- either (throw . WrapParseError file) pure $ lexParseInput file inp
2023-03-31 13:23:30 -04:00
concat <$> traverse fromPTopLevel tl
Nothing => pure []
||| populates the `defs` field of the state
export covering
2023-05-01 21:06:25 -04:00
fromPTopLevel : PTopLevel -> Eff FromParserIO (List NDefinition)
fromPTopLevel (PD decl) = lift $ fromPDecl decl
fromPTopLevel (PLoad file loc) = loadProcessFile loc file
2023-03-13 14:33:09 -04:00
export
2023-05-01 21:06:25 -04:00
fromParserPure : NameSuf -> Definitions ->
Eff FromParserPure a ->
2023-05-01 21:06:25 -04:00
(Either Error (a, Definitions), NameSuf)
fromParserPure suf defs act =
extract $
2023-05-01 21:06:25 -04:00
runStateAt GEN suf $
runExcept $
evalStateAt NS [<] $
runStateAt DEFS defs act
2023-03-13 14:33:09 -04:00
export
2023-03-31 13:23:30 -04:00
fromParserIO : (MonadRec io, HasIO io) =>
2023-05-01 21:06:25 -04:00
IncludePath ->
IORef SeenFiles -> IORef NameSuf -> IORef Definitions ->
Eff FromParserIO a -> io (Either Error a)
fromParserIO inc seen suf defs act =
2023-03-31 13:23:30 -04:00
runIO $
2023-05-01 21:06:25 -04:00
runStateIORefAt GEN suf $
2023-03-31 13:23:30 -04:00
runExcept $
evalStateAt NS [<] $
runStateIORefAt SEEN seen $
runStateIORefAt DEFS defs $
runReader inc act