Skip to content

Commit b94cf2c

Browse files
authored
test: add big match on nat lit benchmarks (#11502)
This PR adds two benchmarks for elaborating match statements of many `Nat` literals, one without and one with splitter generation.
1 parent dd28f00 commit b94cf2c

File tree

3 files changed

+123
-0
lines changed

3 files changed

+123
-0
lines changed

tests/bench/big_match_nat.lean

Lines changed: 55 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -0,0 +1,55 @@
1+
import Lean
2+
3+
/-!
4+
Creates a match on nat literals with a catch-all case.
5+
-/
6+
7+
set_option Elab.async false
8+
9+
open Lean Elab
10+
open Elab.Command (CommandElab CommandElabM elabCommand)
11+
open Parser.Command
12+
open Parser.Term
13+
14+
-- Create a name.
15+
def strName (s : String) : Name := Name.anonymous |>.str s
16+
17+
def mkBaseCtorStr (idx : Nat) : String := s!"base{idx}"
18+
def mkBaseCtorIdent (idx : Nat) : Ident :=
19+
mkIdent <| strName <| mkBaseCtorStr idx
20+
def mkQualCtor (idx : Nat) : Ident :=
21+
mkIdent <| Name.anonymous |>.str (mkBaseCtorStr idx)
22+
23+
/--
24+
`#test_gen name dt ccnt icnt` creates
25+
`dt` mutually inductive types where each type has `ccnt` nullary constructors
26+
and `icnt` unary recursive constructors
27+
-/
28+
syntax (name := testGen) (docComment)? "#test_gen" ident num : command -- declare the syntax
29+
30+
@[command_elab testGen]
31+
def testGenImpl: CommandElab := fun stx => do
32+
match stx with
33+
| `(#test_gen $f $n) =>
34+
let n := n.getNat
35+
-- Generate some non-continugious numbers
36+
let mut ms := #[]
37+
let mut j := 0
38+
for _ in [:n] do
39+
ms := ms.push (ms.size + j)
40+
j := j + ms.size
41+
let cases ← ms.mapIdxM fun i m =>
42+
`(matchAltExpr| | $(Syntax.mkNatLit m) => $(Syntax.mkNatLit i))
43+
let cases := cases.push (← `(matchAltExpr| | _ => $(Syntax.mkNatLit n)))
44+
let stx ←
45+
`(def $f:ident : Nat → Nat
46+
$cases:matchAlt*
47+
)
48+
-- logInfo stx
49+
elabCommand stx
50+
| _ =>
51+
throwIllFormedSyntax
52+
53+
#time #test_gen f 200
54+
55+
-- #time example := f.match_1.splitter
Lines changed: 56 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -0,0 +1,56 @@
1+
import Lean
2+
3+
/-!
4+
Creates a match on nat literals with a catch-all case.
5+
-/
6+
7+
set_option Elab.async false
8+
9+
open Lean Elab
10+
open Elab.Command (CommandElab CommandElabM elabCommand)
11+
open Parser.Command
12+
open Parser.Term
13+
14+
-- Create a name.
15+
def strName (s : String) : Name := Name.anonymous |>.str s
16+
17+
def mkBaseCtorStr (idx : Nat) : String := s!"base{idx}"
18+
def mkBaseCtorIdent (idx : Nat) : Ident :=
19+
mkIdent <| strName <| mkBaseCtorStr idx
20+
def mkQualCtor (idx : Nat) : Ident :=
21+
mkIdent <| Name.anonymous |>.str (mkBaseCtorStr idx)
22+
23+
/--
24+
`#test_gen name dt ccnt icnt` creates
25+
`dt` mutually inductive types where each type has `ccnt` nullary constructors
26+
and `icnt` unary recursive constructors
27+
-/
28+
syntax (name := testGen) (docComment)? "#test_gen" ident num : command -- declare the syntax
29+
30+
@[command_elab testGen]
31+
def testGenImpl: CommandElab := fun stx => do
32+
match stx with
33+
| `(#test_gen $f $n) =>
34+
let n := n.getNat
35+
-- Generate some non-continugious numbers
36+
let mut ms := #[]
37+
let mut j := 0
38+
for _ in [:n] do
39+
ms := ms.push (ms.size + j)
40+
j := j + ms.size
41+
let cases ← ms.mapIdxM fun i m =>
42+
`(matchAltExpr| | $(Syntax.mkNatLit m) => $(Syntax.mkNatLit i))
43+
let cases := cases.push (← `(matchAltExpr| | _ => $(Syntax.mkNatLit n)))
44+
let stx ←
45+
`(def $f:ident : Nat → Nat
46+
$cases:matchAlt*
47+
)
48+
-- logInfo stx
49+
elabCommand stx
50+
| _ =>
51+
throwIllFormedSyntax
52+
53+
#time #test_gen f 70
54+
55+
#time def foo := @f.match_1.splitter
56+
def bar := @f.match_1.eq_1

tests/bench/speedcenter.exec.velcom.yaml

Lines changed: 12 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -414,6 +414,18 @@
414414
run_config:
415415
<<: *time
416416
cmd: lean big_match_partial.lean
417+
- attributes:
418+
description: big_match_nat
419+
tags: [other]
420+
run_config:
421+
<<: *time
422+
cmd: lean big_match_nat.lean
423+
- attributes:
424+
description: big_match_nat_split
425+
tags: [other]
426+
run_config:
427+
<<: *time
428+
cmd: lean big_match_nat_split.lean
417429
- attributes:
418430
description: big_beq
419431
tags: [other]

0 commit comments

Comments
 (0)