Skip to content

Commit 72e952e

Browse files
authored
chore: avoid runtime array bounds checks (leanprover#6134)
This PR avoids runtime array bounds checks in places where it can trivially be done at compile time. None of these changes are of particular consequence: I mostly wanted to learn how much we do this, and what the obstacles are to doing it less.
1 parent 56a80de commit 72e952e

34 files changed

+130
-130
lines changed

src/Init/Meta.lean

Lines changed: 8 additions & 9 deletions
Original file line numberDiff line numberDiff line change
@@ -431,21 +431,20 @@ def getSubstring? (stx : Syntax) (withLeading := true) (withTrailing := true) :
431431
}
432432
| _, _ => none
433433

434-
@[specialize] private partial def updateLast {α} [Inhabited α] (a : Array α) (f : α → Option α) (i : Nat) : Option (Array α) :=
435-
if i == 0 then
436-
none
437-
else
438-
let i := i - 1
439-
let v := a[i]!
434+
@[specialize] private partial def updateLast {α} (a : Array α) (f : α → Option α) (i : Fin (a.size + 1)) : Option (Array α) :=
435+
match i with
436+
| 0 => none
437+
| ⟨i + 1, h⟩ =>
438+
let v := a[i]'(Nat.succ_lt_succ_iff.mp h)
440439
match f v with
441-
| some v => some <| a.set! i v
442-
| none => updateLast a f i
440+
| some v => some <| a.set i v (Nat.succ_lt_succ_iff.mp h)
441+
| none => updateLast a f ⟨i, Nat.lt_of_succ_lt h⟩
443442

444443
partial def setTailInfoAux (info : SourceInfo) : Syntax → Option Syntax
445444
| atom _ val => some <| atom info val
446445
| ident _ rawVal val pre => some <| ident info rawVal val pre
447446
| node info' k args =>
448-
match updateLast args (setTailInfoAux info) args.size with
447+
match updateLast args (setTailInfoAux info) args.size, by simp⟩ with
449448
| some args => some <| node info' k args
450449
| none => none
451450
| _ => none

src/Init/NotationExtra.lean

Lines changed: 10 additions & 10 deletions
Original file line numberDiff line numberDiff line change
@@ -22,28 +22,28 @@ syntax explicitBinders := (ppSpace bracketedExplicitBinders)+ <|> unb
2222

2323
open TSyntax.Compat in
2424
def expandExplicitBindersAux (combinator : Syntax) (idents : Array Syntax) (type? : Option Syntax) (body : Syntax) : MacroM Syntax :=
25-
let rec loop (i : Nat) (acc : Syntax) := do
25+
let rec loop (i : Nat) (h : i ≤ idents.size) (acc : Syntax) := do
2626
match i with
2727
| 0 => pure acc
28-
| i+1 =>
29-
let ident := idents[i]![0]
28+
| i + 1 =>
29+
let ident := idents[i][0]
3030
let acc ← match ident.isIdent, type? with
3131
| true, none => `($combinator fun $ident => $acc)
3232
| true, some type => `($combinator fun $ident : $type => $acc)
3333
| false, none => `($combinator fun _ => $acc)
3434
| false, some type => `($combinator fun _ : $type => $acc)
35-
loop i acc
36-
loop idents.size body
35+
loop i (Nat.le_of_succ_le h) acc
36+
loop idents.size (by simp) body
3737

3838
def expandBrackedBindersAux (combinator : Syntax) (binders : Array Syntax) (body : Syntax) : MacroM Syntax :=
39-
let rec loop (i : Nat) (acc : Syntax) := do
39+
let rec loop (i : Nat) (h : i ≤ binders.size) (acc : Syntax) := do
4040
match i with
4141
| 0 => pure acc
4242
| i+1 =>
43-
let idents := binders[i]![1].getArgs
44-
let type := binders[i]![3]
45-
loop i (← expandExplicitBindersAux combinator idents (some type) acc)
46-
loop binders.size body
43+
let idents := binders[i][1].getArgs
44+
let type := binders[i][3]
45+
loop i (Nat.le_of_succ_le h) (← expandExplicitBindersAux combinator idents (some type) acc)
46+
loop binders.size (by simp) body
4747

4848
def expandExplicitBinders (combinatorDeclName : Name) (explicitBinders : Syntax) (body : Syntax) : MacroM Syntax := do
4949
let combinator := mkCIdentFrom (← getRef) combinatorDeclName

