Skip to content

Commit 24ac0e4

Browse files
fpvandoornkim-em
andcommitted
feat: better name guessing for to_additive (#779)
* `@[to_additive]` will now correctly guess names with `CoeTC`, `CoeT` and `CoeHTCT` in it * rewrite function `targetName`. - It will now warn the user if it gives a composite name that can be auto-generated (before `to_additive` would never warn if a composite name was given). - the logic when `allowAutoName = true` now corresponds to the docstring - Fix a bug where declarations were silently allowed to translate to itself (maybe because the `return` statements returned a value for the whole function?) - Add some more tracing - The behavior of namespaces when giving a composite name has been changed. It will always generate a name with the same number of components. Example: ```lean namespace MeasureTheory.MulMeasure @[to_additive AddMeasure.myOperation] def myOperation := ... -- before: generates `AddMeasure.myOperation` (and never gives a warning) -- after: generates `MeasureTheory.AddMeasure.myOperation` (and probably warns that the name can be auto-generated) end MeasureTheory.MulMeasure ``` * This should fix all problems in #659 other than #660 Minor changes: * When applying `@[to_additive]` to a structure, add a trace message if no translation is inserted for a field. * Define `Name.fromComponents` and `Name.splitAt` * Reduce transitive imports of `Tactic/toAdditive` * Move some auxiliary declarations from `Tactic/Simps` to more appropriate places - swap arguments of `String.isPrefixOf?` - improve `Name.getString` Co-authored-by: Scott Morrison <scott@tqft.net>
1 parent 716d703 commit 24ac0e4

File tree

5 files changed

+91
-57
lines changed

5 files changed

+91
-57
lines changed

Mathlib/Data/String/Defs.lean

Lines changed: 5 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -20,4 +20,9 @@ then reassembles the string by intercalating the separator token `c` over the ma
2020
def mapTokens (c : Char) (f : String → String) : String → String :=
2121
intercalate (singleton c) ∘ List.map f ∘ (·.split (· = c))
2222

23+
/-- `isPrefixOf? pre s` returns `some post` if `s = pre ++ post`.
24+
If `pre` is not a prefix of `s`, it returns `none`. -/
25+
def isPrefixOf? (pre s : String) : Option String :=
26+
if startsWith s pre then some <| s.drop pre.length else none
27+
2328
end String

Mathlib/Lean/Expr/Basic.lean

Lines changed: 29 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -6,6 +6,7 @@ Floris van Doorn, E.W.Ayers, Arthur Paulino
66
-/
77
import Lean
88
import Mathlib.Util.MapsTo
9+
import Std.Data.List.Basic
910

1011
/-!
1112
# Additional operations on Expr and related types
@@ -43,6 +44,34 @@ def mapPrefix (f : Name → Option Name) (n : Name) : Name := Id.run do
4344
| str n' s => mkStr (mapPrefix f n') s
4445
| num n' i => mkNum (mapPrefix f n') i
4546

47+
/-- Build a name from components. For example ``from_components [`foo, `bar]`` becomes
48+
``` `foo.bar```.
49+
It is the inverse of `Name.components` on list of names that have single components. -/
50+
def fromComponents : List Name → Name := go .anonymous where
51+
/-- Auxiliary for `Name.fromComponents` -/
52+
go : Name → List Name → Name
53+
| n, [] => n
54+
| n, s :: rest => go (s.updatePrefix n) rest
55+
56+
/-- Update the last component of a name. -/
57+
def updateLast (f : String → String) : Name → Name
58+
| .str n s => .str n (f s)
59+
| n => n
60+
61+
/-- Get the last field of a name as a string.
62+
Doesn't raise an error when the last component is a numeric field. -/
63+
def getString : Name → String
64+
| .str _ s => s
65+
| .num _ n => toString n
66+
| .anonymous => ""
67+
68+
/-- `nm.splitAt n` splits a name `nm` in two parts, such that the *second* part has depth `n`, i.e.
69+
`(nm.splitAt n).2.getNumParts = n` (assuming `nm.getNumParts ≥ n`).
70+
Example: ``splitAt `foo.bar.baz.back.bat 1 = (`foo.bar.baz.back, `bat)``. -/
71+
def splitAt (nm : Name) (n : Nat) : Name × Name :=
72+
let (nm2, nm1) := (nm.componentsRev.splitAt n)
73+
(.fromComponents <| nm1.reverse, .fromComponents <| nm2.reverse)
74+
4675
end Name
4776

4877

Mathlib/Tactic/Simps/Basic.lean

