quox/lib/Quox/Name.idr

201 lines
4.2 KiB
Idris
Raw Normal View History

2021-07-20 16:05:19 -04:00
module Quox.Name
2023-05-01 21:06:25 -04:00
import Quox.Loc
2023-03-17 16:50:04 -04:00
import Quox.CharExtra
2021-07-20 16:05:19 -04:00
import public Data.SnocList
2022-05-06 18:57:23 -04:00
import Data.List
2023-05-01 21:06:25 -04:00
import Control.Eff
2023-03-16 13:34:49 -04:00
import Text.Lexer
2023-03-17 16:50:04 -04:00
import Derive.Prelude
2022-05-13 01:05:55 -04:00
%hide TT.Name
2021-07-20 16:05:19 -04:00
%default total
2022-05-13 01:05:55 -04:00
%language ElabReflection
2021-07-20 16:05:19 -04:00
2023-05-01 21:06:25 -04:00
public export
NameSuf : Type
NameSuf = Nat
2021-07-20 16:05:19 -04:00
public export
2022-05-13 01:05:55 -04:00
data BaseName
2023-05-01 21:06:25 -04:00
= UN String -- user-given name
| MN String NameSuf -- machine-generated name
| Unused -- "_"
2023-03-02 13:52:32 -05:00
%runElab derive "BaseName" [Eq, Ord]
2021-07-20 16:05:19 -04:00
export
baseStr : BaseName -> String
2023-05-01 21:06:25 -04:00
baseStr (UN x) = x
baseStr (MN x i) = "\{x}#\{show i}"
baseStr Unused = "_"
2021-07-20 16:05:19 -04:00
2023-05-25 12:34:13 -04:00
export Show BaseName where show = show . baseStr
export FromString BaseName where fromString = UN
public export
Mods : Type
Mods = SnocList String
2021-07-20 16:05:19 -04:00
public export
record Name where
constructor MakeName
mods : Mods
2021-07-20 16:05:19 -04:00
base : BaseName
2023-03-02 13:52:32 -05:00
%runElab derive "Name" [Eq, Ord]
2021-07-20 16:05:19 -04:00
2023-02-22 01:40:19 -05:00
public export %inline
unq : BaseName -> Name
unq = MakeName [<]
||| 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 MakePName
mods : Mods
base : PBaseName
%runElab derive "PName" [Eq, Ord]
export %inline
fromPName : PName -> Name
fromPName p = MakeName p.mods $ UN p.base
export %inline
toPName : Name -> PName
toPName p = MakePName p.mods $ baseStr p.base
export %inline
fromPBaseName : PBaseName -> Name
fromPBaseName = MakeName [<] . UN
export
Show PName where
2023-05-25 12:34:13 -04:00
show (MakePName mods base) =
show $ concat $ intersperse "." $ toList $ mods :< base
export Show Name where show = show . toPName
export FromString PName where fromString = MakePName [<]
export FromString Name where fromString = fromPBaseName
2023-05-01 21:06:25 -04:00
public export
record BindName where
constructor BN
name : BaseName
loc_ : Loc
%runElab derive "BindName" [Eq, Ord, Show]
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
2021-07-20 16:05:19 -04:00
export
toDots : Name -> String
toDots x = fastConcat $ cast $ map (<+> ".") x.mods :< baseStr x.base
2023-02-28 14:51:54 -05:00
export
fromListP : List1 String -> PName
fromListP (x ::: xs) = go [<] x xs where
go : SnocList String -> String -> List String -> PName
go mods x [] = MakePName mods x
2023-02-28 14:51:54 -05:00
go mods x (y :: ys) = go (mods :< x) y ys
export %inline
fromList : List1 String -> Name
fromList = fromPName . fromListP
2023-03-16 13:34:49 -04:00
export
syntaxChars : List Char
2023-05-21 14:09:34 -04:00
syntaxChars = ['(', ')', '[', ']', '{', '}', '"', '\'', ',', '.', ';', '^']
2023-03-16 13:34:49 -04:00
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)
2023-03-16 13:39:24 -04:00
export
isName : String -> Bool
isName str =
case scan name [] (unpack str) of
Just (_, []) => True
_ => False
2023-05-01 21:06:25 -04:00
public export
data GenTag = GEN
public export
NameGen : Type -> Type
NameGen = StateL GEN NameSuf
export
runNameGenWith : Has NameGen fs =>
NameSuf -> Eff fs a -> Eff (fs - NameGen) (a, NameSuf)
runNameGenWith = runStateAt GEN
export
runNameGen : Has NameGen fs => Eff fs a -> Eff (fs - NameGen) a
runNameGen = map fst . runNameGenWith 0
||| 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
||| (optionally) location `loc`
export
mnb : Has NameGen fs =>
PBaseName -> {default noLoc loc : Loc} -> Eff fs BindName
mnb base = 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}