src/Init/System/Uri.lean

Lines changed: 6 additions & 6 deletions
Original file line numberDiff line numberDiff line change
@@ -29,13 +29,13 @@ def decodeUri (uri : String) : String := Id.run do
2929
let len := rawBytes.size
3030
let mut i := 0
3131
let percent := '%'.toNat.toUInt8
32-
while i < len do
33-
let c := rawBytes[i]!
34-
(decoded, i) := if c == percent && i + 1 < len then
35-
let h1 := rawBytes[i + 1]!
32+
while h : i < len do
33+
let c := rawBytes[i]
34+
(decoded, i) := if h₁ : c == percent i + 1 < len then
35+
let h1 := rawBytes[i + 1]
3636
if let some hd1 := hexDigitToUInt8? h1 then
37-
if i + 2 < len then
38-
let h2 := rawBytes[i + 2]!
37+
if h₂ : i + 2 < len then
38+
let h2 := rawBytes[i + 2]
3939
if let some hd2 := hexDigitToUInt8? h2 then
4040
-- decode the hex digits into a byte.
4141
(decoded.push (hd1 * 16 + hd2), i + 3)

src/Lean/Compiler/IR/EmitC.lean

Lines changed: 3 additions & 3 deletions
Original file line numberDiff line numberDiff line change
@@ -271,9 +271,9 @@ def emitTag (x : VarId) (xType : IRType) : M Unit := do
271271
emit x
272272

273273
def isIf (alts : Array Alt) : Option (Nat × FnBody × FnBody) :=
274-
if alts.size != 2 then none
275-
else match alts[0]! with
276-
| Alt.ctor c b => some (c.cidx, b, alts[1]!.body)
274+
if h : alts.size 2 then none
275+
else match alts[0] with
276+
| Alt.ctor c b => some (c.cidx, b, alts[1].body)
277277
| _ => none
278278

279279
def emitInc (x : VarId) (n : Nat) (checkRef : Bool) : M Unit := do

src/Lean/Compiler/IR/EmitLLVM.lean

Lines changed: 7 additions & 6 deletions
Original file line numberDiff line numberDiff line change
@@ -1172,8 +1172,8 @@ def emitFnArgs (builder : LLVM.Builder llvmctx)
11721172
(needsPackedArgs? : Bool) (llvmfn : LLVM.Value llvmctx) (params : Array Param) : M llvmctx Unit := do
11731173
if needsPackedArgs? then do
11741174
let argsp ← LLVM.getParam llvmfn 0 -- lean_object **args
1175-
for i in List.range params.size do
1176-
let param := params[i]!
1175+
for h : i in [:params.size] do
1176+
let param := params[i]
11771177
-- argsi := (args + i)
11781178
let argsi ← LLVM.buildGEP2 builder (← LLVM.voidPtrType llvmctx) argsp #[← constIntUnsigned i] s!"packed_arg_{i}_slot"
11791179
let llvmty ← toLLVMType param.ty
@@ -1182,15 +1182,16 @@ def emitFnArgs (builder : LLVM.Builder llvmctx)
11821182
-- slot for arg[i] which is always void* ?
11831183
let alloca ← buildPrologueAlloca builder llvmty s!"arg_{i}"
11841184
LLVM.buildStore builder pv alloca
1185-
addVartoState params[i]!.x alloca llvmty
1185+
addVartoState param.x alloca llvmty
11861186
else
11871187
let n ← LLVM.countParams llvmfn
1188-
for i in (List.range n.toNat) do
1189-
let llvmty ← toLLVMType params[i]!.ty
1188+
for i in [:n.toNat] do
1189+
let param := params[i]!
1190+
let llvmty ← toLLVMType param.ty
11901191
let alloca ← buildPrologueAlloca builder llvmty s!"arg_{i}"
11911192
let arg ← LLVM.getParam llvmfn (UInt64.ofNat i)
11921193
let _ ← LLVM.buildStore builder arg alloca
1193-
addVartoState params[i]!.x alloca llvmty
1194+
addVartoState param.x alloca llvmty
11941195

11951196
def emitDeclAux (mod : LLVM.Module llvmctx) (builder : LLVM.Builder llvmctx) (d : Decl) : M llvmctx Unit := do
11961197
let env ← getEnv

src/Lean/Compiler/IR/ExpandResetReuse.lean

