Skip to content

Commit 7f885ee

Browse files
committed
feat: pretty print as ↥S instead of { x // x ∈ S } for SetLike (#7927)
Adds a delaborator detecting `{ x // x ∈ S }` where the membership uses `SetLike.instMembership`, since then this matches the `CoeSort` instance for `SetLike`.
1 parent 65b7d16 commit 7f885ee

File tree

2 files changed

+32
-1
lines changed

2 files changed

+32
-1
lines changed

Mathlib/Data/SetLike/Basic.lean

Lines changed: 18 additions & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -108,12 +108,29 @@ variable {A : Type*} {B : Type*} [i : SetLike A B]
108108

109109
instance : CoeTC A (Set B) where coe := SetLike.coe
110110

111-
instance (priority := 100) : Membership B A :=
111+
instance (priority := 100) instMembership : Membership B A :=
112112
fun x p => x ∈ (p : Set B)⟩
113113

114114
instance (priority := 100) : CoeSort A (Type _) :=
115115
fun p => { x : B // x ∈ p }⟩
116116

117+
section Delab
118+
open Lean PrettyPrinter.Delaborator SubExpr
119+
120+
/-- For terms that match the `CoeSort` instance's body, pretty print as `↥S`
121+
rather than as `{ x // x ∈ S }`. The discriminating feature is that membership
122+
uses the `SetLike.instMembership` instance. -/
123+
@[delab app.Subtype]
124+
def delabSubtypeSetLike : Delab := whenPPOption getPPNotation do
125+
let #[_, .lam n _ body _] := (← getExpr).getAppArgs | failure
126+
guard <| body.isAppOf ``Membership.mem
127+
let #[_, _, inst, .bvar 0, _] := body.getAppArgs | failure
128+
guard <| inst.isAppOfArity ``instMembership 3
129+
let S ← withAppArg <| withBindingBody n <| withNaryArg 4 delab
130+
`(↥$S)
131+
132+
end Delab
133+
117134
variable (p q : A)
118135

119136
@[simp, norm_cast]

test/set_like.lean

Lines changed: 14 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -5,6 +5,20 @@ import Mathlib.FieldTheory.Subfield
55

66
set_option autoImplicit true
77

8+
section Delab
9+
variable {M : Type u} [Monoid M] (S S' : Submonoid M)
10+
11+
/-- info: ↥S → ↥S' : Type u -/
12+
#guard_msgs in #check S → S'
13+
14+
/-- info: ↥S : Type u -/
15+
#guard_msgs in #check {x // x ∈ S}
16+
17+
/-- info: { x // 1 * x ∈ S } : Type u -/
18+
#guard_msgs in #check {x // 1 * x ∈ S}
19+
20+
end Delab
21+
822
example [Ring R] (S : Subring R) (hx : x ∈ S) (hy : y ∈ S) (hz : z ∈ S) (n m : ℕ) :
923
n • x ^ 3 - 2 • y + z ^ m ∈ S := by
1024
aesop

0 commit comments

Comments
 (0)