Skip to content

Commit

Permalink
chore: cleaunp grind tests (leanprover#6616)
Browse files Browse the repository at this point in the history
Tests using `logInfo` were taking an additional two seconds on my
machine. This is a performance issue with the old code generator, where
we spend all this time specializing the logging functions for `GoalM`. I
have not checked whether the new code generator is also affected by this
performance issue.

Here is a small example that exposes the issue:
```lean
import Lean

set_option profiler true
open Lean Meta Grind in
def test (e : Expr): GoalM Unit := do
  logInfo e
```

cc @zwarich
  • Loading branch information
leodemoura authored and luisacicolini committed Jan 21, 2025
1 parent 24321ad commit 5c65173
Show file tree
Hide file tree
Showing 7 changed files with 57 additions and 91 deletions.
9 changes: 6 additions & 3 deletions tests/lean/run/grind_canon_insts.lean
Original file line number Diff line number Diff line change
Expand Up @@ -53,11 +53,13 @@ theorem left_comm [CommMonoid α] (a b c : α) : a * (b * c) = b * (a * c) := by
open Lean Meta Elab Tactic Grind in
def fallback : Fallback := do
let nodes ← filterENodes fun e => return e.self.isAppOf ``HMul.hMul
logInfo (nodes.toList.map (·.self))
trace[Meta.debug] "{nodes.toList.map (·.self)}"
(← get).mvarId.admit

set_option trace.Meta.debug true

/--
info: [b * c, a * (b * c), d * (b * c)]
info: [Meta.debug] [b * c, a * (b * c), d * (b * c)]
-/
#guard_msgs (info) in
example (a b c d : Nat) : b * (a * c) = d * (b * c) → False := by
Expand All @@ -68,7 +70,8 @@ example (a b c d : Nat) : b * (a * c) = d * (b * c) → False := by
set_option pp.notation false in
set_option pp.explicit true in
/--
info: [@HMul.hMul Nat Nat Nat (@instHMul Nat instMulNat) b a, @HMul.hMul Nat Nat Nat (@instHMul Nat instMulNat) b d]
info: [Meta.debug] [@HMul.hMul Nat Nat Nat (@instHMul Nat instMulNat) b a,
@HMul.hMul Nat Nat Nat (@instHMul Nat instMulNat) b d]
-/
#guard_msgs (info) in
example (a b c d : Nat) : b * a = d * b → False := by
Expand Down
5 changes: 3 additions & 2 deletions tests/lean/run/grind_canon_types.lean
Original file line number Diff line number Diff line change
Expand Up @@ -6,12 +6,13 @@ def f (a : α) := a
open Lean Meta Grind in
def fallback : Fallback := do
let nodes ← filterENodes fun e => return e.self.isAppOf ``f
logInfo (nodes.toList.map (·.self))
trace[Meta.debug] "{nodes.toList.map (·.self)}"
(← get).mvarId.admit

set_option trace.Meta.debug true
set_option pp.explicit true
/--
info: [@f Nat a, @f Nat b]
info: [Meta.debug] [@f Nat a, @f Nat b]
-/
#guard_msgs (info) in
example (a b c d : Nat) : @f Nat a = b → @f (g Nat) a = c → @f (g Nat) b = d → a = b → False := by
Expand Down
11 changes: 6 additions & 5 deletions tests/lean/run/grind_congr.lean
Original file line number Diff line number Diff line change
Expand Up @@ -8,35 +8,36 @@ open Lean Meta Grind in
def fallback : Fallback := do
let #[n, _] ← filterENodes fun e => return e.self.isAppOf ``f | unreachable!
let eqc ← getEqc n.self
logInfo eqc
trace[Meta.debug] eqc
(← get).mvarId.admit

set_option trace.Meta.debug true
set_option grind.debug true
set_option grind.debug.proofs true

/--
info: [d, f b, c, f a]
info: [Meta.debug] [d, f b, c, f a]
-/
#guard_msgs (info) in
example (a b c d : Nat) : a = b → f a = c → f b = d → False := by
grind on_failure fallback

/--
info: [d, f b, c, f a]
info: [Meta.debug] [d, f b, c, f a]
-/
#guard_msgs (info) in
example (a b c d : Nat) : f a = c → f b = d → a = b → False := by
grind on_failure fallback

/--
info: [d, f (g b), c, f (g a)]
info: [Meta.debug] [d, f (g b), c, f (g a)]
-/
#guard_msgs (info) in
example (a b c d e : Nat) : f (g a) = c → f (g b) = d → a = e → b = e → False := by
grind on_failure fallback

/--
info: [d, f (g b), c, f v]
info: [Meta.debug] [d, f (g b), c, f v]
-/
#guard_msgs (info) in
example (a b c d e v : Nat) : f v = c → f (g b) = d → a = e → b = e → v = g a → False := by
Expand Down
30 changes: 13 additions & 17 deletions tests/lean/run/grind_nested_proofs.lean
Original file line number Diff line number Diff line change
Expand Up @@ -5,12 +5,13 @@ def f (α : Type) [Add α] (a : α) := a + a + a
open Lean Meta Grind in
def fallback : Fallback := do
let nodes ← filterENodes fun e => return e.self.isAppOf ``Lean.Grind.nestedProof
logInfo (nodes.toList.map (·.self))
trace[Meta.debug] "{nodes.toList.map (·.self)}"
let nodes ← filterENodes fun e => return e.self.isAppOf ``GetElem.getElem
let [_, n, _] := nodes.toList | unreachable!
logInfo (← getEqc n.self)
trace[Meta.debug] "{← getEqc n.self}"
(← get).mvarId.admit

set_option trace.Meta.debug true
set_option grind.debug true
set_option grind.debug.proofs true

Expand All @@ -20,26 +21,21 @@ The following test relies on `grind` `nestedProof` wrapper to
detect equalities between array access terms.
-/

/-
info: [Lean.Grind.nestedProof (i < a.toList.length) (_example.proof_1 i j a b h1 h2),
Lean.Grind.nestedProof (j < a.toList.length) h1,
Lean.Grind.nestedProof (j < b.toList.length) h]
---
info: [a[i], b[j], a[j]]
---
warning: declaration uses 'sorry'
/--
info: [Meta.debug] [Lean.Grind.nestedProof (i < a.toList.length) (_example.proof_1 i j a b h1 h2),
Lean.Grind.nestedProof (j < a.toList.length) h1,
Lean.Grind.nestedProof (j < b.toList.length) h]
[Meta.debug] [a[i], b[j], a[j]]
-/
-- #guard_msgs in

#guard_msgs (info) in
example (i j : Nat) (a b : Array Nat) (h1 : j < a.size) (h : j < b.size) (h2 : i ≤ j) : a[i] < a[j] + b[j] → i = j → a = b → False := by
grind on_failure fallback

/--
info: [Lean.Grind.nestedProof (i < a.toList.length) (_example.proof_1 i j a b h1 h2),
Lean.Grind.nestedProof (j < a.toList.length) h1,
Lean.Grind.nestedProof (j < b.toList.length) h]
---
info: [a[i], a[j]]
info: [Meta.debug] [Lean.Grind.nestedProof (i < a.toList.length) (_example.proof_1 i j a b h1 h2),
Lean.Grind.nestedProof (j < a.toList.length) h1,
Lean.Grind.nestedProof (j < b.toList.length) h]
[Meta.debug] [a[i], a[j]]
-/
#guard_msgs (info) in
example (i j : Nat) (a b : Array Nat) (h1 : j < a.size) (h : j < b.size) (h2 : i ≤ j) : a[i] < a[j] + b[j] → i = j → False := by
Expand Down
5 changes: 3 additions & 2 deletions tests/lean/run/grind_norm_levels.lean
Original file line number Diff line number Diff line change
Expand Up @@ -5,12 +5,13 @@ def g {α : Sort u} (a : α) := a
open Lean Meta Grind in
def fallback : Fallback := do
let nodes ← filterENodes fun e => return e.self.isAppOf ``g
logInfo (nodes.toList.map (·.self))
trace[Meta.debug] "{nodes.toList.map (·.self)}"
(← get).mvarId.admit

-- `grind` final state must contain only two `g`-applications
set_option trace.Meta.debug true in
/--
info: [g (a, b), g (g (a, b))]
info: [Meta.debug] [g (a, b), g (g (a, b))]
-/
#guard_msgs (info) in
example {β : Type v} {α : Type u} (a c : α) (b d : β) : g.{max u v + 1} (a, b) = (c, d) → g (g.{max (u+1) (v+1)} (a, b)) = (c, d) → False := by
Expand Down
61 changes: 26 additions & 35 deletions tests/lean/run/grind_propagate_connectives.lean
Original file line number Diff line number Diff line change
@@ -1,107 +1,98 @@
import Lean.Meta.Tactic.Grind

set_option trace.Meta.debug true

open Lean Meta Grind in
def fallback : Fallback := do
let trueExprs := (← filterENodes fun e => return e.self.isFVar && (← isEqTrue e.self)).toList.map (·.self)
let falseExprs := (← filterENodes fun e => return e.self.isFVar && (← isEqFalse e.self)).toList.map (·.self)
logInfo m!"true: {trueExprs}"
logInfo m!"false: {falseExprs}"
trace[Meta.debug] "true: {trueExprs}"
trace[Meta.debug] "false: {falseExprs}"
forEachEqcRoot fun n => do
unless (← isProp n.self) || (← isType n.self) || n.size == 1 do
let eqc ← getEqc n.self
logInfo eqc
trace[Meta.debug] eqc
(← get).mvarId.admit

set_option grind.debug true
set_option grind.debug.proofs true

/--
info: true: [q, w]
---
info: false: [p, r]
info: [Meta.debug] true: [q, w]
[Meta.debug] false: [p, r]
-/
#guard_msgs (info) in
example : (p ∨ (q ∧ ¬r ∧ w)) → ¬p → False := by
grind on_failure fallback


/--
info: true: [r]
---
info: false: [p, q]
info: [Meta.debug] true: [r]
[Meta.debug] false: [p, q]
-/
#guard_msgs (info) in
example : (p ∨ q ∨ r) → (p ∨ ¬q) → ¬p → False := by
grind on_failure fallback


/--
info: true: [r]
---
info: false: [p₁, q]
info: [Meta.debug] true: [r]
[Meta.debug] false: [p₁, q]
-/
#guard_msgs (info) in
example : ((p₁ ∧ p₂) ∨ q ∨ r) → (p₁ ∨ ¬q) → p₁ = False → False := by
grind on_failure fallback

/--
info: true: [r]
---
info: false: [p₂, q]
info: [Meta.debug] true: [r]
[Meta.debug] false: [p₂, q]
-/
#guard_msgs (info) in
example : ((p₁ ∧ p₂) ∨ q ∨ r) → ((p₂ ∧ p₁) ∨ ¬q) → p₂ = False → False := by
grind on_failure fallback

/--
info: true: [q, r]
---
info: false: [p]
info: [Meta.debug] true: [q, r]
[Meta.debug] false: [p]
-/
#guard_msgs (info) in
example (p q r : Prop) : p ∨ (q ↔ r) → p = False → q → False := by
grind on_failure fallback

/--
info: true: [r]
---
info: false: [p, s]
info: [Meta.debug] true: [r]
[Meta.debug] false: [p, s]
-/
#guard_msgs (info) in
example (p q r : Prop) : p ∨ ¬(s ∨ (p ↔ r)) → p = False → False := by
grind on_failure fallback

/--
info: true: [p]
---
info: false: []
---
info: [a, b]
info: [Meta.debug] true: [p]
[Meta.debug] false: []
[Meta.debug] [a, b]
-/
#guard_msgs (info) in
example (p : Prop) (a : Vector Nat 5) (b : Vector Nat 6) : (p → HEq a b) → p → False := by
grind on_failure fallback

/--
info: true: [p, q]
---
info: false: [r]
info: [Meta.debug] true: [p, q]
[Meta.debug] false: [r]
-/
#guard_msgs (info) in
example (p q r : Prop) : p ∨ (q ↔ r) → q → ¬r → False := by
grind on_failure fallback

/--
info: hello world
---
info: true: [p, q]
---
info: false: [r]
info: [Meta.debug] hello world
[Meta.debug] true: [p, q]
[Meta.debug] false: [r]
-/
#guard_msgs (info) in
example (p q r : Prop) : p ∨ (q ↔ r) → ¬r → q → False := by
grind on_failure do
Lean.logInfo "hello world"
trace[Meta.debug] "hello world"
fallback

example (a b : Nat) : (a = b) = (b = a) := by
Expand Down
27 changes: 0 additions & 27 deletions tests/lean/run/grind_revertAll.lean

This file was deleted.

0 comments on commit 5c65173

Please sign in to comment.