Lines changed: 2 additions & 23 deletions
Original file line numberDiff line numberDiff line change
@@ -63,35 +63,14 @@ There are some small changes in the attribute. None of them should have great ef
6363
structures, projections, simp, simplifier, generates declarations
6464
-/
6565

66-
-- move
67-
namespace String
68-
69-
/-- `isPrefixOf? s pre` returns `some post` if `s = pre ++ post`.
70-
If `pre` is not a prefix of `s`, it returns `none`. -/
71-
def isPrefixOf? (s pre : String) : Option String :=
72-
if startsWith s pre then some <| s.drop pre.length else none
73-
74-
end String
75-
7666
open Lean Meta Parser Elab Term Command
7767

78-
/-- Update the last component of a name. -/
79-
def Lean.Name.updateLast (f : String → String) : Name → Name
80-
| .str n s => .str n (f s)
81-
| n => n
82-
8368
/-- `updateName nm s is_prefix` adds `s` to the last component of `nm`,
8469
either as prefix or as suffix (specified by `isPrefix`), separated by `_`.
8570
Used by `simps_add_projections`. -/
8671
def updateName (nm : Name) (s : String) (isPrefix : Bool) : Name :=
8772
nm.updateLast fun s' ↦ if isPrefix then s ++ "_" ++ s' else s' ++ "_" ++ s
8873

89-
/-- Get the last field of a name as a string.
90-
Doesn't raise an error when the last component is a numeric field. -/
91-
def Lean.Name.getString : Name → String
92-
| .str _ s => s
93-
| _ => ""
94-
9574
-- move
9675
namespace Lean.Meta
9776
open Tactic Simp
@@ -455,7 +434,7 @@ partial def getCompositeOfProjectionsAux (str : Name) (proj : String) (x : Expr)
455434
let env ← getEnv
456435
let projs := (getStructureInfo? env str).get!
457436
let projInfo := projs.fieldNames.toList.mapIdx fun n p ↦
458-
return (← proj.isPrefixOf? ("_" ++ p.getString!), n, p)
437+
return (← ("_" ++ p.getString!).isPrefixOf? proj, n, p)
459438
let some (projRest, index, projName) := (projInfo.filterMap id).getLast?
460439
| throwError "Failed to find constructor {proj.drop 1} in structure {str}."
461440
let strDecl := (env.find? str).get!
@@ -992,7 +971,7 @@ partial def simpsAddProjections (nm : Name) (type lhs rhs : Expr)
992971
}be customly defined for `simps`, and differ from the projection names of the structure."
993972
let nms ← projInfo.concatMapM fun ⟨newRhs, proj, projExpr, projNrs, isDefault, isPrefix⟩ ↦ do
994973
let newType ← inferType newRhs
995-
let newTodo := todo.filterMap fun x ↦ x.isPrefixOf? ("_" ++ proj.getString)
974+
let newTodo := todo.filterMap fun x ↦ ("_" ++ proj.getString).isPrefixOf? x
996975
-- we only continue with this field if it is default or mentioned in todo
997976
if !(isDefault && todo.isEmpty) && newTodo.isEmpty then return #[]
998977
let newLhs := projExpr.instantiateLambdasOrApps #[lhsAp]

Mathlib/Tactic/ToAdditive.lean

Lines changed: 45 additions & 30 deletions
Original file line numberDiff line numberDiff line change
@@ -5,12 +5,7 @@ Authors: Mario Carneiro, Yury Kudryashov, Floris van Doorn, Jon Eugster
55
Ported by: E.W.Ayers
66
-/
77
import Mathlib.Data.String.Defs
8-
import Mathlib.Lean.Expr.Basic
98
import Mathlib.Lean.Expr.ReplaceRec
10-
import Mathlib.Lean.Expr
11-
import Lean
12-
import Lean.Data
13-
import Lean.Elab.Term
149
import Std.Lean.NameMapAttribute
1510
import Std.Data.Option.Basic
1611

@@ -50,9 +45,20 @@ macro "to_additive!?" rest:to_additiveRest : attr => `(attr| to_additive ! ? $re
5045
/-- The `to_additive` attribute. -/
5146
macro "to_additive?!" rest:to_additiveRest : attr => `(attr| to_additive ! ? $rest)
5247

