Skip to content

Commit a5c7037

Browse files
committed
feat(Data/Finsupp): add notation (#6367)
This PR provides `fun₀ | 3 => a | 7 => b` notation for `Finsupp`, which desugars to `Finsupp.update` and `Finsupp.single`, in the same way that `{a, b}` desugars to `insert` and `singleton`. It also provides a matching delaborator: as an example of the effect of this, `Finsupp.single_add` now shows as: ```lean Finsupp.single_add.{u_2, u_1} {α : Type u_1} {M : Type u_2} [inst✝ : AddZeroClass M] (a : α) (b₁ b₂ : M) : (fun₀ | a => b₁ + b₂) = (fun₀ | a => b₁) + fun₀ | a => b₂ ``` Finally, it provides a `Repr` instance; though this is somewhat misleading as the syntax that is produced by `Repr` isn't actually computable so can't round-trip! [Discussed on Zulip here](https://leanprover.zulipchat.com/#narrow/stream/113488-general/topic/Notation.20for.20.60Finsupp.60/near/381919585).
1 parent 2677935 commit a5c7037

File tree

4 files changed

+89
-0
lines changed

4 files changed

+89
-0
lines changed

Mathlib.lean

Lines changed: 1 addition & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -1437,6 +1437,7 @@ import Mathlib.Data.Finsupp.Interval
14371437
import Mathlib.Data.Finsupp.Lex
14381438
import Mathlib.Data.Finsupp.Multiset
14391439
import Mathlib.Data.Finsupp.NeLocus
1440+
import Mathlib.Data.Finsupp.Notation
14401441
import Mathlib.Data.Finsupp.Order
14411442
import Mathlib.Data.Finsupp.Pointwise
14421443
import Mathlib.Data.Finsupp.Pwo

Mathlib/Data/Finsupp/Basic.lean

Lines changed: 1 addition & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -7,6 +7,7 @@ import Mathlib.Algebra.BigOperators.Finsupp
77
import Mathlib.Algebra.Hom.GroupAction
88
import Mathlib.Algebra.Regular.SMul
99
import Mathlib.Data.Finset.Preimage
10+
import Mathlib.Data.Finsupp.Notation
1011
import Mathlib.Data.Rat.BigOperators
1112
import Mathlib.Data.Set.Countable
1213

Mathlib/Data/Finsupp/Notation.lean

Lines changed: 71 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -0,0 +1,71 @@
1+
/-
2+
Copyright (c) 2023 Eric Wieser. All rights reserved.
3+
Released under Apache 2.0 license as described in the file LICENSE.
4+
Authors: Eric Wieser
5+
-/
6+
import Mathlib.Data.Finsupp.Defs
7+
8+
/-!
9+
# Notation for `Finsupp`
10+
11+
This file provides `fun₀ | 3 => a | 7 => b` notation for `Finsupp`, which desugars to
12+
`Finsupp.update` and `Finsupp.single`, in the same way that `{a, b}` desugars to `insert` and
13+
`singleton`.
14+
-/
15+
16+
namespace Finsupp
17+
18+
open Lean
19+
open Lean.Parser
20+
open Lean.Parser.Term
21+
22+
/-- A variant of `Lean.Parser.Term.matchAlts` with less line wrapping. -/
23+
def fun₀.matchAlts : Parser :=
24+
leading_parser withPosition $ ppRealGroup <| many1Indent (ppSpace >> ppGroup matchAlt)
25+
26+
@[term_parser, inherit_doc Finsupp]
27+
def fun₀ := leading_parser:maxPrec
28+
ppAllowUngrouped >> unicodeSymbol "λ₀" "fun₀" >> fun₀.matchAlts
29+
30+
macro_rules
31+
| `(term| fun₀ $x:matchAlt*) => do
32+
let mut stx : Term ← `(0)
33+
let mut fst : Bool := true
34+
for xi in x do
35+
for xii in (← Elab.Term.expandMatchAlt xi) do
36+
match xii with
37+
| `(matchAltExpr| | $pat => $val) =>
38+
if fst then
39+
stx ← `(Finsupp.single $pat $val)
40+
else
41+
stx ← `(Finsupp.update $stx $pat $val)
42+
fst := false
43+
| _ => Macro.throwUnsupported
44+
pure stx
45+
46+
/-- Unexpander for the `fun₀ | i => x` notation. -/
47+
@[app_unexpander Finsupp.single]
48+
def singleUnexpander : Lean.PrettyPrinter.Unexpander
49+
| `($_ $pat $val) => `(fun₀ | $pat => $val)
50+
| _ => throw ()
51+
52+
/-- Unexpander for the `fun₀ | i => x` notation. -/
53+
@[app_unexpander Finsupp.update]
54+
def updateUnexpander : Lean.PrettyPrinter.Unexpander
55+
| `($_ $f $pat $val) => match f with
56+
| `(fun₀ $xs:matchAlt*) => `(fun₀ $xs:matchAlt* | $pat => $val)
57+
| _ => throw ()
58+
| _ => throw ()
59+
60+
/-- Display `Finsupp` using `fun₀` notation. -/
61+
unsafe instance {α β} [Repr α] [Repr β] [Zero β] : Repr (α →₀ β) where
62+
reprPrec f p :=
63+
if f.support.card = 0 then
64+
"0"
65+
else
66+
let ret := "fun₀" ++
67+
Std.Format.join (f.support.val.unquot.map <|
68+
fun a => " | " ++ repr a ++ " => " ++ repr (f a))
69+
if p ≥ leadPrec then Format.paren ret else ret
70+
71+
end Finsupp

test/finsupp_notation.lean

Lines changed: 16 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -0,0 +1,16 @@
1+
import Mathlib.Data.Finsupp.Notation
2+
3+
example : (fun₀ | 1 => 3) 1 = 3 :=
4+
by simp
5+
6+
example : (fun₀ | 1 | 2 | 3 => 3 | 3 => 4) 1 = 3 :=
7+
by simp
8+
example : (fun₀ | 1 | 2 | 3 => 3 | 3 => 4) 2 = 3 :=
9+
by simp
10+
example : (fun₀ | 1 | 2 | 3 => 3 | 3 => 4) 3 = 4 :=
11+
by simp
12+
13+
#eval show Lean.MetaM Unit from
14+
guard <|
15+
reprStr (Finsupp.mk {1, 2} (fun | 1 | 2 => 3 | _ => 0) (fun x => by aesop))
16+
= "fun₀ | 1 => 3 | 2 => 3"

0 commit comments

Comments
 (0)