-
Notifications
You must be signed in to change notification settings - Fork 449
Commit
This commit does not belong to any branch on this repository, and may belong to a fork outside of the repository.
feat: allow explicit mode with field notation (#5528)
Now one can write `@x.f`, `@(x).f`, `@x.1`, `@(x).1`, and so on. This fixes an issue where structure instance update notation (like `{x with a := a'}`) could fail if the field `a` had a type with implicit, optional, or auto parameters. Closes #5406
- Loading branch information
Showing
3 changed files
with
114 additions
and
10 deletions.
There are no files selected for viewing
This file contains bidirectional Unicode text that may be interpreted or compiled differently than what appears below. To review, open the file in an editor that reveals hidden Unicode characters.
Learn more about bidirectional Unicode characters
This file contains bidirectional Unicode text that may be interpreted or compiled differently than what appears below. To review, open the file in an editor that reveals hidden Unicode characters.
Learn more about bidirectional Unicode characters
This file contains bidirectional Unicode text that may be interpreted or compiled differently than what appears below. To review, open the file in an editor that reveals hidden Unicode characters.
Learn more about bidirectional Unicode characters
Original file line number | Diff line number | Diff line change |
---|---|---|
@@ -0,0 +1,99 @@ | ||
/-! | ||
# Add explicit mode projections | ||
This is to fix a bug where structure instance notation was not working when there were opt params. | ||
-/ | ||
|
||
/-! | ||
Motivating issue from /~https://github.com/leanprover/lean4/issues/5406 | ||
The `example` had an elaboration error because the structure instance was expanding to `{b := m.b}`. | ||
Now it expands to `{b := @m.b}`. | ||
-/ | ||
structure Methods where | ||
b : Nat → (opt : Nat := 42) → Nat | ||
|
||
example (m : Methods) : Methods := { m with } | ||
|
||
/-- info: fun m => { b := @Methods.b m } : Methods → Methods -/ | ||
#guard_msgs in #check fun (m : Methods) => { m with } | ||
|
||
|
||
/-! | ||
Tests of explicit mode, with and without arguments. | ||
-/ | ||
|
||
/-- info: fun m => @Methods.b m : Methods → Nat → optParam Nat 42 → Nat -/ | ||
#guard_msgs in #check fun (m : Methods) => @m.b | ||
|
||
/-- info: fun m => @Methods.b m : Methods → Nat → optParam Nat 42 → Nat -/ | ||
#guard_msgs in #check fun (m : Methods) => @(m).b | ||
|
||
/-- info: fun m => @Methods.b m : Methods → Nat → optParam Nat 42 → Nat -/ | ||
#guard_msgs in #check fun (m : Methods) => @m.1 | ||
|
||
/-- info: fun m => @Methods.b m : Methods → Nat → optParam Nat 42 → Nat -/ | ||
#guard_msgs in #check fun (m : Methods) => @(m).1 | ||
|
||
/-- info: fun m => @Methods.b m 1 : Methods → optParam Nat 42 → Nat -/ | ||
#guard_msgs in #check fun (m : Methods) => @m.b 1 | ||
|
||
/-- info: fun m => @Methods.b m 1 : Methods → optParam Nat 42 → Nat -/ | ||
#guard_msgs in #check fun (m : Methods) => @(m).b 1 | ||
|
||
/-- info: fun m => @Methods.b m 1 : Methods → optParam Nat 42 → Nat -/ | ||
#guard_msgs in #check fun (m : Methods) => @m.1 1 | ||
|
||
/-- info: fun m => @Methods.b m 1 : Methods → optParam Nat 42 → Nat -/ | ||
#guard_msgs in #check fun (m : Methods) => @(m).1 1 | ||
|
||
|
||
/-! | ||
Tests of explicit mode with class instances. The type parameter remains implicit. | ||
We need this so that structure instances work properly. | ||
-/ | ||
|
||
class C (α : Type) [Inhabited α] where | ||
f (x : α := default) : α | ||
|
||
/-- info: fun inst => C.f : C Nat → Nat -/ | ||
#guard_msgs in #check fun (inst : C Nat) => inst.f | ||
|
||
/-- info: fun inst => @C.f Nat instInhabitedNat inst : C Nat → optParam Nat default → Nat -/ | ||
#guard_msgs in #check fun (inst : C Nat) => @inst.f | ||
|
||
/-- info: fun inst => @C.f Nat instInhabitedNat inst : C Nat → optParam Nat default → Nat -/ | ||
#guard_msgs in #check fun (inst : C Nat) => @C.f _ _ inst | ||
|
||
/-- info: fun inst => { f := @C.f Nat instInhabitedNat inst } : C Nat → C Nat -/ | ||
#guard_msgs in #check fun (inst : C Nat) => { inst with } | ||
|
||
|
||
/-! | ||
Tests of deeper updates and accesses. | ||
-/ | ||
|
||
structure A (α : Type) where | ||
x : α | ||
structure B (α β : Type) extends A α where | ||
y : β | ||
|
||
/-- info: fun α β x' b => { x := x', y := b.y } : (α β : Type) → α → B α β → B α β -/ | ||
#guard_msgs in #check fun (α β) (x' : α) (b : B α β) => {b with x := x'} | ||
|
||
/-- info: fun α β b => b.x : (α β : Type) → B α β → α -/ | ||
#guard_msgs in #check fun (α β) (b : B α β) => @b.toA.x | ||
|
||
/-- info: fun α β b => b.x : (α β : Type) → B α β → α -/ | ||
#guard_msgs in #check fun (α β) (b : B α β) => @b.x | ||
|
||
|
||
/-! | ||
Tests of implicit arguments in updates. | ||
-/ | ||
|
||
structure I where | ||
f : {_ : Nat} → Nat | ||
|
||
-- used to give `fun i ↦ ?m.369 i : I → I` | ||
/-- info: fun i => { f := @I.f i } : I → I -/ | ||
#guard_msgs in #check fun (i : I) => {i with} |