Skip to content

Commit ce5b414

Browse files
committed
feat: allow structures to have non-bracketed binders
1 parent 9b9dd85 commit ce5b414

File tree

6 files changed

+49
-38
lines changed

6 files changed

+49
-38
lines changed

src/Lean/Elab/Structure.lean

Lines changed: 19 additions & 21 deletions
Original file line numberDiff line numberDiff line change
@@ -30,7 +30,7 @@ namespace Structure
3030

3131
/-! Recall that the `structure command syntax is
3232
```
33-
leading_parser (structureTk <|> classTk) >> declId >> many Term.bracketedBinder >> Term.optType >> optional «extends» >> optional (" := " >> optional structCtor >> structFields)
33+
leading_parser (structureTk <|> classTk) >> declId >> optDeclSig >> optional «extends» >> optional (" := " >> optional structCtor >> structFields)
3434
```
3535
-/
3636

@@ -206,10 +206,10 @@ private def expandCtor (structStx : Syntax) (structModifiers : Modifiers) (struc
206206
let ref := structStx[1].mkSynthetic
207207
addDeclarationRangesFromSyntax declName ref
208208
pure { ref, declId := ref, modifiers := default, declName }
209-
if structStx[5].isNone then
209+
if structStx[4].isNone then
210210
useDefault
211211
else
212-
let optCtor := structStx[5][1]
212+
let optCtor := structStx[4][1]
213213
if optCtor.isNone then
214214
useDefault
215215
else
@@ -278,12 +278,12 @@ def structFields := leading_parser many (structExplicitBinder <|> struct
278278
```
279279
-/
280280
private def expandFields (structStx : Syntax) (structModifiers : Modifiers) (structDeclName : Name) : TermElabM (Array StructFieldView) := do
281-
if structStx[5][0].isToken ":=" then
281+
if structStx[4][0].isToken ":=" then
282282
-- https://github.com/leanprover/lean4/issues/5236
283283
let cmd := if structStx[0].getKind == ``Parser.Command.classTk then "class" else "structure"
284-
withRef structStx[0] <| Linter.logLintIf Linter.linter.deprecated structStx[5][0]
284+
withRef structStx[0] <| Linter.logLintIf Linter.linter.deprecated structStx[4][0]
285285
s!"{cmd} ... :=' has been deprecated in favor of '{cmd} ... where'."
286-
let fieldBinders := if structStx[5].isNone then #[] else structStx[5][2][0].getArgs
286+
let fieldBinders := if structStx[4].isNone then #[] else structStx[4][2][0].getArgs
287287
fieldBinders.foldlM (init := #[]) fun (views : Array StructFieldView) fieldBinder => withRef fieldBinder do
288288
let mut fieldBinder := fieldBinder
289289
if fieldBinder.getKind == ``Parser.Command.structSimpleBinder then
@@ -342,14 +342,12 @@ private def expandFields (structStx : Syntax) (structModifiers : Modifiers) (str
342342
}
343343

344344
/-
345-
leading_parser (structureTk <|> classTk) >> declId >> many Term.bracketedBinder >> Term.optType >> optional «extends» >>
345+
leading_parser (structureTk <|> classTk) >> declId >> optDeclSig >> optional «extends» >>
346346
optional (("where" <|> ":=") >> optional structCtor >> structFields) >> optDeriving
347347
348348
where
349349
def structParent := leading_parser optional (atomic (ident >> " : ")) >> termParser
350-
def «extends» := leading_parser " extends " >> sepBy1 structParent ", "
351-
def typeSpec := leading_parser " : " >> termParser
352-
def optType : Parser := optional typeSpec
350+
def «extends» := leading_parser " extends " >> sepBy1 structParent ", " >> optType
353351
354352
def structFields := leading_parser many (structExplicitBinder <|> structImplicitBinder <|> structInstBinder)
355353
def structCtor := leading_parser try (declModifiers >> ident >> " :: ")
@@ -361,24 +359,24 @@ def structureSyntaxToView (modifiers : Modifiers) (stx : Syntax) : TermElabM Str
361359
let declId := stx[1]
362360
let ⟨name, declName, levelNames⟩ ← Term.expandDeclId (← getCurrNamespace) (← Term.getLevelNames) declId modifiers
363361
addDeclarationRangesForBuiltin declName modifiers.stx stx
364-
let binders := stx[2]
365-
let (optType, exts) ←
362+
let (binders, type?) := expandOptDeclSig stx[2]
363+
let exts := stx[3]
364+
let type? ←
366365
-- Compatibility mode for `structure S extends P : Type` syntax
367-
if stx[3].isNone && !stx[4].isNone && !stx[4][0][2].isNone then
368-
logWarningAt stx[4][0][2][0] "\
366+
if type?.isNone && !exts.isNone && !exts[0][2].isNone then
367+
logWarningAt exts[0][2][0] "\
369368
The syntax is now 'structure S : Type extends P' rather than 'structure S extends P : Type'.\n\n\
370369
The purpose of this change is to accommodate 'structure S extends toP : P' syntax for naming parent projections."
371-
pure (stx[4][0][2], stx[4])
370+
pure (some exts[0][2][0][1])
372371
else
373-
if !stx[4].isNone && !stx[4][0][2].isNone then
374-
logErrorAt stx[4][0][2][0] "\
372+
if !exts.isNone && !exts[0][2].isNone then
373+
logErrorAt exts[0][2][0] "\
375374
Unexpected additional resulting type. \
376375
The syntax is now 'structure S : Type extends P' rather than 'structure S extends P : Type'.\n\n\
377376
The purpose of this change is to accommodate 'structure S extends toP : P' syntax for naming parent projections."
378-
pure (stx[3], stx[4])
379-
let parents ← expandParents exts
380-
let derivingClasses ← getOptDerivingClasses stx[6]
381-
let type? := if optType.isNone then none else some optType[0][1]
377+
pure type?
378+
let parents ← expandParents exts
379+
let derivingClasses ← getOptDerivingClasses stx[5]
382380
let ctor ← expandCtor stx modifiers declName
383381
let fields ← expandFields stx modifiers declName
384382
fields.forM fun field => do

src/Lean/Linter/MissingDocs.lean

Lines changed: 2 additions & 2 deletions
Original file line numberDiff line numberDiff line change
@@ -149,7 +149,7 @@ def checkDecl : SimpleHandler := fun stx => do
149149
if declModifiersPubNoDoc head then -- no doc string
150150
lintField rest[1][0] stx[1] "computed field"
151151
else if rest.getKind == ``«structure» then
152-
unless rest[5][2].isNone do
152+
unless rest[4][2].isNone do
153153
let redecls : Std.HashSet String.Pos :=
154154
(← get).infoState.trees.foldl (init := {}) fun s tree =>
155155
tree.foldInfo (init := s) fun _ info s =>
@@ -163,7 +163,7 @@ def checkDecl : SimpleHandler := fun stx => do
163163
if let some range := stx.getRange? then
164164
if redecls.contains range.start then return
165165
lintField parent stx "public field"
166-
for stx in rest[5][2][0].getArgs do
166+
for stx in rest[4][2][0].getArgs do
167167
let head := stx[0]
168168
if declModifiersPubNoDoc head then
169169
if stx.getKind == ``structSimpleBinder then

src/Lean/Linter/UnusedVariables.lean

Lines changed: 7 additions & 9 deletions
Original file line numberDiff line numberDiff line change
@@ -189,16 +189,14 @@ builtin_initialize
189189
builtin_initialize addBuiltinUnusedVariablesIgnoreFn (fun _ stack _ =>
190190
stack.matches [`null, none, `null, ``Lean.Parser.Command.variable])
191191

192-
/-- `structure Foo where unused : Nat` -/
193-
builtin_initialize addBuiltinUnusedVariablesIgnoreFn (fun _ stack _ =>
194-
stack.matches [`null, none, `null, ``Lean.Parser.Command.structure])
195-
196-
/-- `inductive Foo where | unused : Foo` -/
192+
/--
193+
* `structure Foo (unused : Nat)`
194+
* `inductive Foo (unused : Foo)`
195+
-/
197196
builtin_initialize addBuiltinUnusedVariablesIgnoreFn (fun _ stack _ =>
198-
stack.matches [`null, none, `null, none, ``Lean.Parser.Command.inductive] &&
199-
(stack[3]? |>.any fun (stx, pos) =>
200-
pos == 0 &&
201-
[``Lean.Parser.Command.optDeclSig, ``Lean.Parser.Command.declSig].any (stx.isOfKind ·)))
197+
stack.matches [`null, none, `null, ``Lean.Parser.Command.optDeclSig, none] &&
198+
(stack[4]? |>.any fun (stx, _) =>
199+
[``Lean.Parser.Command.structure, ``Lean.Parser.Command.inductive].any (stx.isOfKind ·)))
202200

