Skip to content

Commit

Permalink
feat: show inductive constructor argument names (#249)
Browse files Browse the repository at this point in the history
* feat: show inductive constructor argument names

* refactor: move argToHtml to DocGen4.Output.Arg
  • Loading branch information
cppio authored Dec 17, 2024
1 parent 589875a commit e473211
Show file tree
Hide file tree
Showing 6 changed files with 35 additions and 30 deletions.
22 changes: 22 additions & 0 deletions DocGen4/Output/Arg.lean
Original file line number Diff line number Diff line change
@@ -0,0 +1,22 @@
import DocGen4.Output.Base

namespace DocGen4
namespace Output

open scoped DocGen4.Jsx

/--
Render an `Arg` as HTML, adding opacity effects etc. depending on what
type of binder it has.
-/
def argToHtml (arg : Process.Arg) : HtmlM Html := do
let node ← infoFormatToHtml arg.binder
let inner := <span class="fn">[node]</span>
let html := Html.element "span" false #[("class", "decl_args")] #[inner]
if arg.implicit then
return <span class="impl_arg">{html}</span>
else
return html

end Output
end DocGen4
8 changes: 5 additions & 3 deletions DocGen4/Output/Inductive.lean
Original file line number Diff line number Diff line change
@@ -1,3 +1,4 @@
import DocGen4.Output.Arg
import DocGen4.Output.Template
import DocGen4.Output.DocString
import DocGen4.Process
Expand All @@ -15,20 +16,21 @@ def instancesForToHtml (typeName : Name) : BaseHtmlM Html := do
<ul class="instances-for-enum"></ul>
</details>

def ctorToHtml (c : Process.NameInfo) : HtmlM Html := do
def ctorToHtml (c : Process.ConstructorInfo) : HtmlM Html := do
let shortName := c.name.componentsRev.head!.toString
let name := c.name.toString
let args ← c.args.mapM argToHtml
if let some doc := c.doc then
let renderedDoc ← docStringToHtml doc name
pure
<li class="constructor" id={name}>
{shortName} : [← infoFormatToHtml c.type]
{shortName} [args] {" : "} [← infoFormatToHtml c.type]
<div class="inductive_ctor_doc">[renderedDoc]</div>
</li>
else
pure
<li class="constructor" id={name}>
{shortName} : [← infoFormatToHtml c.type]
{shortName} [args] {" : "} [← infoFormatToHtml c.type]
</li>

def inductiveToHtml (i : Process.InductiveInfo) : HtmlM (Array Html) := do
Expand Down
13 changes: 0 additions & 13 deletions DocGen4/Output/Module.lean
Original file line number Diff line number Diff line change
Expand Up @@ -20,19 +20,6 @@ namespace Output
open scoped DocGen4.Jsx
open Lean Process

/--
Render an `Arg` as HTML, adding opacity effects etc. depending on what
type of binder it has.
-/
def argToHtml (arg : Arg) : HtmlM Html := do
let node ← infoFormatToHtml arg.binder
let inner := <span class="fn">[node]</span>
let html := Html.element "span" false #[("class", "decl_args")] #[inner]
if arg.implicit then
return <span class="impl_arg">{html}</span>
else
return html

/--
Render the structures this structure extends from as HTML so it can be
added to the top level.
Expand Down
12 changes: 6 additions & 6 deletions DocGen4/Process/Base.lean
Original file line number Diff line number Diff line change
Expand Up @@ -117,14 +117,19 @@ structure InstanceInfo extends DefinitionInfo where
typeNames : Array Name
deriving Inhabited

/--
Information about a constructor of an inductive type
-/
abbrev ConstructorInfo := Info

/--
Information about an `inductive` declaration
-/
structure InductiveInfo extends Info where
/--
List of all constructors of this inductive type.
-/
ctors : List NameInfo
ctors : List ConstructorInfo
isUnsafe : Bool
deriving Inhabited

Expand Down Expand Up @@ -166,11 +171,6 @@ Information about a `class inductive` declaration.
abbrev ClassInductiveInfo := InductiveInfo


/--
Information about a constructor of an inductive type
-/
abbrev ConstructorInfo := Info

/--
A general type for informations about declarations.
-/
Expand Down
8 changes: 1 addition & 7 deletions DocGen4/Process/InductiveInfo.lean
Original file line number Diff line number Diff line change
Expand Up @@ -12,15 +12,9 @@ namespace DocGen4.Process

open Lean Meta

def getConstructorType (ctor : Name) : MetaM Expr := do
let env ← getEnv
match env.find? ctor with
| some (ConstantInfo.ctorInfo i) => pure i.type
| _ => panic! s!"Constructor {ctor} was requested but does not exist"

def InductiveInfo.ofInductiveVal (v : InductiveVal) : MetaM InductiveInfo := do
let info ← Info.ofConstantVal v.toConstantVal
let ctors ← v.ctors.mapM (fun name => do NameInfo.ofTypedName name (← getConstructorType name))
let ctors ← v.ctors.mapM (fun name => do Info.ofConstantVal (← getConstInfoCtor name).toConstantVal)
return {
toInfo := info,
ctors,
Expand Down
2 changes: 1 addition & 1 deletion static/style.css
Original file line number Diff line number Diff line change
Expand Up @@ -753,7 +753,7 @@ a:hover {
transition: opacity 300ms ease-in;
}

.decl_header:not(:hover) .impl_arg {
.decl_header:not(:hover) .impl_arg, .constructor:not(:hover) .impl_arg {
opacity: 30%;
transition: opacity 1000ms ease-out;
}
Expand Down

0 comments on commit e473211

Please sign in to comment.