Lines changed: 2 additions & 2 deletions
Original file line numberDiff line numberDiff line change
@@ -54,15 +54,15 @@ abbrev Mask := Array (Option VarId)
5454
partial def eraseProjIncForAux (y : VarId) (bs : Array FnBody) (mask : Mask) (keep : Array FnBody) : Array FnBody × Mask :=
5555
let done (_ : Unit) := (bs ++ keep.reverse, mask)
5656
let keepInstr (b : FnBody) := eraseProjIncForAux y bs.pop mask (keep.push b)
57-
if bs.size < 2 then done ()
57+
if h : bs.size < 2 then done ()
5858
else
5959
let b := bs.back!
6060
match b with
6161
| .vdecl _ _ (.sproj _ _ _) _ => keepInstr b
6262
| .vdecl _ _ (.uproj _ _) _ => keepInstr b
6363
| .inc z n c p _ =>
6464
if n == 0 then done () else
65-
let b' := bs[bs.size - 2]!
65+
let b' := bs[bs.size - 2]
6666
match b' with
6767
| .vdecl w _ (.proj i x) _ =>
6868
if w == z && y == x then

src/Lean/Compiler/LCNF/SpecInfo.lean

Lines changed: 5 additions & 5 deletions
Original file line numberDiff line numberDiff line change
@@ -152,8 +152,8 @@ def saveSpecParamInfo (decls : Array Decl) : CompilerM Unit := do
152152
let specArgs? := getSpecializationArgs? (← getEnv) decl.name
153153
let contains (i : Nat) : Bool := specArgs?.getD #[] |>.contains i
154154
let mut paramsInfo : Array SpecParamInfo := #[]
155-
for i in [:decl.params.size] do
156-
let param := decl.params[i]!
155+
for h :i in [:decl.params.size] do
156+
let param := decl.params[i]
157157
let info ←
158158
if contains i then
159159
pure .user
@@ -181,14 +181,14 @@ def saveSpecParamInfo (decls : Array Decl) : CompilerM Unit := do
181181
declsInfo := declsInfo.push paramsInfo
182182
if declsInfo.any fun paramsInfo => paramsInfo.any (· matches .user | .fixedInst | .fixedHO) then
183183
let m := mkFixedParamsMap decls
184-
for i in [:decls.size] do
185-
let decl := decls[i]!
184+
for hi : i in [:decls.size] do
185+
let decl := decls[i]
186186
let mut paramsInfo := declsInfo[i]!
187187
let some mask := m.find? decl.name | unreachable!
188188
trace[Compiler.specialize.info] "{decl.name} {mask}"
189189
paramsInfo := paramsInfo.zipWith mask fun info fixed => if fixed || info matches .user then info else .other
190190
for j in [:paramsInfo.size] do
191-
let mut info := paramsInfo[j]!
191+
let mut info := paramsInfo[j]!
192192
if info matches .fixedNeutral && !hasFwdDeps decl paramsInfo j then
193193
paramsInfo := paramsInfo.set! j .other
194194
if paramsInfo.any fun info => info matches .fixedInst | .fixedHO | .user then

src/Lean/Compiler/LCNF/ToLCNF.lean

Lines changed: 2 additions & 2 deletions
Original file line numberDiff line numberDiff line change
@@ -499,8 +499,8 @@ where
499499
match app with
500500
| .fvar f =>
501501
let mut argsNew := #[]
502-
for i in [arity : args.size] do
503-
argsNew := argsNew.push (← visitAppArg args[i]!)
502+
for h :i in [arity : args.size] do
503+
argsNew := argsNew.push (← visitAppArg args[i])
504504
letValueToArg <| .fvar f argsNew
505505
| .erased | .type .. => return .erased
506506

src/Lean/Compiler/Specialize.lean

