191 lines
3.9 KiB
Idris
191 lines
3.9 KiB
Idris
module Quox.Name
|
|
|
|
import Quox.Loc
|
|
import Quox.CharExtra
|
|
import Quox.PrettyValExtra
|
|
import public Data.SnocList
|
|
import Data.List
|
|
import Control.Eff
|
|
import Text.Lexer
|
|
import Derive.Prelude
|
|
|
|
%hide TT.Name
|
|
|
|
%default total
|
|
%language ElabReflection
|
|
|
|
|
|
public export
|
|
NameSuf : Type
|
|
NameSuf = Nat
|
|
|
|
public export
|
|
data BaseName
|
|
= UN String -- user-given name
|
|
| MN String NameSuf -- machine-generated name
|
|
| Unused -- "_"
|
|
%runElab derive "BaseName" [Eq, Ord, PrettyVal]
|
|
|
|
export
|
|
baseStr : BaseName -> String
|
|
baseStr (UN x) = x
|
|
baseStr (MN x i) = "\{x}#\{show i}"
|
|
baseStr Unused = "_"
|
|
|
|
export Show BaseName where show = show . baseStr
|
|
export FromString BaseName where fromString = UN
|
|
|
|
|
|
public export
|
|
Mods : Type
|
|
Mods = SnocList String
|
|
|
|
|
|
public export
|
|
record Name where
|
|
constructor MkName
|
|
mods : Mods
|
|
base : BaseName
|
|
%runElab derive "Name" [Eq, Ord]
|
|
|
|
public export %inline
|
|
unq : BaseName -> Name
|
|
unq = MkName [<]
|
|
|
|
||| add some namespaces to the beginning of a name
|
|
public export %inline
|
|
addMods : Mods -> Name -> Name
|
|
addMods ms = {mods $= (ms <+>)}
|
|
|
|
|
|
public export
|
|
PBaseName : Type
|
|
PBaseName = String
|
|
|
|
public export
|
|
record PName where
|
|
constructor MkPName
|
|
mods : Mods
|
|
base : PBaseName
|
|
%runElab derive "PName" [Eq, Ord, PrettyVal]
|
|
|
|
export %inline
|
|
fromPName : PName -> Name
|
|
fromPName p = MkName p.mods $ UN p.base
|
|
|
|
export %inline
|
|
toPName : Name -> PName
|
|
toPName p = MkPName p.mods $ baseStr p.base
|
|
|
|
export %inline
|
|
fromPBaseName : PBaseName -> Name
|
|
fromPBaseName = MkName [<] . UN
|
|
|
|
export
|
|
Show PName where
|
|
show (MkPName mods base) =
|
|
show $ concat $ intersperse "." $ toList $ mods :< base
|
|
|
|
export Show Name where show = show . toPName
|
|
|
|
export FromString PName where fromString = MkPName [<]
|
|
|
|
export FromString Name where fromString = fromPBaseName
|
|
|
|
|
|
public export
|
|
record BindName where
|
|
constructor BN
|
|
val : BaseName
|
|
loc_ : Loc
|
|
%runElab derive "BindName" [Eq, Ord, Show, PrettyVal]
|
|
|
|
export Located BindName where n.loc = n.loc_
|
|
export Relocatable BindName where setLoc loc (BN x _) = BN x loc
|
|
|
|
|
|
export
|
|
toDotsP : PName -> String
|
|
toDotsP x = fastConcat $ cast $ map (<+> ".") x.mods :< x.base
|
|
|
|
export
|
|
toDots : Name -> String
|
|
toDots x = fastConcat $ cast $ map (<+> ".") x.mods :< baseStr x.base
|
|
|
|
export
|
|
fromListP : List1 String -> PName
|
|
fromListP (x ::: xs) = go [<] x xs where
|
|
go : SnocList String -> String -> List String -> PName
|
|
go mods x [] = MkPName mods x
|
|
go mods x (y :: ys) = go (mods :< x) y ys
|
|
|
|
export %inline
|
|
fromList : List1 String -> Name
|
|
fromList = fromPName . fromListP
|
|
|
|
|
|
export
|
|
syntaxChars : List Char
|
|
syntaxChars = ['(', ')', '[', ']', '{', '}', '"', '\'', ',', '.', ';', '^']
|
|
|
|
export
|
|
isSymStart, isSymCont : Char -> Bool
|
|
isSymStart c = not (c `elem` syntaxChars) && isSymChar c
|
|
isSymCont c = c == '\'' || isSymStart c
|
|
|
|
export
|
|
idStart, idCont, idEnd, idContEnd : Lexer
|
|
idStart = pred isIdStart
|
|
idCont = pred isIdCont
|
|
idEnd = pred $ \c => c `elem` unpack "?!#"
|
|
idContEnd = idCont <|> idEnd
|
|
|
|
export
|
|
symStart, symCont : Lexer
|
|
symStart = pred isSymStart
|
|
symCont = pred isSymCont
|
|
|
|
export
|
|
baseName : Lexer
|
|
baseName = idStart <+> many idCont <+> many idEnd
|
|
<|> symStart <+> many symCont
|
|
|
|
export
|
|
name : Lexer
|
|
name = baseName <+> many (is '.' <+> baseName)
|
|
|
|
|
|
export
|
|
isName : String -> Bool
|
|
isName str =
|
|
case scan name [] (unpack str) of
|
|
Just (_, []) => True
|
|
_ => False
|
|
|
|
|
|
public export
|
|
data GenTag = GEN
|
|
|
|
public export
|
|
NameGen : Type -> Type
|
|
NameGen = StateL GEN NameSuf
|
|
|
|
|
|
||| generate a fresh name with the given base
|
|
export
|
|
mn : Has NameGen fs => PBaseName -> Eff fs BaseName
|
|
mn base = do
|
|
i <- getAt GEN
|
|
modifyAt GEN S
|
|
pure $ MN base i
|
|
|
|
||| generate a fresh binding name with the given base and location `loc`
|
|
export
|
|
mnb : Has NameGen fs => PBaseName -> Loc -> Eff fs BindName
|
|
mnb base loc = pure $ BN !(mn base) loc
|
|
|
|
export
|
|
fresh : Has NameGen fs => BindName -> Eff fs BindName
|
|
fresh (BN (UN str) loc) = mnb str loc
|
|
fresh (BN (MN str k) loc) = mnb str loc
|
|
fresh (BN Unused loc) = mnb "x" loc
|