Skip to content

Commit

Permalink
doc: split the docstring of LocalContext (#6340)
Browse files Browse the repository at this point in the history
This results in better hovers in VSCode, without duplicating information
in a way that might go out of sync.
  • Loading branch information
eric-wieser authored Dec 15, 2024
1 parent 80fb404 commit a8dc619
Showing 1 changed file with 10 additions and 7 deletions.
17 changes: 10 additions & 7 deletions src/Lean/LocalContext.lean
Original file line number Diff line number Diff line change
Expand Up @@ -42,16 +42,15 @@ inductive LocalDeclKind
| auxDecl
deriving Inhabited, Repr, DecidableEq, Hashable

/-- A declaration for a LocalContext. This is used to register which free variables are in scope.
Each declaration comes with
- `index` the position of the decl in the local context
- `fvarId` the unique id of the free variables
- `userName` the pretty-printable name of the variable
- `type` the type.
A `cdecl` is a local variable, a `ldecl` is a let-bound free variable with a `value : Expr`.
/-- A declaration for a `LocalContext`. This is used to register which free variables are in scope.
See `LocalDecl.index`, `LocalDecl.fvarId`, `LocalDecl.userName`, `LocalDecl.type` for accessors for
arguments common to both constructors.
-/
inductive LocalDecl where
/-- A local variable. -/
| cdecl (index : Nat) (fvarId : FVarId) (userName : Name) (type : Expr) (bi : BinderInfo) (kind : LocalDeclKind)
/-- A let-bound free variable, with a `value : Expr`. -/
| ldecl (index : Nat) (fvarId : FVarId) (userName : Name) (type : Expr) (value : Expr) (nonDep : Bool) (kind : LocalDeclKind)
deriving Inhabited

Expand All @@ -71,6 +70,7 @@ def isLet : LocalDecl → Bool
| cdecl .. => false
| ldecl .. => true

/-- The position of the decl in the local context. -/
def index : LocalDecl → Nat
| cdecl (index := i) .. => i
| ldecl (index := i) .. => i
Expand All @@ -79,14 +79,17 @@ def setIndex : LocalDecl → Nat → LocalDecl
| cdecl _ id n t bi k, idx => cdecl idx id n t bi k
| ldecl _ id n t v nd k, idx => ldecl idx id n t v nd k

/-- The unique id of the free variable. -/
def fvarId : LocalDecl → FVarId
| cdecl (fvarId := id) .. => id
| ldecl (fvarId := id) .. => id

/-- The pretty-printable name of the variable. -/
def userName : LocalDecl → Name
| cdecl (userName := n) .. => n
| ldecl (userName := n) .. => n

/-- The type of the variable. -/
def type : LocalDecl → Expr
| cdecl (type := t) .. => t
| ldecl (type := t) .. => t
Expand Down

0 comments on commit a8dc619

Please sign in to comment.