[ doc ] Add constructor docstrings (#2789)

Co-authored-by: Guillaume Allais <guillaume.allais@ens-lyon.org>
This commit is contained in:
pinselimo 2023-02-19 11:15:39 +01:00 committed by GitHub
parent 7f9db70e15
commit 2dbb824a93
No known key found for this signature in database
GPG Key ID: 4AEE18F83AFDEB23
8 changed files with 31 additions and 17 deletions

View File

@ -944,7 +944,7 @@ mutual
let consb = map (\ (nm, tm) => (nm, doBind bnames tm)) cons' let consb = map (\ (nm, tm) => (nm, doBind bnames tm)) cons'
body' <- traverse (desugarDecl (ps ++ mnames ++ paramNames)) body body' <- traverse (desugarDecl (ps ++ mnames ++ paramNames)) body
pure [IPragma fc (maybe [tn] (\n => [tn, n]) conname) pure [IPragma fc (maybe [tn] (\n => [tn, snd n]) conname)
(\nest, env => (\nest, env =>
elabInterface fc vis env nest consb elabInterface fc vis env nest consb
tn paramsb det conname tn paramsb det conname
@ -1032,7 +1032,8 @@ mutual
map fst params) (mkNamespace recName)) map fst params) (mkNamespace recName))
fields fields
let _ = the (List IField) fields' let _ = the (List IField) fields'
let conname = maybe (mkConName tn) id conname_in let conname = maybe (mkConName tn) snd conname_in
whenJust (fst <$> conname_in) (addDocString conname)
let _ = the Name conname let _ = the Name conname
pure [IRecord fc (Just recName) pure [IRecord fc (Just recName)
vis mbtot (MkImpRecord fc tn paramsb opts conname fields')] vis mbtot (MkImpRecord fc tn paramsb opts conname fields')]

View File

@ -273,7 +273,7 @@ getDocsForName fc n config
showVisible : Visibility -> Doc IdrisDocAnn showVisible : Visibility -> Doc IdrisDocAnn
showVisible vis = header "Visibility" <++> annotate (Syntax Keyword) (pretty0 vis) showVisible vis = header "Visibility" <++> annotate (Syntax Keyword) (pretty0 vis)
getDConDoc : Name -> Core (Doc IdrisDocAnn) getDConDoc : {default True showType : Bool} -> Name -> Core (Doc IdrisDocAnn)
getDConDoc con getDConDoc con
= do defs <- get Ctxt = do defs <- get Ctxt
Just def <- lookupCtxtExact con (gamma defs) Just def <- lookupCtxtExact con (gamma defs)
@ -281,7 +281,11 @@ getDocsForName fc n config
| Nothing => pure Empty | Nothing => pure Empty
syn <- get Syn syn <- get Syn
ty <- prettyType Syntax (type def) ty <- prettyType Syntax (type def)
let conWithTypeDoc = annotate (Decl con) (hsep [dCon con (prettyName con), colon, ty]) let conWithTypeDoc
= annotate (Decl con)
$ ifThenElse showType
(hsep [dCon con (prettyName con), colon, ty])
(dCon con (prettyName con))
case lookupName con (defDocstrings syn) of case lookupName con (defDocstrings syn) of
[(n, "")] => pure conWithTypeDoc [(n, "")] => pure conWithTypeDoc
[(n, str)] => pure $ vcat [(n, str)] => pure $ vcat
@ -347,9 +351,11 @@ getDocsForName fc n config
case !(traverse (pterm . map defaultKindedName) (parents iface)) of case !(traverse (pterm . map defaultKindedName) (parents iface)) of
[] => [] [] => []
ps => [hsep (header "Constraints" :: punctuate comma (map (prettyBy Syntax) ps))] ps => [hsep (header "Constraints" :: punctuate comma (map (prettyBy Syntax) ps))]
let icon = case dropNS (iconstructor iface) of icon <- do cName <- toFullNames (iconstructor iface)
DN _ _ => [] -- machine inserted case dropNS cName of
nm => [hsep [header "Constructor", dCon nm (prettyName nm)]] UN{} => do doc <- getDConDoc {showType = False} cName
pure $ [header "Constructor" <++> annotate Declarations doc]
_ => pure [] -- machine inserted
mdocs <- traverse getMethDoc (methods iface) mdocs <- traverse getMethDoc (methods iface)
let meths = case concat mdocs of let meths = case concat mdocs of
[] => [] [] => []

View File

@ -9,6 +9,7 @@ import Core.Metadata
import Core.TT import Core.TT
import Core.Unify import Core.Unify
import Idris.Doc.String
import Idris.REPL.Opts import Idris.REPL.Opts
import Idris.Syntax import Idris.Syntax
@ -338,15 +339,16 @@ elabInterface : {vars : _} ->
Name -> Name ->
(params : List (Name, (RigCount, RawImp))) -> (params : List (Name, (RigCount, RawImp))) ->
(dets : List Name) -> (dets : List Name) ->
(conName : Maybe Name) -> (conName : Maybe (String, Name)) ->
List ImpDecl -> List ImpDecl ->
Core () Core ()
elabInterface {vars} ifc vis env nest constraints iname params dets mcon body elabInterface {vars} ifc vis env nest constraints iname params dets mcon body
= do fullIName <- getFullName iname = do fullIName <- getFullName iname
ns_iname <- inCurrentNS fullIName ns_iname <- inCurrentNS fullIName
let conName_in = maybe (mkCon vfc fullIName) id mcon let conName_in = maybe (mkCon vfc fullIName) snd mcon
-- Machine generated names need to be qualified when looking them up -- Machine generated names need to be qualified when looking them up
conName <- inCurrentNS conName_in conName <- inCurrentNS conName_in
whenJust (fst <$> mcon) (addDocString conName)
let meth_sigs = mapMaybe getSig body let meth_sigs = mapMaybe getSig body
let meth_decls = map sigToDecl meth_sigs let meth_decls = map sigToDecl meth_sigs
let meth_names = map name meth_decls let meth_names = map name meth_decls

View File

@ -1551,11 +1551,12 @@ getVisibility (Just vis) (Left x :: xs)
= fatalError "Multiple visibility modifiers" = fatalError "Multiple visibility modifiers"
getVisibility v (_ :: xs) = getVisibility v xs getVisibility v (_ :: xs) = getVisibility v xs
recordConstructor : OriginDesc -> Rule Name recordConstructor : OriginDesc -> Rule (String, Name)
recordConstructor fname recordConstructor fname
= do decorate fname Keyword $ exactIdent "constructor" = do doc <- optDocumentation fname
decorate fname Keyword $ exactIdent "constructor"
n <- mustWork $ decoratedDataConstructorName fname n <- mustWork $ decoratedDataConstructorName fname
pure n pure (doc, n)
constraints : OriginDesc -> IndentInfo -> EmptyRule (List (Maybe Name, PTerm)) constraints : OriginDesc -> IndentInfo -> EmptyRule (List (Maybe Name, PTerm))
constraints fname indents constraints fname indents

View File

@ -487,7 +487,7 @@ mutual
Core ( Name Core ( Name
, List (Name, RigCount, PiInfo IPTerm, IPTerm) , List (Name, RigCount, PiInfo IPTerm, IPTerm)
, List DataOpt , List DataOpt
, Maybe Name , Maybe (String, Name)
, List (PField' KindedName)) , List (PField' KindedName))
toPRecord (MkImpRecord fc n ps opts con fs) toPRecord (MkImpRecord fc n ps opts con fs)
= do ps' <- traverse (\ (n, c, p, ty) => = do ps' <- traverse (\ (n, c, p, ty) =>
@ -495,7 +495,7 @@ mutual
p' <- mapPiInfo p p' <- mapPiInfo p
pure (n, c, p', ty')) ps pure (n, c, p', ty')) ps
fs' <- traverse toPField fs fs' <- traverse toPField fs
pure (n, ps', opts, Just con, fs') pure (n, ps', opts, Just ("", con), fs')
where where
mapPiInfo : PiInfo IRawImp -> Core (PiInfo IPTerm) mapPiInfo : PiInfo IRawImp -> Core (PiInfo IPTerm)
mapPiInfo Explicit = pure Explicit mapPiInfo Explicit = pure Explicit

View File

@ -279,8 +279,8 @@ mutual
MkPRecord : (tyname : Name) -> MkPRecord : (tyname : Name) ->
(params : List (Name, RigCount, PiInfo (PTerm' nm), PTerm' nm)) -> (params : List (Name, RigCount, PiInfo (PTerm' nm), PTerm' nm)) ->
(opts : List DataOpt) -> (opts : List DataOpt) ->
(conName : Maybe Name) -> (conName : Maybe (String, Name)) ->
List (PField' nm) -> (fields : List (PField' nm)) ->
PRecordDecl' nm PRecordDecl' nm
MkPRecordLater : (tyname : Name) -> MkPRecordLater : (tyname : Name) ->
(params : List (Name, RigCount, PiInfo (PTerm' nm), PTerm' nm)) -> (params : List (Name, RigCount, PiInfo (PTerm' nm), PTerm' nm)) ->
@ -407,7 +407,7 @@ mutual
(doc : String) -> (doc : String) ->
(params : List (Name, (RigCount, PTerm' nm))) -> (params : List (Name, (RigCount, PTerm' nm))) ->
(det : List Name) -> (det : List Name) ->
(conName : Maybe Name) -> (conName : Maybe (String, Name)) ->
List (PDecl' nm) -> List (PDecl' nm) ->
PDecl' nm PDecl' nm
PImplementation : FC -> PImplementation : FC ->

View File

@ -6,6 +6,8 @@ infixr 5 ::
infixr 5 ++ infixr 5 ++
interface Monoid ty where interface Monoid ty where
||| Users can hand-craft their own monoid implementations
constructor MkMonoid
neutral : ty neutral : ty
(++) : ty -> ty -> ty (++) : ty -> ty -> ty

View File

@ -107,6 +107,8 @@ Main> Bye for now!
1/1: Building List (List.idr) 1/1: Building List (List.idr)
List> interface List.Monoid : Type -> Type List> interface List.Monoid : Type -> Type
Parameters: ty Parameters: ty
Constructor: MkMonoid
Users can hand-craft their own monoid implementations
Methods: Methods:
neutral : ty neutral : ty
(++) : ty -> ty -> ty (++) : ty -> ty -> ty