203201
/--
204202
* `structure Foo where foo (unused : Nat) : Nat`

src/Lean/Parser/Command.lean

Lines changed: 1 addition & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -236,7 +236,7 @@ def «structure» := leading_parser
236236
(structureTk <|> classTk) >>
237237
-- Note: no error recovery here due to clashing with the `class abbrev` syntax
238238
declId >>
239-
ppIndent (many (ppSpace >> Term.bracketedBinder) >> Term.optType >> optional «extends») >>
239+
ppIndent (optDeclSig >> optional «extends») >>
240240
optional ((symbol " := " <|> " where ") >> optional structCtor >> structFields) >>
241241
optDeriving
242242
@[builtin_command_parser] def declaration := leading_parser

tests/lean/1705.lean.expected.out

Lines changed: 2 additions & 5 deletions
Original file line numberDiff line numberDiff line change
@@ -1,5 +1,2 @@
1-
1705.lean:1:13-1:16: error: function expected at
2-
T
3-
term has type
4-
?m
5-
1705.lean:1:18-1:19: error: unexpected identifier; expected command
1+
1705.lean:1:22-1:23: error: type expected, got
2+
(R : ?m)
Lines changed: 18 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -0,0 +1,18 @@
1+
structure Δ n where
2+
val : Fin (n + 1)
3+
deriving Inhabited
4+
5+
/--
6+
info: Δ (n : Nat) : Type
7+
-/
8+
#guard_msgs in
9+
#check Δ
10+
11+
structure Prod₃ α β γ : Type u extends toProd : α × β where
12+
trd : (γ : Type u)
13+
14+
/--
15+
info: Prod₃.{u} (α β γ : Type u) : Type u
16+
-/
17+
#guard_msgs in
18+
#check Prod₃

0 commit comments

Comments
 (0)