53-
/-- A set of strings of names that are written in all-caps. -/
54-
def allCapitalNames : RBTree String compare :=
55-
.ofList ["LE", "LT", "WF"]
48+
/-- A set of strings of names that end in a capital letter.
49+
* If the string contains a lowercase letter, the string should be split between the first occurrence
50+
of a lower-case letter followed by a upper-case letter.
51+
* If multiple strings have the same prefix, they should be grouped by prefix
52+
* In this case, the second list should be prefix-free
53+
(no element can be a prefix of a later element)
54+
55+
Todo: automate the translation from `String` to an element in this `RBMap`
56+
(but this would require having something similar to the `rb_lmap` from Lean 3). -/
57+
def endCapitalNames : Lean.RBMap String (List String) compare :=
58+
-- endCapitalNamesOfList ["LE", "LT", "WF", "CoeTC", "CoeT", "CoeHTCT"]
59+
.ofList [("LE", [""]), ("LT", [""]), ("WF", [""]), ("Coe", ["TC", "T", "HTCT"])]
60+
61+
5662

5763
/--
5864
This function takes a String and splits it into separate parts based on the following
@@ -62,21 +68,28 @@ E.g. `#eval "InvHMulLEConjugate₂Smul_ne_top".splitCase` yields
6268
`["Inv", "HMul", "LE", "Conjugate₂", "Smul", "_", "ne", "_", "top"]`.
6369
-/
6470
partial def String.splitCase (s : String) (i₀ : Pos := 0) (r : List String := []) : List String :=
71+
Id.run do
6572
-- We test if we need to split between `i₀` and `i₁`.
6673
let i₁ := s.next i₀
6774
if s.atEnd i₁ then
6875
-- If `i₀` is the last position, return the list.
6976
let r := s::r
70-
r.reverse
77+
return r.reverse
7178
/- We split the string in three cases
7279
* We split on both sides of `_` to keep them there when rejoining the string;
73-
* We split after a sequence of capital letters in `allCapitalNames`;
74-
* We split after a lower-case letter that is followed by an upper-case letter. -/
75-
else if s.get i₀ == '_' || s.get i₁ == '_' ||
76-
((s.get i₁).isUpper && (!(s.get i₀).isUpper || allCapitalNames.contains (s.extract 0 i₁))) then
77-
splitCase (s.extract i₁ s.endPos) 0 <| (s.extract 0 i₁)::r
78-
else
79-
splitCase s i₁ r
80+
* We split after a name in `endCapitalNames`;
81+
* We split after a lower-case letter that is followed by an upper-case letter
82+
(unless it is part of a name in `endCapitalNames`). -/
83+
if s.get i₀ == '_' || s.get i₁ == '_' then
84+
return splitCase (s.extract i₁ s.endPos) 0 <| (s.extract 0 i₁)::r
85+
if (s.get i₁).isUpper then
86+
if let some strs := endCapitalNames.find? (s.extract 0 i₁) then
87+
if let some (pref, newS) := strs.findSome?
88+
fun x ↦ x.isPrefixOf? (s.extract i₁ s.endPos) |>.map (x, ·) then
89+
return splitCase newS 0 <| (s.extract 0 i₁ ++ pref)::r
90+
if !(s.get i₀).isUpper then
91+
return splitCase (s.extract i₁ s.endPos) 0 <| (s.extract 0 i₁)::r
92+
return splitCase s i₁ r
8093

8194
namespace ToAdditive
8295

@@ -271,7 +284,7 @@ def applyReplacementFun : Expr → MetaM Expr :=
271284
match e with
272285
| .lit (.natVal 1) => pure <| mkRawNatLit 0
273286
| .const n₀ ls => do
274-
let n₁ := Name.mapPrefix (findTranslation? <|← getEnv) n₀
287+
let n₁ := n₀.mapPrefix (findTranslation? <|← getEnv)
275288
if n₀ != n₁ then
276289
trace[to_additive_detail] "applyReplacementFun: {n₀} → {n₁}"
277290
let ls : List Level ← (do -- [todo] just get Lean to figure out the levels?
@@ -675,21 +688,21 @@ def guessName : String → String :=
675688

676689
/-- Return the provided target name or autogenerate one if one was not provided. -/
677690
def targetName (src tgt : Name) (allowAutoName : Bool) : CoreM Name := do
678-
let res ← do
679-
if tgt.getPrefix != Name.anonymous || allowAutoName then
680-
return tgt
681-
let .str pre s := src | throwError "to_additive: can't transport {src}"
682-
let tgt_auto := guessName s
683-
if tgt.toString == tgt_auto then
684-
dbg_trace "{src}: correctly autogenerated target name {tgt_auto
685-
}, you may remove the explicit {tgt} argument."
686-
let pre := pre.mapPrefix <| findTranslation? (← getEnv)
687-
if tgt == Name.anonymous then
688-
return Name.mkStr pre tgt_auto
689-
else
690-
return Name.mkStr pre tgt.toString
691+
let .str pre s := src | throwError "to_additive: can't transport {src}"
692+
trace[to_additive_detail] "The name {s} splits as {s.splitCase}"
693+
let tgt_auto := guessName s
694+
let depth := tgt.getNumParts
695+
let pre := pre.mapPrefix <| findTranslation? (← getEnv)
696+
let (pre1, pre2) := pre.splitAt (depth - 1)
697+
if tgt == pre2.str tgt_auto && !allowAutoName then
698+
dbg_trace "to_additive correctly autogenerated target name for {src}. {"\n"
699+
}You may remove the explicit argument {tgt}."
700+
let res := if tgt == .anonymous then pre.str tgt_auto else pre1 ++ tgt
701+
-- we allow translating to itself if `tgt == src`, which is occasionally useful for `additiveTest`
691702
if res == src && tgt != src then
692703
throwError "to_additive: can't transport {src} to itself."
704+
if tgt != .anonymous then
705+
trace[to_additive_detail] "The automatically generated name would be {pre.str tgt_auto}"
693706
return res
694707

