Skip to content

Commit

Permalink
feat: add option pp.mvars.delayed
Browse files Browse the repository at this point in the history
Where before we had
```lean
#check fun x : Nat => ?a
-- fun x ↦ ?m.7 x : (x : Nat) → ?m.6 x
```
Now by default we have
```lean
#check fun x : Nat => ?a
-- fun x => ?a : (x : Nat) → ?m.6 x
```
in particular, delayed assignment metavariables such as `?m.7` pretty print using the name of the metavariable they are delayed assigned to, suppressing the arguments used in the delayed assignment (hence `?a` rather than `?a x`). Hovering over `?a` shows `?m.7 x`.

The benefit is that users can see the user-provided name in local contexts.

The original pretty printing behavior can be recovered using `set_option pp.mvars.delayed true`.

This PR also extends the documentation for holes and synthetic holes, with useful high-level technical details about what delayed assignments are.

(This PR supercedes #3494, which has a round-trippable notation for delayed assignments. The notation in this PR is unlikely to round trip, but it is better than the current situation, where delayed assignment metavariables never round trip.)
  • Loading branch information
kmill committed Oct 7, 2024
1 parent 9dac514 commit fe6f13d
Show file tree
Hide file tree
Showing 14 changed files with 102 additions and 26 deletions.
59 changes: 55 additions & 4 deletions src/Lean/Parser/Term.lean
Original file line number Diff line number Diff line change
Expand Up @@ -140,11 +140,62 @@ def optSemicolon (p : Parser) : Parser :=
/-- The universe of propositions. `Prop ≡ Sort 0`. -/
@[builtin_term_parser] def prop := leading_parser
"Prop"
/-- A placeholder term, to be synthesized by unification. -/
/--
A *hole* (or *placeholder term*), which stands for an unknown term that is expected to be inferred based on context.
For example, in `@id _ Nat.zero`, the `_` must be the type of `Nat.zero`, which is `Nat`.
The way this works is that holes create fresh metavariables.
The elaborator is allowed to assign terms to metavariables while it is checking definitional equalities.
This is often known as *unification*.
Normally, all holes must be solved for. However, there are a few contexts where this is not necessary:
* In `match` patterns, holes are catch-all patterns.
* In some tactics, such as `refine'` and `apply`, unsolved-for placeholders become new goals.
Related concept: implicit parameters are automatically filled in with holes during the elaboration process.
See also `?m` syntax (synthetic holes).
-/
@[builtin_term_parser] def hole := leading_parser
"_"
/-- Parses a "synthetic hole", that is, `?foo` or `?_`.
This syntax is used to construct named metavariables. -/
/--
A *synthetic hole* (or *synthetic placeholder*), which stands for an unknown term that should should be synthesized using tactics.
- `?_` creates a fresh metavariable with an auto-generated name.
- `?m` either refers to a pre-existing metavariable named `m` or creates a fresh metavariable with that name.
In particular, the synthetic hole syntax creates "synthetic opaque metavariables",
the same kind of metavariable used to represent goals in the tactic state.
Synthetic holes are similar to holes in that `_` also creates metavariables,
but synthetic opaque metavariables have some different properties:
- In tactics such as `refine`, only synthetic holes yield new goals.
- When synthetic holes appear under binders, they capture local variables using a more complicated mechanism known as delayed assignment.
- During elaboration, unification will not solve for synthetic opaque metavariables, they are "opaque".
This is to prevent counterintuitive behavior such as disappearing goals.
## Delayed assigned metavariables
This section gives a high-level overview of technical details of synthetic holes, which you should feel free to skip.
When a synthetic hole appears under a binding construct, such as for example `fun (x : α) (y : β) => ?s`,
the system creates a construct called a *delayed assignment*. This consists of
1. A metavariable `?m` of type `α → β → γ` and whole local context is from the local context outside the `fun`,
where `γ` is the type of `?s`. (The type `γ` can depend on `α` and `β`.)
2. A delayed assigment record associating `?m` to `?s` and the variables `#[x, y]` in the local context of `?s`
Then, this function elaborates as `fun (x : α) (y : β) => ?m x y`, where one should understand `x` and `y` here
as being De Bruijn indexes, since Lean uses the locally nameless encoding of lambda calculus.
Once `?s` is fully solved for, in the sense that it can be fully instantiated as a metavariable-free term `e`,
then we can make the assignment `?m := fun (x' : α) (y' : β) => e[x := x', y := y']`.
(Implementation note: in practice, Lean only instantiates full applications `?m x' y'` of delayed assigned metavariables.)
This delayed assignment mechanism is essential to the operation of basic tactics like `intro`,
and a good mental model is that it is a way to "apply" the metavariable `?s` by substituting values in for some of its local variables.
By default, delayed assigned metavariables pretty print with what they are delayed assigned to.
The delayed assigned metavariables themselves can be pretty printed using `set_option pp.mvars.delayed true`.
For more details, see the module docstrings in `Lean.MetavarContext`.
-/
@[builtin_term_parser] def syntheticHole := leading_parser
"?" >> (ident <|> "_")
/--
Expand Down Expand Up @@ -451,7 +502,7 @@ def withAnonymousAntiquot := leading_parser
@[builtin_term_parser] def «trailing_parser» := leading_parser:leadPrec
"trailing_parser" >> optExprPrecedence >> optExprPrecedence >> ppSpace >> termParser

/--
/--
Indicates that an argument to a function marked `@[extern]` is borrowed.
Being borrowed only affects the ABI and runtime behavior of the function when compiled or interpreted. From the perspective of Lean's type system, this annotation has no effect. It similarly has no effect on functions not marked `@[extern]`.
Expand Down
19 changes: 19 additions & 0 deletions src/Lean/PrettyPrinter/Delaborator/Builtins.lean
Original file line number Diff line number Diff line change
Expand Up @@ -570,6 +570,25 @@ def withOverApp (arity : Nat) (x : Delab) : Delab := do
withAnnotateTermInfo x
delabAppCore (n - arity) delabHead (unexpand := false)

@[builtin_delab app]
def delabDelayedAssignedMVar : Delab := whenNotPPOption getPPMVarsDelayed do
let .mvar mvarId := (← getExpr).getAppFn | failure
let some decl ← getDelayedMVarAssignment? mvarId | failure
withOverApp decl.fvars.size do
let args := (← getExpr).getAppArgs
-- Only delaborate using decl.mvarIdPending if the delayed mvar is applied to fvars
guard <| args.all Expr.isFVar
withTypeAscription (cond := ← getPPOption getPPMVarsWithType) do
if ← getPPOption getPPMVars then
let mvarDecl ← decl.mvarIdPending.getDecl
let n :=
match mvarDecl.userName with
| .anonymous => decl.mvarIdPending.name.replacePrefix `_uniq `m
| n => n
`(?$(mkIdent n))
else
`(?_)

/-- State for `delabAppMatch` and helpers. -/
structure AppMatchState where
info : MatcherInfo
Expand Down
6 changes: 6 additions & 0 deletions src/Lean/PrettyPrinter/Delaborator/Options.lean
Original file line number Diff line number Diff line change
Expand Up @@ -95,6 +95,11 @@ register_builtin_option pp.mvars.withType : Bool := {
group := "pp"
descr := "(pretty printer) display metavariables with a type ascription"
}
register_builtin_option pp.mvars.delayed : Bool := {
defValue := false
group := "pp"
descr := "(pretty printer) display delayed assigned metavariables when true, otherwise display what they are assigned to"
}
register_builtin_option pp.beta : Bool := {
defValue := false
group := "pp"
Expand Down Expand Up @@ -244,6 +249,7 @@ def getPPPrivateNames (o : Options) : Bool := o.get pp.privateNames.name (getPPA
def getPPInstantiateMVars (o : Options) : Bool := o.get pp.instantiateMVars.name pp.instantiateMVars.defValue
def getPPMVars (o : Options) : Bool := o.get pp.mvars.name pp.mvars.defValue
def getPPMVarsWithType (o : Options) : Bool := o.get pp.mvars.withType.name pp.mvars.withType.defValue
def getPPMVarsDelayed (o : Options) : Bool := o.get pp.mvars.delayed.name (pp.mvars.delayed.defValue || getPPAll o)
def getPPBeta (o : Options) : Bool := o.get pp.beta.name pp.beta.defValue
def getPPSafeShadowing (o : Options) : Bool := o.get pp.safeShadowing.name pp.safeShadowing.defValue
def getPPProofs (o : Options) : Bool := o.get pp.proofs.name (pp.proofs.defValue || getPPAll o)
Expand Down
2 changes: 1 addition & 1 deletion tests/lean/funParen.lean.expected.out
Original file line number Diff line number Diff line change
@@ -1,4 +1,4 @@
fun α [Repr α] => repr : (α : Type u_1) → [inst : Repr α] → α → Std.Format
fun x y => x : (x : ?m) → ?m x → ?m
funParen.lean:4:12-4:16: error: invalid pattern, constructor or constant marked with '[match_pattern]' expected
fun x => ?m x : (x : ?m) → ?m x
fun x => ?m : (x : ?m) → ?m x
Original file line number Diff line number Diff line change
Expand Up @@ -4,5 +4,5 @@
{"start": {"line": 4, "character": 29}, "end": {"line": 4, "character": 30}},
"contents":
{"value":
"```lean\ninstDecidableTrue : Decidable True\n```\n***\nA placeholder term, to be synthesized by unification. \n***\n*import Init.Core*",
"```lean\ninstDecidableTrue : Decidable True\n```\n***\nA *hole* (or *placeholder term*), which stands for an unknown term that is expected to be inferred based on context.\nFor example, in `@id _ Nat.zero`, the `_` must be the type of `Nat.zero`, which is `Nat`.\n\nThe way this works is that holes create fresh metavariables.\nThe elaborator is allowed to assign terms to metavariables while it is checking definitional equalities.\nThis is often known as *unification*.\n\nNormally, all holes must be solved for. However, there are a few contexts where this is not necessary:\n* In `match` patterns, holes are catch-all patterns.\n* In some tactics, such as `refine'` and `apply`, unsolved-for placeholders become new goals.\n\nRelated concept: implicit parameters are automatically filled in with holes during the elaboration process.\n\nSee also `?m` syntax (synthetic holes).\n\n***\n*import Init.Core*",
"kind": "markdown"}}
4 changes: 2 additions & 2 deletions tests/lean/interactive/hover.lean.expected.out
Original file line number Diff line number Diff line change
Expand Up @@ -507,7 +507,7 @@
"end": {"line": 252, "character": 10}},
"contents":
{"value":
"```lean\nℕ\n```\n***\nA placeholder term, to be synthesized by unification. ",
"```lean\nℕ\n```\n***\nA *hole* (or *placeholder term*), which stands for an unknown term that is expected to be inferred based on context.\nFor example, in `@id _ Nat.zero`, the `_` must be the type of `Nat.zero`, which is `Nat`.\n\nThe way this works is that holes create fresh metavariables.\nThe elaborator is allowed to assign terms to metavariables while it is checking definitional equalities.\nThis is often known as *unification*.\n\nNormally, all holes must be solved for. However, there are a few contexts where this is not necessary:\n* In `match` patterns, holes are catch-all patterns.\n* In some tactics, such as `refine'` and `apply`, unsolved-for placeholders become new goals.\n\nRelated concept: implicit parameters are automatically filled in with holes during the elaboration process.\n\nSee also `?m` syntax (synthetic holes).\n",
"kind": "markdown"}}
{"textDocument": {"uri": "file:///hover.lean"},
"position": {"line": 252, "character": 11}}
Expand Down Expand Up @@ -573,7 +573,7 @@
{"start": {"line": 269, "character": 2}, "end": {"line": 269, "character": 3}},
"contents":
{"value":
"```lean\nlet x :=\n match 0 with\n | x => 0;\nℕ\n```\n***\nA placeholder term, to be synthesized by unification. ",
"```lean\nlet x :=\n match 0 with\n | x => 0;\nℕ\n```\n***\nA *hole* (or *placeholder term*), which stands for an unknown term that is expected to be inferred based on context.\nFor example, in `@id _ Nat.zero`, the `_` must be the type of `Nat.zero`, which is `Nat`.\n\nThe way this works is that holes create fresh metavariables.\nThe elaborator is allowed to assign terms to metavariables while it is checking definitional equalities.\nThis is often known as *unification*.\n\nNormally, all holes must be solved for. However, there are a few contexts where this is not necessary:\n* In `match` patterns, holes are catch-all patterns.\n* In some tactics, such as `refine'` and `apply`, unsolved-for placeholders become new goals.\n\nRelated concept: implicit parameters are automatically filled in with holes during the elaboration process.\n\nSee also `?m` syntax (synthetic holes).\n",
"kind": "markdown"}}
{"textDocument": {"uri": "file:///hover.lean"},
"position": {"line": 272, "character": 4}}
Expand Down
12 changes: 6 additions & 6 deletions tests/lean/interactive/hoverBinderUndescore.lean.expected.out
Original file line number Diff line number Diff line change
Expand Up @@ -4,45 +4,45 @@
{"start": {"line": 1, "character": 5}, "end": {"line": 1, "character": 6}},
"contents":
{"value":
"```lean\nNat\n```\n***\nA placeholder term, to be synthesized by unification. ",
"```lean\nNat\n```\n***\nA *hole* (or *placeholder term*), which stands for an unknown term that is expected to be inferred based on context.\nFor example, in `@id _ Nat.zero`, the `_` must be the type of `Nat.zero`, which is `Nat`.\n\nThe way this works is that holes create fresh metavariables.\nThe elaborator is allowed to assign terms to metavariables while it is checking definitional equalities.\nThis is often known as *unification*.\n\nNormally, all holes must be solved for. However, there are a few contexts where this is not necessary:\n* In `match` patterns, holes are catch-all patterns.\n* In some tactics, such as `refine'` and `apply`, unsolved-for placeholders become new goals.\n\nRelated concept: implicit parameters are automatically filled in with holes during the elaboration process.\n\nSee also `?m` syntax (synthetic holes).\n",
"kind": "markdown"}}
{"textDocument": {"uri": "file:///hoverBinderUndescore.lean"},
"position": {"line": 1, "character": 7}}
{"range":
{"start": {"line": 1, "character": 7}, "end": {"line": 1, "character": 8}},
"contents":
{"value":
"```lean\nBool\n```\n***\nA placeholder term, to be synthesized by unification. ",
"```lean\nBool\n```\n***\nA *hole* (or *placeholder term*), which stands for an unknown term that is expected to be inferred based on context.\nFor example, in `@id _ Nat.zero`, the `_` must be the type of `Nat.zero`, which is `Nat`.\n\nThe way this works is that holes create fresh metavariables.\nThe elaborator is allowed to assign terms to metavariables while it is checking definitional equalities.\nThis is often known as *unification*.\n\nNormally, all holes must be solved for. However, there are a few contexts where this is not necessary:\n* In `match` patterns, holes are catch-all patterns.\n* In some tactics, such as `refine'` and `apply`, unsolved-for placeholders become new goals.\n\nRelated concept: implicit parameters are automatically filled in with holes during the elaboration process.\n\nSee also `?m` syntax (synthetic holes).\n",
"kind": "markdown"}}
{"textDocument": {"uri": "file:///hoverBinderUndescore.lean"},
"position": {"line": 6, "character": 6}}
{"range":
{"start": {"line": 6, "character": 6}, "end": {"line": 6, "character": 7}},
"contents":
{"value":
"```lean\nNat\n```\n***\nA placeholder term, to be synthesized by unification. ",
"```lean\nNat\n```\n***\nA *hole* (or *placeholder term*), which stands for an unknown term that is expected to be inferred based on context.\nFor example, in `@id _ Nat.zero`, the `_` must be the type of `Nat.zero`, which is `Nat`.\n\nThe way this works is that holes create fresh metavariables.\nThe elaborator is allowed to assign terms to metavariables while it is checking definitional equalities.\nThis is often known as *unification*.\n\nNormally, all holes must be solved for. However, there are a few contexts where this is not necessary:\n* In `match` patterns, holes are catch-all patterns.\n* In some tactics, such as `refine'` and `apply`, unsolved-for placeholders become new goals.\n\nRelated concept: implicit parameters are automatically filled in with holes during the elaboration process.\n\nSee also `?m` syntax (synthetic holes).\n",
"kind": "markdown"}}
{"textDocument": {"uri": "file:///hoverBinderUndescore.lean"},
"position": {"line": 6, "character": 8}}
{"range":
{"start": {"line": 6, "character": 8}, "end": {"line": 6, "character": 9}},
"contents":
{"value":
"```lean\nBool\n```\n***\nA placeholder term, to be synthesized by unification. ",
"```lean\nBool\n```\n***\nA *hole* (or *placeholder term*), which stands for an unknown term that is expected to be inferred based on context.\nFor example, in `@id _ Nat.zero`, the `_` must be the type of `Nat.zero`, which is `Nat`.\n\nThe way this works is that holes create fresh metavariables.\nThe elaborator is allowed to assign terms to metavariables while it is checking definitional equalities.\nThis is often known as *unification*.\n\nNormally, all holes must be solved for. However, there are a few contexts where this is not necessary:\n* In `match` patterns, holes are catch-all patterns.\n* In some tactics, such as `refine'` and `apply`, unsolved-for placeholders become new goals.\n\nRelated concept: implicit parameters are automatically filled in with holes during the elaboration process.\n\nSee also `?m` syntax (synthetic holes).\n",
"kind": "markdown"}}
{"textDocument": {"uri": "file:///hoverBinderUndescore.lean"},
"position": {"line": 11, "character": 6}}
{"range":
{"start": {"line": 11, "character": 6}, "end": {"line": 11, "character": 7}},
"contents":
{"value":
"```lean\nNat\n```\n***\nA placeholder term, to be synthesized by unification. ",
"```lean\nNat\n```\n***\nA *hole* (or *placeholder term*), which stands for an unknown term that is expected to be inferred based on context.\nFor example, in `@id _ Nat.zero`, the `_` must be the type of `Nat.zero`, which is `Nat`.\n\nThe way this works is that holes create fresh metavariables.\nThe elaborator is allowed to assign terms to metavariables while it is checking definitional equalities.\nThis is often known as *unification*.\n\nNormally, all holes must be solved for. However, there are a few contexts where this is not necessary:\n* In `match` patterns, holes are catch-all patterns.\n* In some tactics, such as `refine'` and `apply`, unsolved-for placeholders become new goals.\n\nRelated concept: implicit parameters are automatically filled in with holes during the elaboration process.\n\nSee also `?m` syntax (synthetic holes).\n",
"kind": "markdown"}}
{"textDocument": {"uri": "file:///hoverBinderUndescore.lean"},
"position": {"line": 11, "character": 8}}
{"range":
{"start": {"line": 11, "character": 8}, "end": {"line": 11, "character": 9}},
"contents":
{"value":
"```lean\nBool\n```\n***\nA placeholder term, to be synthesized by unification. ",
"```lean\nBool\n```\n***\nA *hole* (or *placeholder term*), which stands for an unknown term that is expected to be inferred based on context.\nFor example, in `@id _ Nat.zero`, the `_` must be the type of `Nat.zero`, which is `Nat`.\n\nThe way this works is that holes create fresh metavariables.\nThe elaborator is allowed to assign terms to metavariables while it is checking definitional equalities.\nThis is often known as *unification*.\n\nNormally, all holes must be solved for. However, there are a few contexts where this is not necessary:\n* In `match` patterns, holes are catch-all patterns.\n* In some tactics, such as `refine'` and `apply`, unsolved-for placeholders become new goals.\n\nRelated concept: implicit parameters are automatically filled in with holes during the elaboration process.\n\nSee also `?m` syntax (synthetic holes).\n",
"kind": "markdown"}}
Loading

0 comments on commit fe6f13d

Please sign in to comment.