Skip to content

Commit 3fe368e

Browse files
feat: allow Verso docstrings to suppose the existence of instances (#11476)
This PR adds a `` {givenInstance}`C` `` documentation role that adds an instance of `C` to the document's local assumptions.
1 parent f8866dc commit 3fe368e

File tree

5 files changed

+140
-8
lines changed

5 files changed

+140
-8
lines changed

src/Lean/Elab/DocString.lean

Lines changed: 13 additions & 7 deletions
Original file line numberDiff line numberDiff line change
@@ -86,6 +86,9 @@ structure State where
8686
-/
8787
lctx : LocalContext
8888
/--
89+
-/
90+
localInstances : LocalInstances
91+
/--
8992
The options.
9093
9194
The `MonadLift TermElabM DocM` instance runs the lifted action with these options, so elaboration
@@ -129,10 +132,10 @@ instance : MonadStateOf State DocM :=
129132

130133
instance : MonadLift TermElabM DocM where
131134
monadLift act := private DocM.mk fun _ _ st' => do
132-
let {openDecls, lctx, options, ..} := (← st'.get)
135+
let {openDecls, lctx, options, localInstances, ..} := (← st'.get)
133136
let v ←
134137
withTheReader Core.Context (fun ρ => { ρ with openDecls, options }) <|
135-
withTheReader Meta.Context (fun ρ => { ρ with lctx }) <|
138+
withTheReader Meta.Context (fun ρ => { ρ with lctx, localInstances }) <|
136139
act
137140
return v
138141

@@ -144,16 +147,19 @@ private builtin_initialize modDocstringStateExt : EnvExtension (Option ModuleDoc
144147
registerEnvExtension (pure none)
145148

146149
private def getModState
147-
[Monad m] [MonadEnv m] [MonadLiftT IO m] [MonadLCtx m]
150+
[Monad m] [MonadEnv m] [MonadLiftT IO m] [MonadLiftT MetaM m] [MonadLCtx m]
148151
[MonadResolveName m] [MonadOptions m] : m ModuleDocstringState := do
149152
if let some st := modDocstringStateExt.getState (← getEnv) then
150153
return st
151154
else
152-
let lctx ← getLCtx
155+
let scopes := [{header := "", isPublic := true}]
153156
let openDecls ← getOpenDecls
157+
let lctx ← getLCtx
158+
let localInstances ← Meta.getLocalInstances
154159
let options ← getOptions
155-
let scopes := [{header := "", isPublic := true}]
156-
let st : ModuleDocstringState := { scopes, openDecls, lctx, options, scopedExts := #[] }
160+
let scopedExts := #[]
161+
let st : ModuleDocstringState :=
162+
{ scopes, openDecls, lctx, localInstances, options, scopedExts }
157163
modifyEnv fun env =>
158164
modDocstringStateExt.setState env st
159165
return st
@@ -197,7 +203,7 @@ def DocM.exec (declName : Name) (binders : Syntax) (act : DocM α)
197203
let options ← getOptions
198204
let scopes := [{header := "", isPublic := true}]
199205
let ((v, _), _) ← withTheReader Meta.Context (fun ρ => { ρ with localInstances }) <|
200-
act.run { suggestionMode } |>.run {} |>.run { scopes, openDecls, lctx, options }
206+
act.run { suggestionMode } |>.run {} |>.run { scopes, openDecls, lctx, localInstances, options }
201207
pure v
202208
finally
203209
scopedEnvExtensionsRef.set sc

src/Lean/Elab/DocString/Builtin.lean

Lines changed: 92 additions & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -522,7 +522,6 @@ private def givenContents : ParserFn :=
522522
optionalFn (symbolFn ":" >> termParser.fn)))
523523
(symbolFn ",")
524524

525-
526525
/--
527526
A metavariable to be discussed in the remainder of the docstring.
528527
@@ -612,6 +611,98 @@ def given (type : Option StrLit := none) (typeIsMeta : flag false) («show» : f
612611
|> Inline.concat
613612
else return .empty
614613

614+
private def givenInstanceContents : ParserFn :=
615+
whitespace >>
616+
sepBy1Fn false
617+
(nodeFn nullKind
618+
(optionalFn (atomicFn (identFn >> symbolFn ":")) >>
619+
termParser.fn))
620+
(symbolFn ",")
621+
622+
/--
623+
An instance metavariable to be discussed in the remainder of the docstring.
624+
625+
This is similar to {given}, but the resulting variable is marked for instance synthesis
626+
(with `BinderInfo.instImplicit`), and the name is optional.
627+
628+
There are two syntaxes that can be used:
629+
* `` {givenInstance}`T` `` establishes an unnamed instance of type `T`.
630+
* `` {givenInstance}`x : T` `` establishes a named instance `x` of type `T`.
631+
632+
Additionally, the contents of the code literal can be repeated, with comma separators.
633+
634+
If the `show` flag is `false` (default `true`), then the instance is not shown in the docstring.
635+
-/
636+
@[builtin_doc_role]
637+
def givenInstanceshow» : flag true) (xs : TSyntaxArray `inline) :
638+
DocM (Inline ElabInline) := do
639+
let s ← onlyCode xs
640+
641+
let stxs ← parseStrLit givenInstanceContents s
642+
let stxs := stxs.getArgs.mapIdx Prod.mk |>.filterMap fun (n, s) =>
643+
if n % 2 = 0 then some s else none
644+
let mut lctx ← getLCtx
645+
let mut localInstances ← Meta.getLocalInstances
646+
let mut instCounter := 0
647+
for stx in stxs do
648+
let nameColonOpt := stx[0][0]
649+
let tyStx := stx[1]
650+
651+
let ty' : Expr ← elabType tyStx
652+
let class? ← Meta.isClass? ty'
653+
let some className := class?
654+
| throwError m!"Expected a type class, but got `{.ofExpr ty'}`"
655+
656+
-- Generate a fresh name if no name is provided
657+
let (userName, hasUserName) ←
658+
if nameColonOpt.isMissing then
659+
instCounter := instCounter + 1
660+
let n ← mkFreshUserName (`inst ++ className)
661+
pure (n, false)
662+
else
663+
pure (nameColonOpt.getId, true)
664+
665+
let fv ← mkFreshFVarId
666+
lctx := lctx.mkLocalDecl fv userName ty' BinderInfo.instImplicit
667+
localInstances := localInstances.push { fvar := .fvar fv, className }
668+
669+
if hasUserName then
670+
addTermInfo' nameColonOpt[0] (.fvar fv)
671+
(lctx? := some lctx) (isBinder := true) (expectedType? := some ty')
672+
673+
modify fun st => { st with lctx, localInstances }
674+
675+
if «show» then
676+
let text ← getFileMap
677+
let mut outStrs := #[]
678+
let mut failed := false
679+
for stx in stxs do
680+
let nameColonOpt := stx[0][0]
681+
let thisStr ←
682+
if let some ⟨b', e'⟩ := stx[1].getRange? then
683+
-- Has type annotation
684+
if nameColonOpt.isMissing then
685+
-- No name, just show type
686+
pure <| String.Pos.Raw.extract text.source b' e'
687+
else
688+
-- Has name and type, nameColonOpt is `identFn >> symbolFn ":"`
689+
if let some ⟨b, e⟩ := nameColonOpt[0].getRange? then
690+
pure <| s!"{b.extract text.source e} : {b'.extract text.source e'}"
691+
else
692+
failed := true
693+
break
694+
else
695+
failed := true
696+
break
697+
outStrs := outStrs.push thisStr
698+
if failed then
699+
return .code s.getString
700+
else
701+
return outStrs.map Inline.code
702+
|>.toList |>.intersperse (Inline.text ", ") |>.toArray
703+
|> Inline.concat
704+
else return .empty
705+
615706

616707
private def firstToken? (stx : Syntax) : Option Syntax :=
617708
stx.find? fun

stage0/src/stdlib_flags.h

Lines changed: 2 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -1,5 +1,7 @@
11
#include "util/options.h"
22

3+
// Dear CI, please update stage0
4+
35
namespace lean {
46
options get_default_options() {
57
options opts;

tests/lean/run/versoDocMarkdown.lean

Lines changed: 27 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -133,6 +133,7 @@ Lean.Doc.assert'
133133
Lean.Doc.attr
134134
Lean.Doc.conv
135135
Lean.Doc.given
136+
Lean.Doc.givenInstance
136137
Lean.Doc.kw
137138
Lean.Doc.kw!
138139
Lean.Doc.kw?
@@ -190,6 +191,32 @@ Visible:
190191
-/
191192
def givenTests (x y : Nat) : Nat := x - y
192193
194+
/--
195+
{given -show}`α : Type, β : Type, γ : Type` {given -show}`x : α, y : α, z : α`
196+
Invisible: {givenInstance -show}`Add α` {givenInstance -show}`addInst : Add β`
197+
198+
There is an {lean}`addInst : Add β` and an {lean}`inferInstance : Add α`, and {lean}`x + y + z`.
199+
200+
Visible: {givenInstance}`Add γ`&{givenInstance}`addInst : OfNat γ 5`
201+
202+
Check: {lean}`(5 : γ) + 5`
203+
204+
-/
205+
def givenInstanceTests (x y : Nat) : Nat := x - y
206+
207+
/--
208+
info: ⏎
209+
Invisible: ⏎
210+
211+
There is an `addInst : Add β` and an `inferInstance : Add α`, and `x + y + z`\.
212+
213+
Visible: `Add γ`&`addInst : OfNat γ 5`
214+
215+
Check: `(5 : γ) + 5`
216+
-/
217+
#guard_msgs in
218+
#verso_to_markdown givenInstanceTests
219+
193220
/--
194221
info: Invisible: ⏎
195222

tests/lean/run/versoDocs.lean

Lines changed: 6 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -96,6 +96,12 @@ def rot (n : Nat) (xs : Array α) : Array α :=
9696
#eval rot 5 #[1, 2, 3]
9797

9898

99+
/--
100+
Given {given}`α : Type` and {given}`x : α, y : α` with an instance of {givenInstance}`Add α`,
101+
{lean}`x + y : α`.
102+
-/
103+
def givens := ()
104+
99105
/--
100106
Given {given}`xs : List α`, finds lists {given}`ys` and {given}`zs` such that {lean}`xs = ys ++ zs`
101107
and {lean}`∀x ∈ xs, p x` and {lean}`zs.head?.all (¬p ·)`.

0 commit comments

Comments
 (0)