Skip to content

Commit f0738c2

Browse files
authored
perf: in CaseValues, subst only once (#11510)
This PR avoids running substCore twice in caseValues.
1 parent 5f561bf commit f0738c2

File tree

7 files changed

+52
-60
lines changed

7 files changed

+52
-60
lines changed

src/Lean/Elab/Match.lean

Lines changed: 1 addition & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -12,6 +12,7 @@ public import Lean.Elab.BindersUtil
1212
public import Lean.Elab.PatternVar
1313
public import Lean.Elab.Quotation.Precheck
1414
public import Lean.Elab.SyntheticMVars
15+
import Lean.Meta.Match.Value
1516
import Lean.Meta.Match.NamedPatterns
1617

1718
public section

src/Lean/Meta/Match/Basic.lean

Lines changed: 7 additions & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -6,8 +6,14 @@ Authors: Leonardo de Moura
66
module
77

88
prelude
9+
public import Lean.Meta.Basic
10+
public import Lean.Meta.Tactic.FVarSubst
911
public import Lean.Meta.CollectFVars
10-
public import Lean.Meta.Match.CaseArraySizes
12+
import Lean.Meta.Match.Value
13+
import Lean.Meta.AppBuilder
14+
import Lean.Meta.Tactic.Util
15+
import Lean.Meta.Tactic.Assert
16+
import Lean.Meta.Tactic.Subst
1117
import Lean.Meta.Match.NamedPatterns
1218

1319
public section

src/Lean/Meta/Match/CaseArraySizes.lean

Lines changed: 25 additions & 23 deletions
Original file line numberDiff line numberDiff line change
@@ -6,32 +6,36 @@ Authors: Leonardo de Moura
66
module
77

88
prelude
9-
public import Lean.Meta.Match.CaseValues
10-
11-
public section
9+
public import Lean.Meta.Basic
10+
public import Lean.Meta.Tactic.FVarSubst
11+
import Lean.Meta.Match.CaseValues
12+
import Lean.Meta.AppBuilder
13+
import Lean.Meta.Tactic.Util
14+
import Lean.Meta.Tactic.Assert
15+
import Lean.Meta.Tactic.Subst
1216

1317
namespace Lean.Meta
1418

15-
structure CaseArraySizesSubgoal where
19+
public structure CaseArraySizesSubgoal where
1620
mvarId : MVarId
1721
elems : Array FVarId := #[]
1822
diseqs : Array FVarId := #[]
1923
subst : FVarSubst := {}
2024
deriving Inhabited
2125

22-
def getArrayArgType (a : Expr) : MetaM Expr := do
26+
public def getArrayArgType (a : Expr) : MetaM Expr := do
2327
let aType ← inferType a
2428
let aType ← whnfD aType
2529
unless aType.isAppOfArity ``Array 1 do
2630
throwError "array expected{indentExpr a}"
2731
pure aType.appArg!
2832

29-
private def mkArrayGetLit (a : Expr) (i : Nat) (n : Nat) (h : Expr) : MetaM Expr := do
33+
def mkArrayGetLit (a : Expr) (i : Nat) (n : Nat) (h : Expr) : MetaM Expr := do
3034
let lt ← mkLt (mkRawNatLit i) (mkRawNatLit n)
3135
let ltPrf ← mkDecideProof lt
3236
mkAppM `Array.getLit #[a, mkRawNatLit i, h, ltPrf]
3337

34-
private partial def introArrayLit (mvarId : MVarId) (a : Expr) (n : Nat) (xNamePrefix : Name) (aSizeEqN : Expr) : MetaM MVarId := do
38+
partial def introArrayLit (mvarId : MVarId) (a : Expr) (n : Nat) (xNamePrefix : Name) (aSizeEqN : Expr) : MetaM MVarId := do
3539
let α ← getArrayArgType a
3640
let rec loop (i : Nat) (xs : Array Expr) (args : Array Expr) := do
3741
if i < n then
@@ -61,7 +65,7 @@ private partial def introArrayLit (mvarId : MVarId) (a : Expr) (n : Nat) (xNameP
6165
n) `..., x_1 ... x_{sizes[n-1]} |- C #[x_1, ..., x_{sizes[n-1]}]`
6266
n+1) `..., (h_1 : a.size != sizes[0]), ..., (h_n : a.size != sizes[n-1]) |- C a`
6367
where `n = sizes.size` -/
64-
def caseArraySizes (mvarId : MVarId) (fvarId : FVarId) (sizes : Array Nat) (xNamePrefix := `x) (hNamePrefix := `h) : MetaM (Array CaseArraySizesSubgoal) :=
68+
public def caseArraySizes (mvarId : MVarId) (fvarId : FVarId) (sizes : Array Nat) (xNamePrefix := `x) (hNamePrefix := `h) : MetaM (Array CaseArraySizesSubgoal) :=
6569
mvarId.withContext do
6670
let a := mkFVar fvarId
6771
let aSize ← mkAppM `Array.size #[a]
@@ -72,22 +76,20 @@ def caseArraySizes (mvarId : MVarId) (fvarId : FVarId) (sizes : Array Nat) (xNam
7276
subgoals.mapIdxM fun i subgoal => do
7377
let subst := subgoal.subst
7478
let mvarId := subgoal.mvarId
75-
let hEqSz := (subst.get hEq).fvarId!
7679
if h : i < sizes.size then
77-
let n := sizes[i]
78-
let mvarId ← mvarId.clear subgoal.newHs[0]!
79-
let mvarId ← mvarId.clear (subst.get aSizeFVarId).fvarId!
80-
mvarId.withContext do
81-
let hEqSzSymm ← mkEqSymm (mkFVar hEqSz)
82-
let mvarId ← introArrayLit mvarId a n xNamePrefix hEqSzSymm
83-
let (xs, mvarId) ← mvarId.introN n
84-
let (hEqLit, mvarId) ← mvarId.intro1
85-
let mvarId ← mvarId.clear hEqSz
86-
let (subst, mvarId) ← substCore mvarId hEqLit false subst
87-
pure { mvarId := mvarId, elems := xs, subst := subst }
80+
let hEqSz := (subst.get hEq).fvarId!
81+
let n := sizes[i]
82+
mvarId.withContext do
83+
let hEqSzSymm ← mkEqSymm (mkFVar hEqSz)
84+
let mvarId ← introArrayLit mvarId a n xNamePrefix hEqSzSymm
85+
let (xs, mvarId) ← mvarId.introN n
86+
let (hEqLit, mvarId) ← mvarId.intro1
87+
let mvarId ← mvarId.clear hEqSz
88+
let (subst, mvarId) ← substCore mvarId hEqLit (symm := false) subst
89+
pure { mvarId := mvarId, elems := xs, subst := subst }
8890
else
89-
let (subst, mvarId) ← substCore mvarId hEq false subst
90-
let diseqs := subgoal.newHs.map fun fvarId => (subst.get fvarId).fvarId!
91-
pure { mvarId := mvarId, diseqs := diseqs, subst := subst }
91+
let (subst, mvarId) ← substCore mvarId hEq (symm := false) subst
92+
let diseqs := subgoal.newHs.map fun fvarId => (subst.get fvarId).fvarId!
93+
pure { mvarId := mvarId, diseqs := diseqs, subst := subst }
9294

9395
end Lean.Meta

src/Lean/Meta/Match/CaseValues.lean

Lines changed: 14 additions & 35 deletions
Original file line numberDiff line numberDiff line change
@@ -6,28 +6,25 @@ Authors: Leonardo de Moura
66
module
77

88
prelude
9-
public import Lean.Meta.Tactic.Subst
10-
public import Lean.Meta.Match.Value
11-
12-
public section
9+
public import Lean.Meta.Basic
10+
public import Lean.Meta.Tactic.FVarSubst
11+
import Lean.Meta.Tactic.Subst
1312

1413
namespace Lean.Meta
1514

1615
structure CaseValueSubgoal where
1716
mvarId : MVarId
1817
newH : FVarId
19-
subst : FVarSubst := {}
2018
deriving Inhabited
2119

2220
/--
2321
Split goal `... |- C x` into two subgoals
24-
`..., (h : x = value) |- C value`
22+
`..., (h : x = value) |- C x`
2523
`..., (h : x != value) |- C x`
2624
where `fvarId` is `x`s id.
2725
The type of `x` must have decidable equality.
28-
29-
Remark: `subst` field of the second subgoal is equal to the input `subst`. -/
30-
private def caseValueAux (mvarId : MVarId) (fvarId : FVarId) (value : Expr) (hName : Name := `h) (subst : FVarSubst := {})
26+
-/
27+
def caseValue (mvarId : MVarId) (fvarId : FVarId) (value : Expr) (hName : Name := `h)
3128
: MetaM (CaseValueSubgoal × CaseValueSubgoal) :=
3229
mvarId.withContext do
3330
let tag ← mvarId.getTag
@@ -42,27 +39,16 @@ private def caseValueAux (mvarId : MVarId) (fvarId : FVarId) (value : Expr) (hNa
4239
let val ← mkAppOptM `dite #[none, xEqValue, none, thenMVar, elseMVar]
4340
mvarId.assign val
4441
let (elseH, elseMVarId) ← elseMVar.mvarId!.intro1P
45-
let elseSubgoal := { mvarId := elseMVarId, newH := elseH, subst := subst : CaseValueSubgoal }
42+
let elseSubgoal := { mvarId := elseMVarId, newH := elseH }
4643
let (thenH, thenMVarId) ← thenMVar.mvarId!.intro1P
47-
let symm := false
48-
let clearH := false
49-
let (thenSubst, thenMVarId) ← substCore thenMVarId thenH symm subst clearH
5044
thenMVarId.withContext do
51-
trace[Meta] "subst domain: {thenSubst.domain.map (·.name)}"
52-
let thenH := (thenSubst.get thenH).fvarId!
5345
trace[Meta] "searching for decl"
5446
let _ ← thenH.getDecl
5547
trace[Meta] "found decl"
56-
let thenSubgoal := { mvarId := thenMVarId, newH := (thenSubst.get thenH).fvarId!, subst := thenSubst : CaseValueSubgoal }
48+
let thenSubgoal := { mvarId := thenMVarId, newH := thenH }
5749
pure (thenSubgoal, elseSubgoal)
5850

59-
def caseValue (mvarId : MVarId) (fvarId : FVarId) (value : Expr) : MetaM (CaseValueSubgoal × CaseValueSubgoal) := do
60-
let s ← caseValueAux mvarId fvarId value
61-
appendTagSuffix s.1.mvarId `thenBranch
62-
appendTagSuffix s.2.mvarId `elseBranch
63-
pure s
64-
65-
structure CaseValuesSubgoal where
51+
public structure CaseValuesSubgoal where
6652
mvarId : MVarId
6753
newHs : Array FVarId := #[]
6854
subst : FVarSubst := {}
@@ -83,22 +69,15 @@ structure CaseValuesSubgoal where
8369
8470
If `substNewEqs = true`, then the new `h_i` equality hypotheses are substituted in the first `n` cases.
8571
-/
86-
def caseValues (mvarId : MVarId) (fvarId : FVarId) (values : Array Expr) (hNamePrefix := `h) (substNewEqs := false) : MetaM (Array CaseValuesSubgoal) :=
72+
public def caseValues (mvarId : MVarId) (fvarId : FVarId) (values : Array Expr) (hNamePrefix := `h) : MetaM (Array CaseValuesSubgoal) :=
8773
let rec loop : Nat → MVarId → List Expr → Array FVarId → Array CaseValuesSubgoal → MetaM (Array CaseValuesSubgoal)
8874
| _, mvarId, [], _, _ => throwTacticEx `caseValues mvarId "list of values must not be empty"
8975
| i, mvarId, v::vs, hs, subgoals => do
90-
let (thenSubgoal, elseSubgoal) ← caseValueAux mvarId fvarId v (hNamePrefix.appendIndexAfter i) {}
76+
let (thenSubgoal, elseSubgoal) ← caseValue mvarId fvarId v (hNamePrefix.appendIndexAfter i)
9177
appendTagSuffix thenSubgoal.mvarId ((`case).appendIndexAfter i)
92-
let thenMVarId ← hs.foldlM
93-
(fun thenMVarId h => match thenSubgoal.subst.get h with
94-
| Expr.fvar fvarId => thenMVarId.tryClear fvarId
95-
| _ => pure thenMVarId)
96-
thenSubgoal.mvarId
97-
let subgoals ← if substNewEqs then
98-
let (subst, mvarId) ← substCore thenMVarId thenSubgoal.newH false thenSubgoal.subst true
99-
pure <| subgoals.push { mvarId := mvarId, newHs := #[], subst := subst }
100-
else
101-
pure <| subgoals.push { mvarId := thenMVarId, newHs := #[thenSubgoal.newH], subst := thenSubgoal.subst }
78+
let thenMVarId ← thenSubgoal.mvarId.tryClearMany hs
79+
let (subst, mvarId) ← substCore thenMVarId thenSubgoal.newH (symm := false) {} (clearH := true)
80+
let subgoals := subgoals.push { mvarId := mvarId, newHs := #[], subst := subst }
10281
match vs with
10382
| [] => do
10483
appendTagSuffix elseSubgoal.mvarId ((`case).appendIndexAfter (i+1))

src/Lean/Meta/Match/Match.lean

Lines changed: 3 additions & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -16,6 +16,8 @@ public import Lean.Meta.Match.MVarRenaming
1616
import Lean.Meta.Match.SimpH
1717
import Lean.Meta.Match.SolveOverlap
1818
import Lean.Meta.HasNotBit
19+
import Lean.Meta.Match.CaseArraySizes
20+
import Lean.Meta.Match.CaseValues
1921
import Lean.Meta.Match.NamedPatterns
2022

2123
public section
@@ -724,7 +726,7 @@ private def processValue (p : Problem) : MetaM (Array Problem) := do
724726
trace[Meta.Match.match] "value step"
725727
let x :: xs := p.vars | unreachable!
726728
let values := collectValues p
727-
let subgoals ← caseValues p.mvarId x.fvarId! values (substNewEqs := true)
729+
let subgoals ← caseValues p.mvarId x.fvarId! values
728730
subgoals.mapIdxM fun i subgoal => do
729731
trace[Meta.Match.match] "processValue subgoal\n{MessageData.ofGoal subgoal.mvarId}"
730732
if h : i < values.size then

src/Lean/Meta/Match/MatchEqsExt.lean

Lines changed: 1 addition & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -8,6 +8,7 @@ module
88
prelude
99
public import Lean.Meta.Basic
1010
public import Lean.Meta.Match.Basic
11+
public import Lean.Meta.Match.MatcherInfo
1112
import Lean.Meta.Eqns
1213

1314
public section

src/Lean/Meta/Match/MatcherApp/Transform.lean

Lines changed: 1 addition & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -10,6 +10,7 @@ prelude
1010
public import Lean.Meta.Match.MatcherApp.Basic
1111
public import Lean.Meta.Match.MatchEqsExt
1212
public import Lean.Meta.Match.AltTelescopes
13+
public import Lean.Meta.AppBuilder
1314
import Lean.Meta.Tactic.Split
1415
import Lean.Meta.Tactic.Refl
1516

0 commit comments

Comments
 (0)