Skip to content

Commit

Permalink
refactor: nicer modifiers/ranges API
Browse files Browse the repository at this point in the history
  • Loading branch information
Kha committed Oct 21, 2024
1 parent 682173d commit d8d9203
Show file tree
Hide file tree
Showing 9 changed files with 31 additions and 29 deletions.
2 changes: 1 addition & 1 deletion src/Lean/Elab/BuiltinCommand.lean
Original file line number Diff line number Diff line change
Expand Up @@ -343,7 +343,7 @@ def failIfSucceeds (x : CommandElabM Unit) : CommandElabM Unit := do
if let .none ← findDeclarationRangesCore? declName then
-- this is only relevant for declarations added without a declaration range
-- in particular `Quot.mk` et al which are added by `init_quot`
addAuxDeclarationRanges declName stx id
addDeclarationRangesFromSyntax declName stx id
addDocString declName (← getDocStringText doc)
| _ => throwUnsupportedSyntax

Expand Down
2 changes: 1 addition & 1 deletion src/Lean/Elab/BuiltinEvalCommand.lean
Original file line number Diff line number Diff line change
Expand Up @@ -84,7 +84,7 @@ private def addAndCompileExprForEval (declName : Name) (value : Expr) (allowSorr
-- An alternative design would be to make `elabTermForEval` into a term elaborator and elaborate the command all at once
-- with `unsafe def _eval := term_for_eval% $t`, which we did try, but unwanted error messages
-- such as "failed to infer definition type" can surface.
let defView := mkDefViewOfDef { isUnsafe := true, stx := ⟨.missing⟩ }
let defView := mkDefViewOfDef { isUnsafe := true }
(← `(Parser.Command.definition|
def $(mkIdent <| `_root_ ++ declName) := $(← Term.exprToSyntax value)))
Term.elabMutualDef #[] { header := "" } #[defView]
Expand Down
3 changes: 2 additions & 1 deletion src/Lean/Elab/DeclModifiers.lean
Original file line number Diff line number Diff line change
Expand Up @@ -57,7 +57,8 @@ inductive RecKind where

