Skip to content

Commit 3ee2842

Browse files
authored
feat: remove runtime bounds checks and partial from qsort (#6241)
This PR refactors `Array.qsort` to remove runtime array bounds checks, and avoids the use of `partial`. We use the `Vector` API, along with auto_params, to avoid having to write any proofs. The new code benchmarks indistinguishably from the old.
1 parent 7b8504c commit 3ee2842

File tree

6 files changed

+52
-27
lines changed

6 files changed

+52
-27
lines changed

src/Init/Data/Array/QSort.lean

Lines changed: 27 additions & 27 deletions
Original file line numberDiff line numberDiff line change
@@ -4,46 +4,46 @@ Released under Apache 2.0 license as described in the file LICENSE.
44
Authors: Leonardo de Moura
55
-/
66
prelude
7-
import Init.Data.Array.Basic
7+
import Init.Data.Vector.Basic
88
import Init.Data.Ord
99

1010
namespace Array
11-
-- TODO: remove the [Inhabited α] parameters as soon as we have the tactic framework for automating proof generation and using Array.fget
1211

13-
def qpartition (as : Array α) (lt : α → α → Bool) (lo hi : Nat) : Nat × Array α :=
14-
if h : as.size = 0 then (0, as) else have : Inhabited α := ⟨as[0]'(by revert h; cases as.size <;> simp)⟩ -- TODO: remove
12+
private def qpartition {n} (as : Vector α n) (lt : α → α → Bool) (lo hi : Nat)
13+
(hlo : lo < n := by omega) (hhi : hi < n := by omega) : {n : Nat // lo ≤ n} × Vector α n :=
1514
let mid := (lo + hi) / 2
16-
let as := if lt (as.get! mid) (as.get! lo) then as.swapIfInBounds lo mid else as
17-
let as := if lt (as.get! hi) (as.get! lo) then as.swapIfInBounds lo hi else as
18-
let as := if lt (as.get! mid) (as.get! hi) then as.swapIfInBounds mid hi else as
19-
let pivot := as.get! hi
20-
let rec loop (as : Array α) (i j : Nat) :=
15+
let as := if lt as[mid] as[lo] then as.swap lo mid else as
16+
let as := if lt as[hi] as[lo] then as.swap lo hi else as
17+
let as := if lt as[mid] as[hi] then as.swap mid hi else as
18+
let pivot := as[hi]
19+
let rec loop (as : Vector α n) (i j : Nat)
20+
(ilo : lo ≤ i := by omega) (jh : j < n := by omega) (w : i ≤ j := by omega) :=
2121
if h : j < hi then
22-
if lt (as.get! j) pivot then
23-
let as := as.swapIfInBounds i j
24-
loop as (i+1) (j+1)
22+
if lt as[j] pivot then
23+
loop (as.swap i j) (i+1) (j+1)
2524
else
2625
loop as i (j+1)
2726
else
28-
let as := as.swapIfInBounds i hi
29-
(i, as)
30-
termination_by hi - j
31-
decreasing_by all_goals simp_wf; decreasing_trivial_pre_omega
27+
(⟨i, ilo⟩, as.swap i hi)
3228
loop as lo lo
3329

34-
@[inline] partial def qsort (as : Array α) (lt : α → α → Bool) (low := 0) (high := as.size - 1) : Array α :=
35-
let rec @[specialize] sort (as : Array α) (low high : Nat) :=
36-
if low < high then
37-
let p := qpartition as lt low high;
38-
-- TODO: fix `partial` support in the equation compiler, it breaks if we use `let (mid, as) := partition as lt low high`
39-
let mid := p.1
40-
let as := p.2
41-
if mid >= high then as
30+
@[inline] def qsort (as : Array α) (lt : α → α → Bool := by exact (· < ·))
31+
(low := 0) (high := as.size - 1) : Array α :=
32+
let rec @[specialize] sort {n} (as : Vector α n) (lo hi : Nat)
33+
(hlo : lo < n := by omega) (hhi : hi < n := by omega) :=
34+
if h₁ : lo < hi then
35+
let ⟨⟨mid, hmid⟩, as⟩ := qpartition as lt lo hi
36+
if h₂ : mid ≥ hi then
37+
as
4238
else
43-
let as := sort as low mid
44-
sort as (mid+1) high
39+
sort (sort as lo mid) (mid+1) hi
4540
else as
46-
sort as low high
41+
if h : as.size = 0 then
42+
as
43+
else
44+
let low := min low (as.size - 1)
45+
let high := min high (as.size - 1)
46+
sort ⟨as, rfl⟩ low high |>.toArray
4747

4848
set_option linter.unusedVariables.funArgs false in
4949
/--

tests/bench/qsort/.gitignore

Lines changed: 1 addition & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -0,0 +1 @@
1+
/.lake

tests/bench/qsort/Main.lean

Lines changed: 15 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -0,0 +1,15 @@
1+
set_option linter.unusedVariables false
2+
3+
abbrev Elem := UInt32
4+
5+
def badRand (seed : Elem) : Elem :=
6+
seed * 1664525 + 1013904223
7+
8+
def mkRandomArray : Nat → Elem → Array Elem → Array Elem
9+
| 0, seed, as => as
10+
| i+1, seed, as => mkRandomArray i (badRand seed) (as.push seed)
11+
12+
def main (args : List String) : IO UInt32 := do
13+
let a := mkRandomArray 4000000 0 (Array.mkEmpty 4000000)
14+
IO.println (a.qsort (· < ·)).size
15+
return 0

tests/bench/qsort/README.md

Lines changed: 1 addition & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -0,0 +1 @@
1+
# insertionSort

tests/bench/qsort/lakefile.toml

Lines changed: 7 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -0,0 +1,7 @@
1+
name = "qsort"
2+
version = "0.1.0"
3+
defaultTargets = ["qsort"]
4+
5+
[[lean_exe]]
6+
name = "qsort"
7+
root = "Main"

tests/bench/qsort/lean-toolchain

Lines changed: 1 addition & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -0,0 +1 @@
1+
lean4

0 commit comments

Comments
 (0)