Skip to content

Commit ffa1e9e

Browse files
authored
doc: add recommended spellings for many term notations (#6886)
This PR adds recommended spellings for many notations defined in Lean core, using the `recommended_spelling` command from #6869.
1 parent 030daff commit ffa1e9e

File tree

14 files changed

+140
-6
lines changed

14 files changed

+140
-6
lines changed

src/Init/Control/Basic.lean

+10
Original file line numberDiff line numberDiff line change
@@ -50,6 +50,8 @@ def Functor.mapRev {f : Type u → Type v} [Functor f] {α β : Type u} : f α
5050

5151
infixr:100 " <&> " => Functor.mapRev
5252

53+
recommended_spelling "mapRev" for "<&>" in [Functor.mapRev, «term_<&>_»]
54+
5355
@[always_inline, inline]
5456
def Functor.discard {f : Type u → Type v} {α : Type u} [Functor f] (x : f α) : f PUnit :=
5557
Functor.mapConst PUnit.unit x
@@ -120,6 +122,8 @@ instance : ToBool Bool where
120122

121123
infixr:30 " <||> " => orM
122124

125+
recommended_spelling "orM" for "<||>" in [orM, «term_<||>_»]
126+
123127
@[macro_inline] def andM {m : Type u → Type v} {β : Type u} [Monad m] [ToBool β] (x y : m β) : m β := do
124128
let b ← x
125129
match toBool b with
@@ -128,6 +132,8 @@ infixr:30 " <||> " => orM
128132

129133
infixr:35 " <&&> " => andM
130134

135+
recommended_spelling "andM" for "<&&>" in [andM, «term_<&&>_»]
136+
131137
@[macro_inline] def notM {m : TypeType v} [Applicative m] (x : m Bool) : m Bool :=
132138
not <$> x
133139

@@ -315,3 +321,7 @@ def Bind.bindLeft [Bind m] (f : α → m β) (ma : m α) : m β :=
315321
@[inherit_doc] infixr:55 " >=> " => Bind.kleisliRight
316322
@[inherit_doc] infixr:55 " <=< " => Bind.kleisliLeft
317323
@[inherit_doc] infixr:55 " =<< " => Bind.bindLeft
324+
325+
recommended_spelling "kleisliRight" for ">=>" in [Bind.kleisliRight, «term_>=>_»]
326+
recommended_spelling "kleisliLeft" for "<=<" in [Bind.kleisliLeft, «term_<=<_»]
327+
recommended_spelling "bindLeft" for "=<<" in [Bind.bindLeft, «term_=<<_»]

src/Init/Core.lean

+23
Original file line numberDiff line numberDiff line change
@@ -133,6 +133,10 @@ structure Iff (a b : Prop) : Prop where
133133
@[inherit_doc] infix:20 " <-> " => Iff
134134
@[inherit_doc] infix:20 " ↔ " => Iff
135135

136+
recommended_spelling "iff" for "↔" in [Iff, «term_↔_»]
137+
/-- prefer `↔` over `<->` -/
138+
recommended_spelling "iff" for "<->" in [Iff, «term_<->_»]
139+
136140
/--
137141
`Sum α β`, or `α ⊕ β`, is the disjoint union of types `α` and `β`.
138142
An element of `α ⊕ β` is either of the form `.inl a` where `a : α`,
@@ -400,6 +404,8 @@ class HasEquiv (α : Sort u) where
400404

401405
@[inherit_doc] infix:50 " ≈ " => HasEquiv.Equiv
402406

407+
recommended_spelling "equiv" for "≈" in [HasEquiv.Equiv, «term_≈_»]
408+
403409
/-! # set notation -/
404410

405411
/-- Notation type class for the subset relation `⊆`. -/
@@ -462,6 +468,16 @@ consisting of all elements in `a` that are not in `b`.
462468
-/
463469
infix:70 " \\ " => SDiff.sdiff
464470

471+
recommended_spelling "subset" for "⊆" in [Subset, «term_⊆_»]
472+
recommended_spelling "ssubset" for "⊂" in [SSubset, «term_⊂_»]
473+
/-- prefer `⊆` over `⊇` -/
474+
recommended_spelling "superset" for "⊇" in [Superset, «term_⊇_»]
475+
/-- prefer `⊂` over `⊃` -/
476+
recommended_spelling "ssuperset" for "⊃" in [SSuperset, «term_⊃_»]
477+
recommended_spelling "union" for "∪" in [Union.union, «term_∪_»]
478+
recommended_spelling "inter" for "∩" in [Inter.inter, «term_∩_»]
479+
recommended_spelling "sdiff" for "\\" in [SDiff.sdiff, «term_\_»]
480+
465481
/-! # collections -/
466482

467483
/-- `EmptyCollection α` is the typeclass which supports the notation `∅`, also written as `{}`. -/
@@ -473,6 +489,9 @@ class EmptyCollection (α : Type u) where
473489
@[inherit_doc] notation "{" "}" => EmptyCollection.emptyCollection
474490
@[inherit_doc] notation "∅" => EmptyCollection.emptyCollection
475491

492+
recommended_spelling "empty" for "{}" in [EmptyCollection.emptyCollection, «term{}»]
493+
recommended_spelling "empty" for "∅" in [EmptyCollection.emptyCollection, «term∅»]
494+
476495
/--
477496
Type class for the `insert` operation.
478497
Used to implement the `{ a, b, c }` syntax.
@@ -650,6 +669,8 @@ Unlike `x ≠ y` (which is notation for `Ne x y`), this is `Bool` valued instead
650669

651670
@[inherit_doc] infix:50 " != " => bne
652671

672+
recommended_spelling "bne" for "!=" in [bne, «term_!=_»]
673+
653674
/--
654675
`LawfulBEq α` is a typeclass which asserts that the `BEq α` implementation
655676
(which supplies the `a == b` notation) coincides with logical equality `a = b`.
@@ -726,6 +747,8 @@ and asserts that `a` and `b` are not equal.
726747

727748
@[inherit_doc] infix:50 " ≠ " => Ne
728749

750+
recommended_spelling "ne" for "≠" in [Ne, «term_≠_»]
751+
729752
section Ne
730753
variable {α : Sort u}
731754
variable {a b : α} {p : Prop}

src/Init/Data/Array/Basic.lean

+4
Original file line numberDiff line numberDiff line change
@@ -18,11 +18,15 @@ universe u v w
1818

1919
/-! ### Array literal syntax -/
2020

21+
/-- Syntax for `Array α`. -/
2122
syntax "#[" withoutPosition(sepBy(term, ", ")) "]" : term
2223

2324
macro_rules
2425
| `(#[ $elems,* ]) => `(List.toArray [ $elems,* ])
2526

27+
recommended_spelling "empty" for "#[]" interm#[_,]»]
28+
recommended_spelling "singleton" for "#[x]" interm#[_,]»]
29+
2630
variable {α : Type u}
2731

2832
namespace Array

src/Init/Data/BitVec/Basic.lean

+5
Original file line numberDiff line numberDiff line change
@@ -157,6 +157,11 @@ section Syntax
157157
syntax:max num noWs "#" noWs term:max : term
158158
macro_rules | `($i:num#$n) => `(BitVec.ofNat $n $i)
159159

160+
/-- not `ofNat_zero` -/
161+
recommended_spelling "zero" for "0#n" in [BitVec.ofNat, «term__#__»]
162+
/-- not `ofNat_one` -/
163+
recommended_spelling "one" for "1#n" in [BitVec.ofNat, «term__#__»]
164+
160165
/-- Unexpander for bit vector literals. -/
161166
@[app_unexpander BitVec.ofNat] def unexpandBitVecOfNat : Lean.PrettyPrinter.Unexpander
162167
| `($(_) $n $i:num) => `($i:num#$n)

src/Init/Data/Bool.lean

+2
Original file line numberDiff line numberDiff line change
@@ -14,6 +14,8 @@ abbrev xor : Bool → Bool → Bool := bne
1414

1515
@[inherit_doc] infixl:33 " ^^ " => xor
1616

17+
recommended_spelling "xor" for "^^" in [xor, «term_^^_»]
18+
1719
instance (p : Bool → Prop) [inst : DecidablePred p] : Decidable (∀ x, p x) :=
1820
match inst true, inst false with
1921
| isFalse ht, _ => isFalse fun h => absurd (h _) ht

src/Init/Data/List/Basic.lean

+9
Original file line numberDiff line numberDiff line change
@@ -959,6 +959,9 @@ def IsPrefix (l₁ : List α) (l₂ : List α) : Prop := Exists fun t => l₁ ++
959959

960960
@[inherit_doc] infixl:50 " <+: " => IsPrefix
961961

962+
/-- not `isPrefix` -/
963+
recommended_spelling "prefix" for "<+:" in [IsPrefix, «term_<+:_»]
964+
962965
/-- `isPrefixOf l₁ l₂` returns `true` Iff `l₁` is a prefix of `l₂`.
963966
That is, there exists a `t` such that `l₂ == l₁ ++ t`. -/
964967
def isPrefixOf [BEq α] : List α → List α → Bool
@@ -1001,6 +1004,9 @@ def IsSuffix (l₁ : List α) (l₂ : List α) : Prop := Exists fun t => t ++ l
10011004

10021005
@[inherit_doc] infixl:50 " <:+ " => IsSuffix
10031006

1007+
/-- not `isSuffix` -/
1008+
recommended_spelling "suffix" for "<:+" in [IsSuffix, «term_<:+_»]
1009+
10041010
/-! ### IsInfix -/
10051011

10061012
/--
@@ -1011,6 +1017,9 @@ def IsInfix (l₁ : List α) (l₂ : List α) : Prop := Exists fun s => Exists f
10111017

10121018
@[inherit_doc] infixl:50 " <:+: " => IsInfix
10131019

1020+
/-- not `isInfix` -/
1021+
recommended_spelling "infix" for "<:+:" in [IsInfix, «term_<:+:_»]
1022+
10141023
/-! ### splitAt -/
10151024

10161025
/--

src/Init/Data/List/Notation.lean

+3
Original file line numberDiff line numberDiff line change
@@ -26,6 +26,9 @@ unless you use side effecting operations like `dbg_trace`.
2626
-/
2727
syntax "[" withoutPosition(term,*,?) "]" : term
2828

29+
recommended_spelling "nil" for "[]" in [List.nil, «term[_]»]
30+
recommended_spelling "singleton" for "[a]" in [List.cons, «term[_]»]
31+
2932
/--
3033
Auxiliary syntax for implementing `[$elem,*]` list literal syntax.
3134
The syntax `%[a,b,c|tail]` constructs a value equivalent to `a::b::c::tail`.

src/Init/Data/Vector/Basic.lean

+3
Original file line numberDiff line numberDiff line change
@@ -36,6 +36,9 @@ open Lean in
3636
macro_rules
3737
| `(#v[ $elems,* ]) => `(Vector.mk (n := $(quote elems.getElems.size)) #[$elems,*] rfl)
3838

39+
recommended_spelling "empty" for "#v[]" in [Vector.mk, «term#v[_,]»]
40+
recommended_spelling "singleton" for "#v[x]" in [Vector.mk, «term#v[_,]»]
41+
3942
/-- Custom eliminator for `Vector α n` through `Array α` -/
4043
@[elab_as_elim]
4144
def elimAsArray {motive : Vector α n → Sort u}

src/Init/GetElem.lean

+5
Original file line numberDiff line numberDiff line change
@@ -114,6 +114,11 @@ panics `i` is out of bounds.
114114
-/
115115
macro:max x:term noWs "[" i:term "]" noWs "!" : term => `(getElem! $x $i)
116116

117+
recommended_spelling "getElem" for "xs[i]" in [GetElem.getElem, «term__[_]»]
118+
recommended_spelling "getElem" for "xs[i]'h" in [GetElem.getElem, «term__[_]'_»]
119+
recommended_spelling "getElem?" for "xs[i]?" in [GetElem?.getElem?, «term__[_]_?»]
120+
recommended_spelling "getElem!" for "xs[i]!" in [GetElem?.getElem!, «term__[_]_!»]
121+
117122
instance (priority := low) [GetElem coll idx elem valid] [∀ xs i, Decidable (valid xs i)] :
118123
GetElem? coll idx elem valid where
119124
getElem? xs i := decidableGetElem? xs i

src/Init/Notation.lean

+59
Original file line numberDiff line numberDiff line change
@@ -273,6 +273,10 @@ syntax (name := rawNatLit) "nat_lit " num : term
273273
@[inherit_doc] infixr:35 " × " => Prod
274274
@[inherit_doc] infixr:35 " ×' " => PProd
275275

276+
recommended_spelling "comp" for "∘" in [Function.comp, «term_∘_»]
277+
recommended_spelling "Prod" for "×" in [Prod, «term_×_»]
278+
recommended_spelling "PProd" for "×'" in [PProd, «term_×'_»]
279+
276280
@[inherit_doc] infix:50 " ∣ " => Dvd.dvd
277281
@[inherit_doc] infixl:55 " ||| " => HOr.hOr
278282
@[inherit_doc] infixl:58 " ^^^ " => HXor.hXor
@@ -306,6 +310,24 @@ macro_rules | `($x ^ $y) => `(rightact% HPow.hPow $x $y)
306310
macro_rules | `($x ++ $y) => `(binop% HAppend.hAppend $x $y)
307311
macro_rules | `(- $x) => `(unop% Neg.neg $x)
308312

313+
recommended_spelling "or" for "|||" in [HOr.hOr, «term_|||_»]
314+
recommended_spelling "xor" for "^^^" in [HXor.hXor, «term_^^^_»]
315+
recommended_spelling "and" for "&&&" in [HAnd.hAnd, «term_&&&_»]
316+
recommended_spelling "add" for "+" in [HAdd.hAdd, «term_+_»]
317+
/-- when used as a binary operator -/
318+
recommended_spelling "sub" for "-" in [HSub.hSub, «term_-_»]
319+
recommended_spelling "mul" for "*" in [HMul.hMul, «term_*_»]
320+
recommended_spelling "div" for "/" in [HDiv.hDiv, «term_/_»]
321+
recommended_spelling "mod" for "%" in [HMod.hMod, «term_%_»]
322+
recommended_spelling "pow" for "^" in [HPow.hPow, «term_^_»]
323+
recommended_spelling "append" for "++" in [HAppend.hAppend, «term_++_»]
324+
/-- when used as a unary operator -/
325+
recommended_spelling "neg" for "-" in [Neg.neg, «term-_»]
326+
recommended_spelling "dvd" for "∣" in [Dvd.dvd, «term_∣_»]
327+
recommended_spelling "shiftLeft" for "<<<" in [HShiftLeft.hShiftLeft, «term_<<<_»]
328+
recommended_spelling "shiftRight" for ">>>" in [HShiftRight.hShiftRight, «term_>>>_»]
329+
recommended_spelling "not" for "~~~" in [Complement.complement, «term~~~_»]
330+
309331
-- declare ASCII alternatives first so that the latter Unicode unexpander wins
310332
@[inherit_doc] infix:50 " <= " => LE.le
311333
@[inherit_doc] infix:50 " ≤ " => LE.le
@@ -330,20 +352,46 @@ macro_rules | `($x ≥ $y) => `(binrel% GE.ge $x $y)
330352
macro_rules | `($x = $y) => `(binrel% Eq $x $y)
331353
macro_rules | `($x == $y) => `(binrel_no_prop% BEq.beq $x $y)
332354

355+
recommended_spelling "le" for "≤" in [LE.le, «term_≤_»]
356+
/-- prefer `≤` over `<=` -/
357+
recommended_spelling "le" for "<=" in [LE.le, «term_<=_»]
358+
recommended_spelling "lt" for "<" in [LT.lt, «term_<_»]
359+
recommended_spelling "gt" for ">" in [GT.gt, «term_>_»]
360+
recommended_spelling "ge" for "≥" in [GE.ge, «term_≥_»]
361+
/-- prefer `≥` over `>=` -/
362+
recommended_spelling "ge" for ">=" in [GE.ge, «term_>=_»]
363+
recommended_spelling "eq" for "=" in [Eq, «term_=_»]
364+
recommended_spelling "beq" for "==" in [BEq.beq, «term_==_»]
365+
333366
@[inherit_doc] infixr:35 " /\\ " => And
334367
@[inherit_doc] infixr:35 " ∧ " => And
335368
@[inherit_doc] infixr:30 " \\/ " => Or
336369
@[inherit_doc] infixr:30 " ∨ " => Or
337370
@[inherit_doc] notation:max "¬" p:40 => Not p
338371

372+
recommended_spelling "and" for "∧" in [And, «term_∧_»]
373+
/-- prefer `∧` over `/\` -/
374+
recommended_spelling "and" for "/\\" in [And, «term_/\_»]
375+
recommended_spelling "or" for "∨" in [Or, «term_∨_»]
376+
/-- prefer `∨` over `\/` -/
377+
recommended_spelling "or" for "\\/" in [Or, «term_\/_»]
378+
recommended_spelling "not" for "¬" in [Not, «term¬_»]
379+
339380
@[inherit_doc] infixl:35 " && " => and
340381
@[inherit_doc] infixl:30 " || " => or
341382
@[inherit_doc] notation:max "!" b:40 => not b
342383

384+
recommended_spelling "and" for "&&" in [and, «term_&&_»]
385+
recommended_spelling "or" for "||" in [and, «term_||_»]
386+
recommended_spelling "not" for "!" in [not, «term!_»]
387+
343388
@[inherit_doc] notation:50 a:50 " ∈ " b:50 => Membership.mem b a
344389
/-- `a ∉ b` is negated elementhood. It is notation for `¬ (a ∈ b)`. -/
345390
notation:50 a:50 " ∉ " b:50 => ¬ (a ∈ b)
346391

392+
recommended_spelling "mem" for "∈" in [Membership.mem, «term_∈_»]
393+
recommended_spelling "not_mem" for "∉" interm_∉_»]
394+
347395
@[inherit_doc] infixr:67 " :: " => List.cons
348396
@[inherit_doc] infixr:100 " <$> " => Functor.map
349397
@[inherit_doc] infixl:55 " >>= " => Bind.bind
@@ -359,6 +407,15 @@ macro_rules | `($x <*> $y) => `(Seq.seq $x fun _ : Unit => $y)
359407
macro_rules | `($x <* $y) => `(SeqLeft.seqLeft $x fun _ : Unit => $y)
360408
macro_rules | `($x *> $y) => `(SeqRight.seqRight $x fun _ : Unit => $y)
361409

410+
recommended_spelling "cons" for "::" in [List.cons, «term_::_»]
411+
recommended_spelling "map" for "<$>" in [Functor.map, «term_<$>_»]
412+
recommended_spelling "bind" for ">>=" in [Bind.bind, «term_>>=_»]
413+
recommended_spelling "orElse" for "<|>" in [HOrElse.hOrElse, «term_<|>_»]
414+
recommended_spelling "andThen" for ">>" in [HAndThen.hAndThen, «term_>>_»]
415+
recommended_spelling "seq" for "<*>" in [Seq.seq, «term_<*>_»]
416+
recommended_spelling "seqLeft" for "<*" in [SeqLeft.seqLeft, «term_<*_»]
417+
recommended_spelling "seqRight" for "*>" in [SeqRight.seqRight, «term_*>_»]
418+
362419
namespace Lean
363420

364421
/--
@@ -463,6 +520,8 @@ macro_rules
463520
| `({ $x : $type // $p }) => ``(Subtype (fun ($x:ident : $type) => $p))
464521
| `({ $x // $p }) => ``(Subtype (fun ($x:ident : _) => $p))
465522

523+
recommended_spelling "subtype" for "{ x // p x }" in [Subtype, «term{_:_//_}»]
524+
466525
/--
467526
`without_expected_type t` instructs Lean to elaborate `t` without an expected type.
468527
Recall that terms such as `match ... with ...` and `⟨...⟩` will postpone elaboration until

src/Init/NotationExtra.lean

+3
Original file line numberDiff line numberDiff line change
@@ -345,12 +345,15 @@ macro:50 e:term:51 " matches " p:sepBy1(term:51, " | ") : term =>
345345

346346
end Lean
347347

348+
/-- `{ a, b, c }` syntax, powered by the `Singleton` and `Insert` typeclasses. -/
348349
syntax "{" term,+ "}" : term
349350

350351
macro_rules
351352
| `({$x:term}) => `(singleton $x)
352353
| `({$x:term, $xs:term,*}) => `(insert $x {$xs:term,*})
353354

355+
recommended_spelling "singleton" for "{x}" in [singleton, «term{_}»]
356+
354357
namespace Lean
355358

356359
/-- Unexpander for the `{ x }` notation. -/

src/Lean/Parser/Term.lean

+2
Original file line numberDiff line numberDiff line change
@@ -254,6 +254,8 @@ do not yield the right result.
254254
@[builtin_term_parser] def tuple := leading_parser
255255
"(" >> optional (withoutPosition (withoutForbidden (termParser >> ", " >> sepBy1 termParser ", " (allowTrailingSep := true)))) >> ")"
256256

257+
recommended_spelling "mk" for "(a, b)" in [Prod.mk, tuple]
258+
257259
/--
258260
Parentheses, used for grouping expressions (e.g., `a * (b + c)`).
259261
Can also be used for creating simple functions when combined with `·`. Here are some examples:

tests/lean/interactive/hover.lean.expected.out

+3-3
Original file line numberDiff line numberDiff line change
@@ -407,7 +407,7 @@
407407
"end": {"line": 206, "character": 13}},
408408
"contents":
409409
{"value":
410-
"```lean\n?m x1✝ x2✝\n```\n***\n`a + b` computes the sum of `a` and `b`.\nThe meaning of this notation is type-dependent. ",
410+
"```lean\n?m x1✝ x2✝\n```\n***\n`a + b` computes the sum of `a` and `b`.\nThe meaning of this notation is type-dependent. \n\nConventions for notations in identifiers:\n\n * The recommended spelling of `+` in identifiers is `add`.",
411411
"kind": "markdown"}}
412412
{"textDocument": {"uri": "file:///hover.lean"},
413413
"position": {"line": 215, "character": 28}}
@@ -612,7 +612,7 @@
612612
"end": {"line": 290, "character": 16}},
613613
"contents":
614614
{"value":
615-
"```lean\nList.nil.{u} {α : Type u} : List α\n```\n***\n`[]` is the empty list. \n***\n*import Init.Prelude*",
615+
"```lean\nList.nil.{u} {α : Type u} : List α\n```\n***\n`[]` is the empty list. \n\nConventions for notations in identifiers:\n\n * The recommended spelling of `[]` in identifiers is `nil`.\n***\n*import Init.Prelude*",
616616
"kind": "markdown"}}
617617
{"textDocument": {"uri": "file:///hover.lean"},
618618
"position": {"line": 292, "character": 13}}
@@ -621,7 +621,7 @@
621621
"end": {"line": 292, "character": 15}},
622622
"contents":
623623
{"value":
624-
"```lean\nList.cons.{u} {α : Type u} (head : α) (tail : List α) : List α\n```\n***\nIf `a : α` and `l : List α`, then `cons a l`, or `a :: l`, is the\nlist whose first element is `a` and with `l` as the rest of the list. \n***\n*import Init.Prelude*",
624+
"```lean\nList.cons.{u} {α : Type u} (head : α) (tail : List α) : List α\n```\n***\nIf `a : α` and `l : List α`, then `cons a l`, or `a :: l`, is the\nlist whose first element is `a` and with `l` as the rest of the list. \n\nConventions for notations in identifiers:\n\n * The recommended spelling of `::` in identifiers is `cons`.\n\n * The recommended spelling of `[a]` in identifiers is `singleton`.\n***\n*import Init.Prelude*",
625625
"kind": "markdown"}}
626626
{"textDocument": {"uri": "file:///hover.lean"},
627627
"position": {"line": 294, "character": 18}}

tests/lean/run/recommendedSpelling.lean

+9-3
Original file line numberDiff line numberDiff line change
@@ -2,11 +2,10 @@ import Lean
22

33
/-!
44
Test the `recommended_spelling` command.
5-
6-
TODO: once we use this command in Init, we can test that recommended spellings from imported
7-
modules are reported correctly.
85
-/
96

7+
recommended_spelling "bland" for "🥤" in [And]
8+
109
/--
1110
Conjuction
1211
@@ -56,6 +55,13 @@ info: some
5655
#guard_msgs in
5756
#eval findDocString? `«term_☋_»
5857

58+
/--
59+
info: some
60+
"`And a b`, or `a ∧ b`, is the conjunction of propositions. It can be\nconstructed and destructed like a pair: if `ha : a` and `hb : b` then\n`⟨ha, hb⟩ : a ∧ b`, and if `h : a ∧ b` then `h.left : a` and `h.right : b`.\n\n\nConventions for notations in identifiers:\n\n * The recommended spelling of `∧` in identifiers is `and`.\n\n * The recommended spelling of `/\\` in identifiers is `and` (prefer `∧` over `/\\`).\n\n * The recommended spelling of `🥤` in identifiers is `bland`."
61+
-/
62+
#guard_msgs in
63+
#eval findDocString? `And
64+
5965
/-- info: 1 -/
6066
#guard_msgs in
6167
#eval do

0 commit comments

Comments
 (0)