Skip to content

Commit

Permalink
feat: add structInstFieldDecl syntax category (#6265)
Browse files Browse the repository at this point in the history
This PR is preparation for changes to structure instance notation in
#6165. It adds a syntax category that will be used for field syntax.
  • Loading branch information
kmill authored Nov 30, 2024
1 parent 27df5e9 commit f3f0045
Show file tree
Hide file tree
Showing 2 changed files with 12 additions and 0 deletions.
5 changes: 5 additions & 0 deletions src/Init/Notation.lean
Original file line number Diff line number Diff line change
Expand Up @@ -48,6 +48,11 @@ def tactic : Category := {}
For example, `let x ← e` is a `doElem`, and a `do` block consists of a list of `doElem`s. -/
def doElem : Category := {}

/-- `structInstFieldDecl` is the syntax category for value declarations for fields in structure instance notation.
For example, the `:= 1` and `where a := 3` in `{ x := 1, y where a := 3 }` are in the `structInstFieldDecl` class.
This category is necessary because structure instance notation is recursive due to the `x where ...` field notation. -/
def structInstFieldDecl : Category := {}

/-- `level` is a builtin syntax category for universe levels.
This is the `u` in `Sort u`: it can contain `max` and `imax`, addition with
constants, and variables. -/
Expand Down
7 changes: 7 additions & 0 deletions src/Lean/Parser/Term.lean
Original file line number Diff line number Diff line change
Expand Up @@ -387,6 +387,13 @@ use `(_)` to inhibit this and have it be solved for by unification instead, like
explicitBinder requireType <|> strictImplicitBinder requireType <|>
implicitBinder requireType <|> instBinder

/-
Syntax category for structure instance notation fields.
Does not initialize `registerBuiltinDynamicParserAttribute` since this category is not meant to be user-extensible.
-/
builtin_initialize
registerBuiltinParserAttribute `builtin_structInstFieldDecl_parser ``Category.structInstFieldDecl

/-
It is feasible to support dependent arrows such as `{α} → α → α` without sacrificing the quality of the error messages for the longer case.
`{α} → α → α` would be short for `{α : Type} → α → α`
Expand Down

0 comments on commit f3f0045

Please sign in to comment.