/- Copyright (c) 2022 Henrik Böving. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Henrik Böving, Xubai Wang -/ import DocGen4.Output.Base import DocGen4.Output.ToHtmlFormat import DocGen4.LeanInk.Process import Lean.Data.Json import LeanInk.Annotation.Alectryon namespace LeanInk.Annotation.Alectryon open DocGen4 Output open scoped DocGen4.Jsx structure AlectryonContext where counter : Nat abbrev AlectryonM := StateT AlectryonContext HtmlM def getNextButtonLabel : AlectryonM String := do let val ← get let newCounter := val.counter + 1 set { val with counter := newCounter } pure s!"plain-lean4-lean-chk{val.counter}" def TypeInfo.toHtml (tyi : TypeInfo) : AlectryonM Html := do pure
{tyi.name}/var> : /b> {tyi.type}/span> /span> /div> /blockquote> /div> /small> /div> def Token.processSemantic (t : Token) : Html := match t.semanticType with | some "Name.Attribute" => {t.raw}/span> | some "Name.Variable" => {t.raw}/span> | some "Keyword" => {t.raw}/span> | _ => Html.text t.raw def Token.toHtml (t : Token) : AlectryonM Html := do -- Right now t.link is always none from LeanInk, ignore it -- TODO: render docstring let mut parts := #[] if let some tyi := t.typeinfo then parts := parts.push $ ←tyi.toHtml parts := parts.push t.processSemantic pure -- TODO: Show rest of token [parts] /span> def Contents.toHtml : Contents → AlectryonM Html | .string value => pure {value} /span> | .experimentalTokens values => do let values ← values.mapM Token.toHtml pure [values] /span> def Hypothesis.toHtml (h : Hypothesis) : AlectryonM Html := do let mut hypParts := #[[h.names.intersperse ", " |>.map Html.text |>.toArray]/var>] if h.body != "" then hypParts := hypParts.push := /b> {h.body}/span> /span> hypParts := hypParts.push : /b> {h.type}/span> /span> pure [hypParts] /span> def Goal.toHtml (g : Goal) : AlectryonM Html := do let mut hypotheses := #[] for hyp in g.hypotheses do let rendered ← hyp.toHtml hypotheses := hypotheses.push rendered hypotheses := hypotheses.push
pure[hypotheses] /div>
{g.name}/span>/hr> /span>{g.conclusion} /div> /blockquote> def Message.toHtml (m : Message) : AlectryonM Html := do pure