Idris2/libs/base/Language/Reflection/TT.idr
Ruslan Feizerakhmanov 7aee7c9b7c
[ new ] --install-with-src; refactoring around FCs (#1450)
Why:

* To implement robust cross-project go-to-definition in LSP
  i.e you can jump to definition of any global name coming
  from library dependencies, as well as from the local project files.

What it does:

*  Modify `FC`s to carry `ModuleIdent` for .idr sources,
   file name for .ipkg sources or nothing for interactive runs.

*  Add `--install-with-src` to install the source code alongside
   the ttc binaries. The source is installed into the same directory
   as the corresponding ttc file. Installed sources are made read-only.

*  As we install the sources pinned to the related ttc files we gain
   the versioning of sources for free.
2021-06-05 12:53:22 +01:00

149 lines
4.1 KiB
Idris

module Language.Reflection.TT
import public Data.List
%default total
public export
data Namespace = MkNS (List String) -- namespace, stored in reverse order
public export
data ModuleIdent = MkMI (List String) -- module identifier, stored in reverse order
export
showSep : String -> List String -> String
showSep sep [] = ""
showSep sep [x] = x
showSep sep (x :: xs) = x ++ sep ++ showSep sep xs
export
Show Namespace where
show (MkNS ns) = showSep "." (reverse ns)
-- 'FilePos' represents the position of
-- the source information in the file (or REPL).
-- in the form of '(line-no, column-no)'.
public export
FilePos : Type
FilePos = (Int, Int)
public export
data VirtualIdent : Type where
Interactive : VirtualIdent
public export
data OriginDesc : Type where
||| Anything that originates in physical Idris source files is assigned a
||| `PhysicalIdrSrc modIdent`,
||| where `modIdent` is the top-level module identifier of that file.
PhysicalIdrSrc : (ident : ModuleIdent) -> OriginDesc
||| Anything parsed from a package file is decorated with `PhysicalPkgSrc fname`,
||| where `fname` is path to the package file.
PhysicalPkgSrc : (fname : String) -> OriginDesc
Virtual : (ident : VirtualIdent) -> OriginDesc
||| A file context is a filename together with starting and ending positions.
||| It's often carried by AST nodes that might have been created from a source
||| file or by the compiler. That makes it useful to have the notion of
||| `EmptyFC` as part of the type.
public export
data FC = MkFC OriginDesc FilePos FilePos
| ||| Virtual FCs are FC attached to desugared/generated code. They can help with marking
||| errors, but we shouldn't attach semantic highlighting metadata to them.
MkVirtualFC OriginDesc FilePos FilePos
| EmptyFC
public export
emptyFC : FC
emptyFC = EmptyFC
public export
data NameType : Type where
Bound : NameType
Func : NameType
DataCon : (tag : Int) -> (arity : Nat) -> NameType
TyCon : (tag : Int) -> (arity : Nat) -> NameType
public export
data Constant
= I Int
| BI Integer
| B8 Int
| B16 Int
| B32 Int
| B64 Integer
| Str String
| Ch Char
| Db Double
| WorldVal
| IntType
| IntegerType
| Bits8Type
| Bits16Type
| Bits32Type
| Bits64Type
| StringType
| CharType
| DoubleType
| WorldType
public export
data Name = UN String -- user defined name
| MN String Int -- machine generated name
| NS Namespace Name -- name in a namespace
| DN String Name -- a name and how to display it
| RF String -- record field name
export
Show Name where
show (NS ns n) = show ns ++ "." ++ show n
show (UN x) = x
show (MN x y) = "{" ++ x ++ ":" ++ show y ++ "}"
show (DN str y) = str
show (RF n) = "." ++ n
public export
data Count = M0 | M1 | MW
public export
data PiInfo t = ImplicitArg | ExplicitArg | AutoImplicit | DefImplicit t
public export
data IsVar : Name -> Nat -> List Name -> Type where
First : IsVar n Z (n :: ns)
Later : IsVar n i ns -> IsVar n (S i) (m :: ns)
public export
data LazyReason = LInf | LLazy | LUnknown
export
data TT : Type where [external]
{-
-- Type checked terms in the core TT
public export
data TT : List Name -> Type where
Local : FC -> (idx : Nat) -> (0 prf : IsVar name idx vars) -> TT vars
Ref : FC -> NameType -> Name -> TT vars
Pi : FC -> Count -> PiInfo (TT vars) ->
(x : Name) -> (argTy : TT vars) -> (retTy : TT (x :: vars)) ->
TT vars
Lam : FC -> Count -> PiInfo (TT vars) ->
(x : Name) -> (argTy : TT vars) -> (scope : TT (x :: vars)) ->
TT vars
App : FC -> TT vars -> TT vars -> TT vars
TDelayed : FC -> LazyReason -> TT vars -> TT vars
TDelay : FC -> LazyReason -> (ty : TT vars) -> (arg : TT vars) -> TT vars
TForce : FC -> LazyReason -> TT vars -> TT vars
PrimVal : FC -> Constant -> TT vars
Erased : FC -> TT vars
TType : FC -> TT vars
-}
public export
data TotalReq = Total | CoveringOnly | PartialOK
public export
data Visibility = Private | Export | Public