mirror of
https://github.com/idris-lang/Idris2.git
synced 2024-11-28 02:23:44 +03:00
[ doc ] Add constructor docstrings (#2789)
Co-authored-by: Guillaume Allais <guillaume.allais@ens-lyon.org>
This commit is contained in:
parent
7f9db70e15
commit
2dbb824a93
@ -944,7 +944,7 @@ mutual
|
||||
let consb = map (\ (nm, tm) => (nm, doBind bnames tm)) cons'
|
||||
|
||||
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 =>
|
||||
elabInterface fc vis env nest consb
|
||||
tn paramsb det conname
|
||||
@ -1032,7 +1032,8 @@ mutual
|
||||
map fst params) (mkNamespace recName))
|
||||
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
|
||||
pure [IRecord fc (Just recName)
|
||||
vis mbtot (MkImpRecord fc tn paramsb opts conname fields')]
|
||||
|
@ -273,7 +273,7 @@ getDocsForName fc n config
|
||||
showVisible : Visibility -> Doc IdrisDocAnn
|
||||
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
|
||||
= do defs <- get Ctxt
|
||||
Just def <- lookupCtxtExact con (gamma defs)
|
||||
@ -281,7 +281,11 @@ getDocsForName fc n config
|
||||
| Nothing => pure Empty
|
||||
syn <- get Syn
|
||||
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
|
||||
[(n, "")] => pure conWithTypeDoc
|
||||
[(n, str)] => pure $ vcat
|
||||
@ -347,9 +351,11 @@ getDocsForName fc n config
|
||||
case !(traverse (pterm . map defaultKindedName) (parents iface)) of
|
||||
[] => []
|
||||
ps => [hsep (header "Constraints" :: punctuate comma (map (prettyBy Syntax) ps))]
|
||||
let icon = case dropNS (iconstructor iface) of
|
||||
DN _ _ => [] -- machine inserted
|
||||
nm => [hsep [header "Constructor", dCon nm (prettyName nm)]]
|
||||
icon <- do cName <- toFullNames (iconstructor iface)
|
||||
case dropNS cName of
|
||||
UN{} => do doc <- getDConDoc {showType = False} cName
|
||||
pure $ [header "Constructor" <++> annotate Declarations doc]
|
||||
_ => pure [] -- machine inserted
|
||||
mdocs <- traverse getMethDoc (methods iface)
|
||||
let meths = case concat mdocs of
|
||||
[] => []
|
||||
|
@ -9,6 +9,7 @@ import Core.Metadata
|
||||
import Core.TT
|
||||
import Core.Unify
|
||||
|
||||
import Idris.Doc.String
|
||||
import Idris.REPL.Opts
|
||||
import Idris.Syntax
|
||||
|
||||
@ -338,15 +339,16 @@ elabInterface : {vars : _} ->
|
||||
Name ->
|
||||
(params : List (Name, (RigCount, RawImp))) ->
|
||||
(dets : List Name) ->
|
||||
(conName : Maybe Name) ->
|
||||
(conName : Maybe (String, Name)) ->
|
||||
List ImpDecl ->
|
||||
Core ()
|
||||
elabInterface {vars} ifc vis env nest constraints iname params dets mcon body
|
||||
= do fullIName <- getFullName iname
|
||||
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
|
||||
conName <- inCurrentNS conName_in
|
||||
whenJust (fst <$> mcon) (addDocString conName)
|
||||
let meth_sigs = mapMaybe getSig body
|
||||
let meth_decls = map sigToDecl meth_sigs
|
||||
let meth_names = map name meth_decls
|
||||
|
@ -1551,11 +1551,12 @@ getVisibility (Just vis) (Left x :: xs)
|
||||
= fatalError "Multiple visibility modifiers"
|
||||
getVisibility v (_ :: xs) = getVisibility v xs
|
||||
|
||||
recordConstructor : OriginDesc -> Rule Name
|
||||
recordConstructor : OriginDesc -> Rule (String, Name)
|
||||
recordConstructor fname
|
||||
= do decorate fname Keyword $ exactIdent "constructor"
|
||||
= do doc <- optDocumentation fname
|
||||
decorate fname Keyword $ exactIdent "constructor"
|
||||
n <- mustWork $ decoratedDataConstructorName fname
|
||||
pure n
|
||||
pure (doc, n)
|
||||
|
||||
constraints : OriginDesc -> IndentInfo -> EmptyRule (List (Maybe Name, PTerm))
|
||||
constraints fname indents
|
||||
|
@ -487,7 +487,7 @@ mutual
|
||||
Core ( Name
|
||||
, List (Name, RigCount, PiInfo IPTerm, IPTerm)
|
||||
, List DataOpt
|
||||
, Maybe Name
|
||||
, Maybe (String, Name)
|
||||
, List (PField' KindedName))
|
||||
toPRecord (MkImpRecord fc n ps opts con fs)
|
||||
= do ps' <- traverse (\ (n, c, p, ty) =>
|
||||
@ -495,7 +495,7 @@ mutual
|
||||
p' <- mapPiInfo p
|
||||
pure (n, c, p', ty')) ps
|
||||
fs' <- traverse toPField fs
|
||||
pure (n, ps', opts, Just con, fs')
|
||||
pure (n, ps', opts, Just ("", con), fs')
|
||||
where
|
||||
mapPiInfo : PiInfo IRawImp -> Core (PiInfo IPTerm)
|
||||
mapPiInfo Explicit = pure Explicit
|
||||
|
@ -279,8 +279,8 @@ mutual
|
||||
MkPRecord : (tyname : Name) ->
|
||||
(params : List (Name, RigCount, PiInfo (PTerm' nm), PTerm' nm)) ->
|
||||
(opts : List DataOpt) ->
|
||||
(conName : Maybe Name) ->
|
||||
List (PField' nm) ->
|
||||
(conName : Maybe (String, Name)) ->
|
||||
(fields : List (PField' nm)) ->
|
||||
PRecordDecl' nm
|
||||
MkPRecordLater : (tyname : Name) ->
|
||||
(params : List (Name, RigCount, PiInfo (PTerm' nm), PTerm' nm)) ->
|
||||
@ -407,7 +407,7 @@ mutual
|
||||
(doc : String) ->
|
||||
(params : List (Name, (RigCount, PTerm' nm))) ->
|
||||
(det : List Name) ->
|
||||
(conName : Maybe Name) ->
|
||||
(conName : Maybe (String, Name)) ->
|
||||
List (PDecl' nm) ->
|
||||
PDecl' nm
|
||||
PImplementation : FC ->
|
||||
|
@ -6,6 +6,8 @@ infixr 5 ::
|
||||
infixr 5 ++
|
||||
|
||||
interface Monoid ty where
|
||||
||| Users can hand-craft their own monoid implementations
|
||||
constructor MkMonoid
|
||||
neutral : ty
|
||||
(++) : ty -> ty -> ty
|
||||
|
||||
|
@ -107,6 +107,8 @@ Main> Bye for now!
|
||||
1/1: Building List (List.idr)
|
||||
List> interface List.Monoid : Type -> Type
|
||||
Parameters: ty
|
||||
Constructor: MkMonoid
|
||||
Users can hand-craft their own monoid implementations
|
||||
Methods:
|
||||
neutral : ty
|
||||
(++) : ty -> ty -> ty
|
||||
|
Loading…
Reference in New Issue
Block a user