382 lines
11 KiB
Idris
382 lines
11 KiB
Idris
||| take freshly-parsed input, scope check, type check, add to env
|
|
module Quox.Parser.FromParser
|
|
|
|
import Quox.Parser.Syntax
|
|
import Quox.Parser.Parser
|
|
import Quox.Typechecker
|
|
|
|
import Data.List
|
|
import Data.Maybe
|
|
import Data.SnocVect
|
|
import Quox.EffExtra
|
|
|
|
import System.File
|
|
import System.Path
|
|
import Data.IORef
|
|
|
|
import public Quox.Parser.FromParser.Error as Quox.Parser.FromParser
|
|
|
|
%default total
|
|
|
|
%hide Typing.Error
|
|
%hide Lexer.Error
|
|
%hide Parser.Error
|
|
|
|
|
|
public export
|
|
NDefinition : Type
|
|
NDefinition = (Name, Definition)
|
|
|
|
public export
|
|
IncludePath : Type
|
|
IncludePath = List String
|
|
|
|
public export
|
|
SeenFiles : Type
|
|
SeenFiles = SortedSet String
|
|
|
|
|
|
public export
|
|
data StateTag = NS | SEEN
|
|
|
|
public export
|
|
FromParserPure : List (Type -> Type)
|
|
FromParserPure =
|
|
[Except Error, DefsState, StateL NS Mods, NameGen]
|
|
|
|
public export
|
|
LoadFile' : List (Type -> Type)
|
|
LoadFile' = [IO, StateL SEEN SeenFiles, Reader IncludePath]
|
|
|
|
public export
|
|
LoadFile : List (Type -> Type)
|
|
LoadFile = LoadFile' ++ [Except Error]
|
|
|
|
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
|
|
fromPDimWith : Has (Except Error) fs =>
|
|
Context' PatVar d -> PDim -> Eff fs (Dim d)
|
|
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
|
|
avoidDim : Has (Except Error) fs =>
|
|
Context' PatVar d -> Loc -> PName -> Eff fs Name
|
|
avoidDim ds loc x =
|
|
fromName (const $ throw $ DimNameInTerm loc x.base) (pure . fromPName) ds x
|
|
|
|
private
|
|
resolveName : Mods -> Loc -> Name -> Maybe Universe ->
|
|
Eff FromParserPure (Term d n)
|
|
resolveName ns loc x u =
|
|
let here = addMods ns x in
|
|
if isJust $ lookup here !(getAt DEFS) then
|
|
pure $ FT here (fromMaybe 0 u) loc
|
|
else do
|
|
let ns :< _ = ns
|
|
| _ => throw $ TermNotInScope loc x
|
|
resolveName ns loc x u
|
|
|
|
export
|
|
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
|
|
|
|
|
|
private
|
|
fromV : Context' PatVar d -> Context' PatVar n ->
|
|
PName -> Maybe Universe -> Loc -> Eff FromParserPure (Term d n)
|
|
fromV ds ns x u loc = fromName bound free ns x where
|
|
bound : Var n -> Eff FromParserPure (Term d n)
|
|
bound i = do whenJust u $ \u => throw $ DisplacedBoundVar loc x
|
|
pure $ E $ B i loc
|
|
free : PName -> Eff FromParserPure (Term d n)
|
|
free x = do x <- avoidDim ds loc x
|
|
resolveName !(getAt NS) loc x u
|
|
|
|
mutual
|
|
export
|
|
fromPTermWith : Context' PatVar d -> Context' PatVar n ->
|
|
PTerm -> Eff FromParserPure (Term d n)
|
|
fromPTermWith ds ns t0 = case t0 of
|
|
TYPE k loc =>
|
|
pure $ TYPE k loc
|
|
|
|
Pi pi x s t loc =>
|
|
Pi (fromPQty pi)
|
|
<$> fromPTermWith ds ns s
|
|
<*> fromPTermTScope ds ns [< x] t
|
|
<*> pure loc
|
|
|
|
Lam x s loc =>
|
|
Lam <$> fromPTermTScope ds ns [< x] s <*> pure loc
|
|
|
|
App s t loc =>
|
|
map E $ App
|
|
<$> fromPTermElim ds ns s
|
|
<*> fromPTermWith ds ns t
|
|
<*> pure loc
|
|
|
|
Sig x s t loc =>
|
|
Sig <$> fromPTermWith ds ns s
|
|
<*> fromPTermTScope ds ns [< x] t
|
|
<*> pure loc
|
|
|
|
Pair s t loc =>
|
|
Pair <$> fromPTermWith ds ns s <*> fromPTermWith ds ns t <*> pure loc
|
|
|
|
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
|
|
<*> pure loc
|
|
|
|
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
|
|
<*> pure loc
|
|
|
|
Nat loc => pure $ Nat loc
|
|
Zero loc => pure $ Zero loc
|
|
Succ n loc => [|Succ (fromPTermWith ds ns n) (pure loc)|]
|
|
|
|
Case pi nat (r, ret) (CaseNat zer (s, pi', ih, suc) _) loc =>
|
|
map E $ CaseNat (fromPQty pi) (fromPQty pi')
|
|
<$> fromPTermElim ds ns nat
|
|
<*> fromPTermTScope ds ns [< r] ret
|
|
<*> fromPTermWith ds ns zer
|
|
<*> fromPTermTScope ds ns [< s, ih] suc
|
|
<*> pure loc
|
|
|
|
Enum strs loc =>
|
|
let set = SortedSet.fromList strs in
|
|
if length strs == length (SortedSet.toList set) then
|
|
pure $ Enum set loc
|
|
else
|
|
throw $ DuplicatesInEnum loc strs
|
|
|
|
Tag str loc => pure $ Tag str loc
|
|
|
|
Eq (i, ty) s t loc =>
|
|
Eq <$> fromPTermDScope ds ns [< i] ty
|
|
<*> fromPTermWith ds ns s
|
|
<*> fromPTermWith ds ns t
|
|
<*> pure loc
|
|
|
|
DLam i s loc =>
|
|
DLam <$> fromPTermDScope ds ns [< i] s <*> pure loc
|
|
|
|
DApp s p loc =>
|
|
map E $ DApp
|
|
<$> fromPTermElim ds ns s
|
|
<*> fromPDimWith ds p
|
|
<*> pure loc
|
|
|
|
BOX q ty loc => BOX (fromPQty q) <$> fromPTermWith ds ns ty <*> pure loc
|
|
|
|
Box val loc => Box <$> fromPTermWith ds ns val <*> pure loc
|
|
|
|
Case pi box (r, ret) (CaseBox b body _) loc =>
|
|
map E $ CaseBox (fromPQty pi)
|
|
<$> fromPTermElim ds ns box
|
|
<*> fromPTermTScope ds ns [< r] ret
|
|
<*> fromPTermTScope ds ns [< b] body
|
|
<*> pure loc
|
|
|
|
V x u loc => fromV ds ns x u loc
|
|
|
|
Ann s a loc =>
|
|
map E $ Ann
|
|
<$> fromPTermWith ds ns s
|
|
<*> fromPTermWith ds ns a
|
|
<*> pure loc
|
|
|
|
Coe (i, ty) p q val loc =>
|
|
map E $ Coe
|
|
<$> fromPTermDScope ds ns [< i] ty
|
|
<*> fromPDimWith ds p
|
|
<*> fromPDimWith ds q
|
|
<*> fromPTermWith ds ns val
|
|
<*> pure loc
|
|
|
|
Comp (i, ty) p q val r (j0, val0) (j1, val1) loc =>
|
|
map E $ CompH'
|
|
<$> 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
|
|
<*> pure loc
|
|
|
|
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
|
|
throw $ AnnotationNeeded t.loc ctx t
|
|
|
|
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
|
|
globalPQty : Loc -> (q : Qty) -> Eff [Except Error] (So $ isGlobal q)
|
|
globalPQty loc pi = case choose $ isGlobal pi of
|
|
Left y => pure y
|
|
Right _ => throw $ QtyNotGlobal loc pi
|
|
|
|
|
|
export
|
|
fromPBaseNameNS : PBaseName -> Eff [StateL NS Mods] Name
|
|
fromPBaseNameNS name = pure $ addMods !(getAt NS) $ fromPBaseName name
|
|
|
|
private
|
|
liftTC : TC a -> Eff FromParserPure a
|
|
liftTC act = do
|
|
res <- lift $ runExcept $ runReaderAt DEFS !(getAt DEFS) act
|
|
rethrow $ mapFst WrapTypeError res
|
|
|
|
export covering
|
|
fromPDef : PDefinition -> Eff FromParserPure NDefinition
|
|
fromPDef (MkPDef qty pname ptype pterm defLoc) = do
|
|
name <- lift $ fromPBaseNameNS pname
|
|
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
|
|
case type of
|
|
Just type => do
|
|
liftTC $ checkTypeC empty type Nothing
|
|
liftTC $ ignore $ checkC empty sqty term type
|
|
let def = mkDef gqty type term defLoc
|
|
modifyAt DEFS $ insert name def
|
|
pure (name, def)
|
|
Nothing => do
|
|
let E elim = term | _ => throw $ AnnotationNeeded term.loc empty term
|
|
res <- liftTC $ inferC empty sqty elim
|
|
let def = mkDef gqty res.type term defLoc
|
|
modifyAt DEFS $ insert name def
|
|
pure (name, def)
|
|
|
|
export covering
|
|
fromPDecl : PDecl -> Eff FromParserPure (List NDefinition)
|
|
fromPDecl (PDef def) = singleton <$> fromPDef def
|
|
fromPDecl (PNs ns) =
|
|
localAt NS (<+> ns.name) $ concat <$> traverse fromPDecl ns.decls
|
|
|
|
|
|
export covering
|
|
loadFile : Loc -> String -> Eff LoadFile (Maybe String)
|
|
loadFile loc file =
|
|
if contains file !(getAt SEEN) then
|
|
pure Nothing
|
|
else do
|
|
Just ifile <- firstExists (map (</> file) !ask)
|
|
| Nothing => throw $ LoadError loc file FileNotFound
|
|
case !(readFile ifile) of
|
|
Right res => modifyAt SEEN (insert file) $> Just res
|
|
Left err => throw $ LoadError loc ifile err
|
|
|
|
mutual
|
|
export covering
|
|
loadProcessFile : Loc -> String -> Eff FromParserIO (List NDefinition)
|
|
loadProcessFile loc file =
|
|
case !(lift $ loadFile loc file) of
|
|
Just inp => do
|
|
tl <- either (throw . WrapParseError file) pure $ lexParseInput file inp
|
|
concat <$> traverse fromPTopLevel tl
|
|
Nothing => pure []
|
|
|
|
||| populates the `defs` field of the state
|
|
export covering
|
|
fromPTopLevel : PTopLevel -> Eff FromParserIO (List NDefinition)
|
|
fromPTopLevel (PD decl) = lift $ fromPDecl decl
|
|
fromPTopLevel (PLoad file loc) = loadProcessFile loc file
|
|
|
|
export
|
|
fromParserPure : NameSuf -> Definitions ->
|
|
Eff FromParserPure a ->
|
|
(Either Error (a, Definitions), NameSuf)
|
|
fromParserPure suf defs act =
|
|
extract $
|
|
runStateAt GEN suf $
|
|
runExcept $
|
|
evalStateAt NS [<] $
|
|
runStateAt DEFS defs act
|
|
|
|
export
|
|
fromParserIO : (MonadRec io, HasIO io) =>
|
|
IncludePath ->
|
|
IORef SeenFiles -> IORef NameSuf -> IORef Definitions ->
|
|
Eff FromParserIO a -> io (Either Error a)
|
|
fromParserIO inc seen suf defs act =
|
|
runIO $
|
|
runStateIORefAt GEN suf $
|
|
runExcept $
|
|
evalStateAt NS [<] $
|
|
runStateIORefAt SEEN seen $
|
|
runStateIORefAt DEFS defs $
|
|
runReader inc act
|