Skip to content

Commit 5bd331e

Browse files
authored
perf: kernel-optimize Mon.mul (#11422)
This PR uses a kernel-reduction optimized variant of Mon.mul in grind.
1 parent 856825a commit 5bd331e

File tree

2 files changed

+75
-7
lines changed

2 files changed

+75
-7
lines changed

src/Init/Grind/Ring/CommSolver.lean

Lines changed: 75 additions & 6 deletions
Original file line numberDiff line numberDiff line change
@@ -180,6 +180,53 @@ where
180180
else
181181
.mult { x := pw₁.x, k := pw₁.k + pw₂.k } (go fuel m₁ m₂)
182182

183+
noncomputable def Mon.mul_k : Mon → Mon → Mon :=
184+
Nat.rec
185+
(fun m₁ m₂ => concat m₁ m₂)
186+
(fun _ ih m₁ m₂ =>
187+
Mon.rec (t := m₂)
188+
m₁
189+
(fun pw₂ m₂' _ => Mon.rec (t := m₁)
190+
m₂
191+
(fun pw₁ m₁' _ =>
192+
Bool.rec (t := pw₁.varLt pw₂)
193+
(Bool.rec (t := pw₂.varLt pw₁)
194+
(.mult { x := pw₁.x, k := Nat.add pw₁.k pw₂.k } (ih m₁' m₂'))
195+
(.mult pw₂ (ih (.mult pw₁ m₁') m₂')))
196+
(.mult pw₁ (ih m₁' (.mult pw₂ m₂'))))))
197+
hugeFuel
198+
199+
theorem Mon.mul_k_eq_mul : Mon.mul_k m₁ m₂ = Mon.mul m₁ m₂ := by
200+
unfold mul_k mul
201+
generalize hugeFuel = fuel
202+
fun_induction mul.go
203+
· rfl
204+
· rfl
205+
case case3 m₂ _ =>
206+
cases m₂
207+
· contradiction
208+
· dsimp
209+
case case4 fuel pw₁ m₁ pw₂ m₂ h ih =>
210+
dsimp only
211+
rw [h]
212+
dsimp only
213+
rw [ih]
214+
case case5 fuel pw₁ m₁ pw₂ m₂ h₁ h₂ ih =>
215+
dsimp only
216+
rw [h₁]
217+
dsimp only
218+
rw [h₂]
219+
dsimp only
220+
rw [ih]
221+
case case6 fuel pw₁ m₁ pw₂ m₂ h₁ h₂ ih =>
222+
dsimp only
223+
rw [h₁]
224+
dsimp only
225+
rw [h₂]
226+
dsimp only
227+
rw [ih]
228+
rfl
229+
183230
def Mon.mul_nc (m₁ m₂ : Mon) : Mon :=
184231
match m₁ with
185232
| .unit => m₂
@@ -190,6 +237,28 @@ def Mon.degree : Mon → Nat
190237
| .unit => 0
191238
| .mult pw m => pw.k + degree m
192239

240+
noncomputable def Mon.degree_k : Mon → Nat :=
241+
Nat.rec
242+
(fun m => m.degree)
243+
(fun _ ih m =>
244+
Mon.rec (t := m)
245+
0
246+
(fun pw m' _ => Nat.add pw.k (ih m')))
247+
hugeFuel
248+
249+
theorem Mon.degree_k_eq_degree : Mon.degree_k m = Mon.degree m := by
250+
unfold degree_k
251+
generalize hugeFuel = fuel
252+
induction fuel generalizing m with
253+
| zero => rfl
254+
| succ fuel ih =>
255+
conv => rhs; unfold degree
256+
split
257+
· rfl
258+
· dsimp only
259+
rw [← ih]
260+
rfl
261+
193262
def Var.revlex (x y : Var) : Ordering :=
194263
bif x.blt y then .gt
195264
else bif y.blt x then .lt
@@ -270,7 +339,7 @@ noncomputable def Mon.grevlex_k (m₁ m₂ : Mon) : Ordering :=
270339
Bool.rec
271340
(Bool.rec .gt .lt (Nat.blt m₁.degree m₂.degree))
272341
(revlex_k m₁ m₂)
273-
(Nat.beq m₁.degree m₂.degree)
342+
(Nat.beq m₁.degree_k m₂.degree_k)
274343

275344
theorem Mon.revlex_k_eq_revlex (m₁ m₂ : Mon) : m₁.revlex_k m₂ = m₁.revlex m₂ := by
276345
unfold revlex_k revlex
@@ -302,18 +371,18 @@ theorem Mon.grevlex_k_eq_grevlex (m₁ m₂ : Mon) : m₁.grevlex_k m₂ = m₁.
302371
next h =>
303372
have h₁ : Nat.blt m₁.degree m₂.degree = true := by simp [h]
304373
have h₂ : Nat.beq m₁.degree m₂.degree = false := by rw [← Bool.not_eq_true, Nat.beq_eq]; omega
305-
simp [h₁, h₂]
374+
simp [degree_k_eq_degree, h₁, h₂]
306375
next h =>
307376
split
308377
next h' =>
309378
have h₂ : Nat.beq m₁.degree m₂.degree = true := by rw [Nat.beq_eq, h']
310-
simp [h₂]
379+
simp [degree_k_eq_degree, h₂]
311380
next h' =>
312381
have h₁ : Nat.blt m₁.degree m₂.degree = false := by
313382
rw [← Bool.not_eq_true, Nat.blt_eq]; assumption
314383
have h₂ : Nat.beq m₁.degree m₂.degree = false := by
315384
rw [← Bool.not_eq_true, Nat.beq_eq]; assumption
316-
simp [h₁, h₂]
385+
simp [degree_k_eq_degree, h₁, h₂]
317386

318387
inductive Poly where
319388
| num (k : Int)
@@ -481,7 +550,7 @@ noncomputable def Poly.mulMon_k (k : Int) (m : Mon) (p : Poly) : Poly :=
481550
(Bool.rec
482551
(Poly.rec
483552
(fun k' => Bool.rec (.add (Int.mul k k') m (.num 0)) (.num 0) (Int.beq' k' 0))
484-
(fun k' m' _ ih => .add (Int.mul k k') (m.mul m') ih)
553+
(fun k' m' _ ih => .add (Int.mul k k') (m.mul_k m') ih)
485554
p)
486555
(p.mulConst_k k)
487556
(Mon.beq' m .unit))
@@ -511,7 +580,7 @@ noncomputable def Poly.mulMon_k (k : Int) (m : Mon) (p : Poly) : Poly :=
511580
next =>
512581
have h : Int.beq' k 0 = false := by simp [*]
513582
simp [h]
514-
next ih => simp [← ih]
583+
next ih => simp [← ih, Mon.mul_k_eq_mul]
515584

516585
def Poly.mulMon_nc (k : Int) (m : Mon) (p : Poly) : Poly :=
517586
bif k == 0 then

tests/lean/run/grind_ring_5.lean

Lines changed: 0 additions & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -1,7 +1,6 @@
11
module
22
open Lean.Grind
33

4-
54
example {α} [CommRing α] [IsCharP α 0] (d t c : α) (d_inv PSO3_inv : α)
65
40 : d^2 * (d + t - d * t - 2) *
76
(d + t + d * t) = 0)

0 commit comments

Comments
 (0)