Lines changed: 4 additions & 3 deletions
Original file line numberDiff line numberDiff line change
@@ -26,10 +26,11 @@ private def elabSpecArgs (declName : Name) (args : Array Syntax) : MetaM (Array
2626
if let some idx := arg.isNatLit? then
2727
if idx == 0 then throwErrorAt arg "invalid specialization argument index, index must be greater than 0"
2828
let idx := idx - 1
29-
if idx >= argNames.size then
29+
if h : idx >= argNames.size then
3030
throwErrorAt arg "invalid argument index, `{declName}` has #{argNames.size} arguments"
31-
if result.contains idx then throwErrorAt arg "invalid specialization argument index, `{argNames[idx]!}` has already been specified as a specialization candidate"
32-
result := result.push idx
31+
else
32+
if result.contains idx then throwErrorAt arg "invalid specialization argument index, `{argNames[idx]}` has already been specified as a specialization candidate"
33+
result := result.push idx
3334
else
3435
let argName := arg.getId
3536
if let some idx := argNames.indexOf? argName then

src/Lean/Elab/Deriving/Util.lean

Lines changed: 4 additions & 4 deletions
Original file line numberDiff line numberDiff line change
@@ -49,9 +49,9 @@ invoking ``mkInstImplicitBinders `BarClass foo #[`α, `n, `β]`` gives `` `([Bar
4949
def mkInstImplicitBinders (className : Name) (indVal : InductiveVal) (argNames : Array Name) : TermElabM (Array Syntax) :=
5050
forallBoundedTelescope indVal.type indVal.numParams fun xs _ => do
5151
let mut binders := #[]
52-
for i in [:xs.size] do
52+
for h : i in [:xs.size] do
5353
try
54-
let x := xs[i]!
54+
let x := xs[i]
5555
let c ← mkAppM className #[x]
5656
if (← isTypeCorrect c) then
5757
let argName := argNames[i]!
@@ -86,8 +86,8 @@ def mkContext (fnPrefix : String) (typeName : Name) : TermElabM Context := do
8686

8787
def mkLocalInstanceLetDecls (ctx : Context) (className : Name) (argNames : Array Name) : TermElabM (Array (TSyntax ``Parser.Term.letDecl)) := do
8888
let mut letDecls := #[]
89-
for i in [:ctx.typeInfos.size] do
90-
let indVal := ctx.typeInfos[i]!
89+
for h : i in [:ctx.typeInfos.size] do
90+
let indVal := ctx.typeInfos[i]
9191
let auxFunName := ctx.auxFunNames[i]!
9292
let currArgNames ← mkInductArgNames indVal
9393
let numParams := indVal.numParams

src/Lean/Elab/Do.lean

Lines changed: 8 additions & 8 deletions
Original file line numberDiff line numberDiff line change
@@ -796,10 +796,10 @@ Note that we are not restricting the macro power since the
796796
actions to be in the same universe.
797797
-/
798798
private def mkTuple (elems : Array Syntax) : MacroM Syntax := do
799-
if elems.size == 0 then
799+
if elems.size = 0 then
800800
mkUnit
801-
else if elems.size == 1 then
802-
return elems[0]!
801+
else if h : elems.size = 1 then
802+
return elems[0]
803803
else
804804
elems.extract 0 (elems.size - 1) |>.foldrM (init := elems.back!) fun elem tuple =>
805805
``(MProd.mk $elem $tuple)
@@ -831,10 +831,10 @@ def isDoExpr? (doElem : Syntax) : Option Syntax :=
831831
We use this method when expanding the `for-in` notation.
832832
-/
833833
private def destructTuple (uvars : Array Var) (x : Syntax) (body : Syntax) : MacroM Syntax := do
834-
if uvars.size == 0 then
834+
if uvars.size = 0 then
835835
return body
836-
else if uvars.size == 1 then
837-
`(let $(uvars[0]!):ident := $x; $body)
836+
else if h : uvars.size = 1 then
837+
`(let $(uvars[0]):ident := $x; $body)
838838
else
839839
destruct uvars.toList x body
840840
where
@@ -1314,9 +1314,9 @@ private partial def expandLiftMethodAux (inQuot : Bool) (inBinder : Bool) : Synt
13141314
else if liftMethodDelimiter k then
13151315
return stx
13161316
-- For `pure` if-then-else, we only lift `(<- ...)` occurring in the condition.
1317-
else if args.size >= 2 && (k == ``termDepIfThenElse || k == ``termIfThenElse) then do
1317+
else if h : args.size >= 2 (k == ``termDepIfThenElse || k == ``termIfThenElse) then do
13181318
let inAntiquot := stx.isAntiquot && !stx.isEscapedAntiquot
1319-
let arg1 ← expandLiftMethodAux (inQuot && !inAntiquot || stx.isQuot) inBinder args[1]!
1319+
let arg1 ← expandLiftMethodAux (inQuot && !inAntiquot || stx.isQuot) inBinder args[1]
13201320
let args := args.set! 1 arg1
13211321
return Syntax.node i k args
13221322
else if k == ``Parser.Term.liftMethod && !inQuot then withFreshMacroScope do

src/Lean/Elab/Inductive.lean

Lines changed: 10 additions & 10 deletions
Original file line numberDiff line numberDiff line change
@@ -173,15 +173,15 @@ private def checkUnsafe (rs : Array ElabHeaderResult) : TermElabM Unit := do
173173
throwErrorAt r.view.ref "invalid inductive type, cannot mix unsafe and safe declarations in a mutually inductive datatypes"
174174

175175
private def InductiveView.checkLevelNames (views : Array InductiveView) : TermElabM Unit := do
176-
if views.size > 1 then
177-
let levelNames := views[0]!.levelNames
176+
if h : views.size > 1 then
177+
let levelNames := views[0].levelNames
178178
for view in views do
179179
unless view.levelNames == levelNames do
180180
throwErrorAt view.ref "invalid inductive type, universe parameters mismatch in mutually inductive datatypes"
181181

182182
private def ElabHeaderResult.checkLevelNames (rs : Array ElabHeaderResult) : TermElabM Unit := do
183-
if rs.size > 1 then
184-
let levelNames := rs[0]!.levelNames
183+
if h : rs.size > 1 then
184+
let levelNames := rs[0].levelNames
185185
for r in rs do
186186
unless r.levelNames == levelNames do
187187
throwErrorAt r.view.ref "invalid inductive type, universe parameters mismatch in mutually inductive datatypes"
@@ -433,8 +433,8 @@ where
433433
let mut args := e.getAppArgs
434434
unless args.size ≥ params.size do
435435
throwError "unexpected inductive type occurrence{indentExpr e}"
436-
for i in [:params.size] do
437-
let param := params[i]!
436+
for h : i in [:params.size] do
437+
let param := params[i]
438438
let arg := args[i]!
439439
unless (← isDefEq param arg) do
440440
throwError "inductive datatype parameter mismatch{indentExpr arg}\nexpected{indentExpr param}"
@@ -694,8 +694,8 @@ private def collectLevelParamsInInductive (indTypes : List InductiveType) : Arra
694694
private def mkIndFVar2Const (views : Array InductiveView) (indFVars : Array Expr) (levelNames : List Name) : ExprMap Expr := Id.run do
695695
let levelParams := levelNames.map mkLevelParam;
696696
let mut m : ExprMap Expr := {}
697-
for i in [:views.size] do
698-
let view := views[i]!
697+
for h : i in [:views.size] do
698+
let view := views[i]
699699
let indFVar := indFVars[i]!
700700
m := m.insert indFVar (mkConst view.declName levelParams)
701701
return m
@@ -856,9 +856,9 @@ private def mkInductiveDecl (vars : Array Expr) (views : Array InductiveView) :
856856
withInductiveLocalDecls rs fun params indFVars => do
857857
trace[Elab.inductive] "indFVars: {indFVars}"
858858
let mut indTypesArray := #[]
859-
for i in [:views.size] do
859+
for h : i in [:views.size] do
860860
let indFVar := indFVars[i]!
861-
Term.addLocalVarInfo views[i]!.declId indFVar
861+
Term.addLocalVarInfo views[i].declId indFVar
862862
let r := rs[i]!
863863
/- At this point, because of `withInductiveLocalDecls`, the only fvars that are in context are the ones related to the first inductive type.
864864
Because of this, we need to replace the fvars present in each inductive type's header of the mutual block with those of the first inductive.

src/Lean/Elab/LetRec.lean

Lines changed: 2 additions & 2 deletions
Original file line numberDiff line numberDiff line change
@@ -87,8 +87,8 @@ private def elabLetRecDeclValues (view : LetRecView) : TermElabM (Array Expr) :=
8787
view.decls.mapM fun view => do
8888
forallBoundedTelescope view.type view.binderIds.size fun xs type => do
8989
-- Add new info nodes for new fvars. The server will detect all fvars of a binder by the binder's source location.
90-
for i in [0:view.binderIds.size] do
91-
addLocalVarInfo view.binderIds[i]! xs[i]!
90+
for h : i in [0:view.binderIds.size] do
91+
addLocalVarInfo view.binderIds[i] xs[i]!
9292
withDeclName view.declName do
9393
withInfoContext' view.valStx
9494
(mkInfo := (pure <| .inl <| mkBodyInfo view.valStx ·))

src/Lean/Elab/Match.lean

Lines changed: 2 additions & 2 deletions
Original file line numberDiff line numberDiff line change
@@ -282,8 +282,8 @@ where
282282
let dArg := dArgs[i]!
283283
unless (← isDefEq tArg dArg) do
284284
return i :: (← goType tArg dArg)
285-
for i in [info.numParams : tArgs.size] do
286-
let tArg := tArgs[i]!
285+
for h : i in [info.numParams : tArgs.size] do
286+
let tArg := tArgs[i]
287287
let dArg := dArgs[i]!
288288
unless (← isDefEq tArg dArg) do
289289
return i :: (← goIndex tArg dArg)

src/Lean/Elab/Open.lean

Lines changed: 4 additions & 4 deletions
Original file line numberDiff line numberDiff line change
@@ -49,12 +49,12 @@ private def resolveNameUsingNamespacesCore (nss : List Name) (idStx : Syntax) :
4949
exs := exs.push ex
5050
if exs.size == nss.length then
5151
withRef idStx do
52-
if exs.size == 1 then
53-
throw exs[0]!
52+
if h : exs.size = 1 then
53+
throw exs[0]
5454
else
5555
throwErrorWithNestedErrors "failed to open" exs
56-
if result.size == 1 then
57-
return result[0]!
56+
if h : result.size = 1 then
57+
return result[0]
5858
else
5959
withRef idStx do throwError "ambiguous identifier '{idStx.getId}', possible interpretations: {result.map mkConst}"
6060

src/Lean/Elab/PreDefinition/Basic.lean

Lines changed: 2 additions & 2 deletions
Original file line numberDiff line numberDiff line change
@@ -244,8 +244,8 @@ def checkCodomainsLevel (preDefs : Array PreDefinition) : MetaM Unit := do
244244
lambdaTelescope preDef.value fun xs _ => return xs.size
245245
forallBoundedTelescope preDefs[0]!.type arities[0]! fun _ type₀ => do
246246
let u₀ ← getLevel type₀
247-
for i in [1:preDefs.size] do
248-
forallBoundedTelescope preDefs[i]!.type arities[i]! fun _ typeᵢ =>
247+
for h : i in [1:preDefs.size] do
248+
forallBoundedTelescope preDefs[i].type arities[i]! fun _ typeᵢ =>
249249
unless ← isLevelDefEq u₀ (← getLevel typeᵢ) do
250250
withOptions (fun o => pp.sanitizeNames.set o false) do
251251
throwError m!"invalid mutual definition, result types must be in the same universe " ++

src/Lean/Elab/PreDefinition/TerminationArgument.lean

Lines changed: 2 additions & 2 deletions
Original file line numberDiff line numberDiff line change
@@ -53,10 +53,10 @@ def TerminationArgument.elab (funName : Name) (type : Expr) (arity extraParams :
5353
(hint : TerminationBy) : TermElabM TerminationArgument := withDeclName funName do
5454
assert! extraParams ≤ arity
5555

56-
if hint.vars.size > extraParams then
56+
if h : hint.vars.size > extraParams then
5757
let mut msg := m!"{parameters hint.vars.size} bound in `termination_by`, but the body of " ++
5858
m!"{funName} only binds {parameters extraParams}."
59-
if let `($ident:ident) := hint.vars[0]! then
59+
if let `($ident:ident) := hint.vars[0] then
6060
if ident.getId.isSuffixOf funName then
6161
msg := msg ++ m!" (Since Lean v4.6.0, the `termination_by` clause no longer " ++
6262
"expects the function name here.)"

src/Lean/Elab/PreDefinition/TerminationHint.lean

Lines changed: 2 additions & 2 deletions
Original file line numberDiff line numberDiff line change
@@ -90,10 +90,10 @@ lambda of `value`, and throws appropriate errors.
9090
-/
9191
def TerminationBy.checkVars (funName : Name) (extraParams : Nat) (tb : TerminationBy) : MetaM Unit := do
9292
unless tb.synthetic do
93-
if tb.vars.size > extraParams then
93+
if h : tb.vars.size > extraParams then
9494
let mut msg := m!"{parameters tb.vars.size} bound in `termination_by`, but the body of " ++
9595
m!"{funName} only binds {parameters extraParams}."
96-
if let `($ident:ident) := tb.vars[0]! then
96+
if let `($ident:ident) := tb.vars[0] then
9797
if ident.getId.isSuffixOf funName then
9898
msg := msg ++ m!" (Since Lean v4.6.0, the `termination_by` clause no longer " ++
9999
"expects the function name here.)"

0 commit comments

Comments
 (0)