/-- Flags and data added to declarations (eg docstrings, attributes, `private`, `unsafe`, `partial`, ...). -/
structure Modifiers where
stx : TSyntax ``Parser.Command.declModifiers
/-- Input syntax, used for adjusting declaration range (unless missing) -/
stx : TSyntax ``Parser.Command.declModifiers := ⟨.missing⟩
docString? : Option String := none
visibility : Visibility := Visibility.regular
isNoncomputable : Bool := false
Expand Down
8 changes: 4 additions & 4 deletions src/Lean/Elab/Declaration.lean
Original file line number Diff line number Diff line change
Expand Up @@ -104,7 +104,7 @@ def elabAxiom (modifiers : Modifiers) (stx : Syntax) : CommandElabM Unit := do
let (binders, typeStx) := expandDeclSig stx[2]
let scopeLevelNames ← getLevelNames
let ⟨_, declName, allUserLevelNames⟩ ← expandDeclId declId modifiers
addDeclarationRanges declName modifiers.stx stx
addDeclarationRangesForBuiltin declName modifiers.stx stx
runTermElabM fun vars =>
Term.withDeclName declName <| Term.withLevelNames allUserLevelNames <| Term.elabBinders binders.getArgs fun xs => do
Term.applyAttributesAt declName modifiers.attrs AttributeApplicationTime.beforeElaboration
Expand Down Expand Up @@ -144,7 +144,7 @@ private def inductiveSyntaxToView (modifiers : Modifiers) (decl : Syntax) : Comm
let (binders, type?) := expandOptDeclSig decl[2]
let declId := decl[1]
let ⟨name, declName, levelNames⟩ ← expandDeclId declId modifiers
addDeclarationRanges declName modifiers.stx decl
addDeclarationRangesForBuiltin declName modifiers.stx decl
let ctors ← decl[4].getArgs.mapM fun ctor => withRef ctor do
-- def ctor := leading_parser optional docComment >> "\n| " >> declModifiers >> rawIdent >> optDeclSig
let mut ctorModifiers ← elabModifiers ⟨ctor[2]⟩
Expand All @@ -162,7 +162,7 @@ private def inductiveSyntaxToView (modifiers : Modifiers) (decl : Syntax) : Comm
let ctorName ← withRef ctor[3] <| applyVisibility ctorModifiers.visibility ctorName
let (binders, type?) := expandOptDeclSig ctor[4]
addDocString' ctorName ctorModifiers.docString?
addAuxDeclarationRanges ctorName ctor ctor[3]
addDeclarationRangesFromSyntax ctorName ctor ctor[3]
return { ref := ctor, modifiers := ctorModifiers, declName := ctorName, binders := binders, type? := type? : CtorView }
let computedFields ← (decl[5].getOptional?.map (·[1].getArgs) |>.getD #[]).mapM fun cf => withRef cf do
return { ref := cf, modifiers := cf[0], fieldId := cf[1].getId, type := ⟨cf[3]⟩, matchAlts := ⟨cf[4]⟩ }
Expand Down Expand Up @@ -400,7 +400,7 @@ def elabMutual : CommandElab := fun stx => do
-- We need to add `id`'s ranges *before* elaborating `initFn` (and then `id` itself) as
-- otherwise the info context created by `with_decl_name` will be incomplete and break the
-- call hierarchy
addDeclarationRanges fullId ⟨defStx.raw[0]⟩ defStx.raw[1]
addDeclarationRangesForBuiltin fullId ⟨defStx.raw[0]⟩ defStx.raw[1]
elabCommand (← `(
$[unsafe%$unsafe?]? def initFn : IO $type := with_decl_name% $(mkIdent fullId) do $doSeq
$defStx:command))
Expand Down
28 changes: 15 additions & 13 deletions src/Lean/Elab/DeclarationRange.lean
Original file line number Diff line number Diff line change
Expand Up @@ -49,29 +49,31 @@ def getDeclarationSelectionRef (stx : Syntax) : Syntax :=
else
stx[0]

/--
Derives and adds declaration ranges from given syntax trees. If `rangeStx` does not have a range,
nothing is added. If `selectionRangeStx` does not have a range, it is defaulted to that of
`rangeStx`.
-/
def addDeclarationRangesFromSyntax [Monad m] [MonadEnv m] [MonadFileMap m] (declName : Name)
(rangeStx : Syntax) (selectionRangeStx : Syntax := .missing) : m Unit := do
-- may fail on partial syntax, ignore in that case
let some range ← getDeclarationRange? rangeStx | return
let selectionRange ← (·.getD range) <$> getDeclarationRange? selectionRangeStx
Lean.addDeclarationRanges declName { range, selectionRange }

/--
Stores the `range` and `selectionRange` for `declName` where `modsStx` is the modifier part and
`cmdStx` the remaining part of the syntax tree for `declName`.
This method is for the builtin declarations only. User-defined commands should use
`Lean.addDeclarationRanges` to store this information for their commands.
`Lean.Elab.addDeclarationRangesFromSyntax` or `Lean.addDeclarationRanges` to store this information
for their commands.
-/
def addDeclarationRanges [Monad m] [MonadEnv m] [MonadFileMap m] (declName : Name)
def addDeclarationRangesForBuiltin [Monad m] [MonadEnv m] [MonadFileMap m] (declName : Name)
(modsStx : TSyntax ``Parser.Command.declModifiers) (declStx : Syntax) : m Unit := do
if declStx.getKind == ``Parser.Command.«example» then
return ()
let stx := mkNullNode #[modsStx, declStx]
-- may fail on partial syntax, ignore in that case
let some range ← getDeclarationRange? stx | return
let some selectionRange ← getDeclarationRange? (getDeclarationSelectionRef declStx) | return
Lean.addDeclarationRanges declName { range, selectionRange }

/-- Auxiliary method for recording ranges for auxiliary declarations (e.g., fields, nested declarations, etc. -/
def addAuxDeclarationRanges [Monad m] [MonadEnv m] [MonadFileMap m] (declName : Name) (stx : Syntax) (header : Syntax) : m Unit := do
-- may fail on partial syntax, ignore in that case
let some range ← getDeclarationRange? stx | return
let some selectionRange ← getDeclarationRange? header | return
Lean.addDeclarationRanges declName { range, selectionRange }
addDeclarationRangesFromSyntax declName stx (getDeclarationSelectionRef declStx)

end Lean.Elab
2 changes: 1 addition & 1 deletion src/Lean/Elab/LetRec.lean
Original file line number Diff line number Diff line change
Expand Up @@ -51,7 +51,7 @@ private def mkLetRecDeclView (letRec : Syntax) : TermElabM LetRecView := do
checkNotAlreadyDeclared declName
applyAttributesAt declName attrs AttributeApplicationTime.beforeElaboration
addDocString' declName docStr?
addAuxDeclarationRanges declName decl declId
addDeclarationRangesFromSyntax declName decl declId
let binders := decl[1].getArgs
let typeStx := expandOptType declId decl[2]
let (type, binderIds) ← elabBindersEx binders fun xs => do
Expand Down
3 changes: 1 addition & 2 deletions src/Lean/Elab/MutualDef.lean
Original file line number Diff line number Diff line change
Expand Up @@ -887,7 +887,6 @@ def getKindForLetRecs (mainHeaders : Array DefViewElabHeader) : DefKind :=
else DefKind.«def»

def getModifiersForLetRecs (mainHeaders : Array DefViewElabHeader) : Modifiers := {
stx := ⟨mkNullNode #[]⟩ -- ignore when computing declaration range
isNoncomputable := mainHeaders.any fun h => h.modifiers.isNoncomputable
recKind := if mainHeaders.any fun h => h.modifiers.isPartial then RecKind.partial else RecKind.default
isUnsafe := mainHeaders.any fun h => h.modifiers.isUnsafe
Expand Down Expand Up @@ -998,7 +997,7 @@ where
for view in views, header in headers do
-- NOTE: this should be the full `ref`, and thus needs to be done after any snapshotting
-- that depends only on a part of the ref
addDeclarationRanges header.declName view.modifiers.stx view.ref
addDeclarationRangesForBuiltin header.declName view.modifiers.stx view.ref


processDeriving (headers : Array DefViewElabHeader) := do
Expand Down
8 changes: 4 additions & 4 deletions src/Lean/Elab/Structure.lean
Original file line number Diff line number Diff line change
Expand Up @@ -94,7 +94,7 @@ private def expandCtor (structStx : Syntax) (structModifiers : Modifiers) (struc
let useDefault := do
let declName := structDeclName ++ defaultCtorName
let ref := structStx[1].mkSynthetic
addAuxDeclarationRanges declName ref ref
addDeclarationRangesFromSyntax declName ref
pure { ref, modifiers := default, name := defaultCtorName, declName }
if structStx[5].isNone then
useDefault
Expand All @@ -115,7 +115,7 @@ private def expandCtor (structStx : Syntax) (structModifiers : Modifiers) (struc
let declName := structDeclName ++ name
let declName ← applyVisibility ctorModifiers.visibility declName
addDocString' declName ctorModifiers.docString?
addAuxDeclarationRanges declName ctor[1] ctor[1]
addDeclarationRangesFromSyntax declName ctor[1]
pure { ref := ctor[1], name, modifiers := ctorModifiers, declName }

def checkValidFieldModifier (modifiers : Modifiers) : TermElabM Unit := do
Expand Down Expand Up @@ -815,7 +815,7 @@ private def elabStructureView (view : StructView) : TermElabM Unit := do
view.fields.forM fun field => do
if field.declName == view.ctor.declName then
throwErrorAt field.ref "invalid field name '{field.name}', it is equal to structure constructor name"
addAuxDeclarationRanges field.declName field.ref field.ref
addDeclarationRangesFromSyntax field.declName field.ref
let type ← Term.elabType view.type
unless validStructType type do throwErrorAt view.type "expected Type"
withRef view.ref do
Expand Down Expand Up @@ -912,7 +912,7 @@ def elabStructure (modifiers : Modifiers) (stx : Syntax) : CommandElabM Unit :=
let scopeLevelNames ← Term.getLevelNames
let ⟨name, declName, allUserLevelNames⟩ ← Elab.expandDeclId (← getCurrNamespace) scopeLevelNames declId modifiers
Term.withAutoBoundImplicitForbiddenPred (fun n => name == n) do
addDeclarationRanges declName modifiers.stx stx
addDeclarationRangesForBuiltin declName modifiers.stx stx
Term.withDeclName declName do
let ctor ← expandCtor stx modifiers declName
let fields ← expandFields stx modifiers declName
Expand Down
4 changes: 2 additions & 2 deletions src/Lean/Elab/Tactic/Ext.lean
Original file line number Diff line number Diff line change
Expand Up @@ -123,7 +123,7 @@ def realizeExtTheorem (structName : Name) (flat : Bool) : Elab.Command.CommandEl
levelParams := info.levelParams
}
modifyEnv fun env => addProtected env extName
addAuxDeclarationRanges extName (← getRef) (← getRef)
addDeclarationRangesFromSyntax extName (← getRef)
catch e =>
throwError m!"\
Failed to generate an 'ext' theorem for '{MessageData.ofConstName structName}': {e.toMessageData}"
Expand Down Expand Up @@ -161,7 +161,7 @@ def realizeExtIffTheorem (extName : Name) : Elab.Command.CommandElabM Name := do
-- Only declarations in a namespace can be protected:
unless extIffName.isAtomic do
modifyEnv fun env => addProtected env extIffName
addAuxDeclarationRanges extName (← getRef) (← getRef)
addDeclarationRangesFromSyntax extName (← getRef)
catch e =>
throwError m!"\
Failed to generate an 'ext_iff' theorem from '{MessageData.ofConstName extName}': {e.toMessageData}\n\
Expand Down

0 comments on commit d8d9203

Please sign in to comment.