695708
private def proceedFieldsAux (src tgt : Name) (f : Name → CoreM (Array Name)) : CoreM Unit := do
@@ -700,6 +713,8 @@ private def proceedFieldsAux (src tgt : Name) (f : Name → CoreM (Array Name))
700713
for (srcField, tgtField) in srcFields.zip tgtFields do
701714
if srcField != tgtField then
702715
insertTranslation (src ++ srcField) (tgt ++ tgtField)
716+
else
717+
trace[to_additive] "Translation {src ++ srcField} ↦ {tgt ++ tgtField} is automatic."
703718

704719
/-- Add the structure fields of `src` to the translations dictionary
705720
so that future uses of `to_additive` will map them to the corresponding `tgt` fields. -/

test/toAdditive.lean

Lines changed: 10 additions & 4 deletions
Original file line numberDiff line numberDiff line change
@@ -103,6 +103,9 @@ run_cmd do
103103
let t ← Elab.Command.liftCoreM <| Lean.Meta.MetaM.run' <| ToAdditive.expand c.type
104104
let decl := c |>.updateName `Test.barr6 |>.updateType t |>.updateValue e |>.toDeclaration!
105105
Elab.Command.liftCoreM <| addAndCompile decl
106+
-- test that we cannot transport a declaration to itself
107+
successIfFail <| Elab.Command.liftCoreM <|
108+
ToAdditive.addToAdditiveAttr `bar11_works { ref := ← getRef }
106109

107110
/-! Test the namespace bug (#8733). This code should *not* generate a lemma
108111
`add_some_def.in_namespace`. -/
@@ -160,7 +163,7 @@ def pi_nat_has_one {I : Type} : One ((x : I) → Nat) := pi.has_one
160163

161164
example : @pi_nat_has_one = @pi_nat_has_zero := rfl
162165

163-
section noncomputablee
166+
section test_noncomputable
164167

165168
@[to_additive Bar.bar]
166169
noncomputable def Foo.foo (h : ∃ _ : α, True) : α := Classical.choose h
@@ -171,9 +174,9 @@ def Foo.foo' : ℕ := 2
171174
theorem Bar.bar'_works : Bar.bar' = 2 := by decide
172175

173176
run_cmd (do
174-
if !isNoncomputable (← getEnv) `Bar.bar then throwError "bar shouldn't be computable"
175-
if isNoncomputable (← getEnv) `Bar.bar' then throwError "bar' should be computable")
176-
end noncomputablee
177+
if !isNoncomputable (← getEnv) `Test.Bar.bar then throwError "bar shouldn't be computable"
178+
if isNoncomputable (← getEnv) `Test.Bar.bar' then throwError "bar' should be computable")
179+
end test_noncomputable
177180

178181
section instances
179182

@@ -254,6 +257,9 @@ run_cmd
254257
checkGuessName "LTHMulHPowLEHDiv" "LTHAddHSMulLEHSub"
255258
checkGuessName "OneLEHMul" "NonnegHAdd"
256259
checkGuessName "OneLTHPow" "PosHSMul"
260+
checkGuessName "instCoeTCOneHom" "instCoeTCZeroHom"
261+
checkGuessName "instCoeTOneHom" "instCoeTZeroHom"
262+
checkGuessName "instCoeOneHom" "instCoeZeroHom"
257263

258264
end guessName
259265

0 commit comments

Comments
 (0)