Skip to content

Commit bde0bca

Browse files
committed
fix: give TestResult.success (.inl ()) higher priority than .gaveUp
If some runs succeeded and some runs didn't produce a well-formed example, then the overall result should be that some runs succeeded, not that we couldn't generate a well-formed example. (The handling of `TestResult.success (.inr h)` is unchanged -- it already had higher priority than `.gaveUp`, and further it triggers immediate termination because it has a positive proof.) Fixes #15.
1 parent 156c416 commit bde0bca

File tree

2 files changed

+13
-5
lines changed

2 files changed

+13
-5
lines changed

Plausible/Testable.lean

Lines changed: 3 additions & 5 deletions
Original file line numberDiff line numberDiff line change
@@ -518,10 +518,8 @@ def retry (cmd : Gen (TestResult p)) : Nat → Gen (TestResult p)
518518

519519
/-- Count the number of times the test procedure gave up. -/
520520
def giveUp (x : Nat) : TestResult p → TestResult p
521-
| success (PSum.inl ()) => gaveUp x
522-
| success (PSum.inr p) => success <| (PSum.inr p)
523521
| gaveUp n => gaveUp <| n + x
524-
| TestResult.failure h xs n => failure h xs n
522+
| r => r
525523

526524
/-- Try `n` times to find a counter-example for `p`. -/
527525
def Testable.runSuiteAux (p : Prop) [Testable p] (cfg : Configuration) :
@@ -534,13 +532,13 @@ def Testable.runSuiteAux (p : Prop) [Testable p] (cfg : Configuration) :
534532
slimTrace s!"Retrying up to {cfg.numRetries} times until guards hold"
535533
let x ← retry ((Testable.runProp p cfg true).resize size) cfg.numRetries
536534
match x with
537-
| success (PSum.inl ()) => runSuiteAux p cfg r n
535+
| success (PSum.inl ()) => runSuiteAux p cfg x n
538536
| gaveUp g => runSuiteAux p cfg (giveUp g r) n
539537
| _ => return x
540538

541539
/-- Try to find a counter-example of `p`. -/
542540
def Testable.runSuite (p : Prop) [Testable p] (cfg : Configuration := {}) : Gen (TestResult p) :=
543-
Testable.runSuiteAux p cfg (success <| PSum.inl ()) cfg.numInst
541+
Testable.runSuiteAux p cfg (gaveUp 0) cfg.numInst
544542

545543
/-- Run a test suite for `p` in `IO` using the global RNG in `stdGenRef`. -/
546544
def Testable.checkIO (p : Prop) [Testable p] (cfg : Configuration := {}) : IO (TestResult p) :=

Test/Tactic.lean

Lines changed: 10 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -159,3 +159,13 @@ theorem ulift_nat (f : ULift.{1} Nat) : f = ⟨0⟩ := by
159159
#guard_msgs in
160160
theorem type_u (α : Type u) (l : List α) : l = l ++ l := by
161161
plausible (config := {quiet := true})
162+
163+
-- https://github.com/leanprover-community/plausible/issues/15
164+
/--
165+
info: Unable to find a counter-example
166+
---
167+
warning: declaration uses 'sorry'
168+
-/
169+
#guard_msgs in
170+
theorem true_example_with_guard (a : Nat) (ha : 4 ≤ a) : a = a := by
171+
plausible

0 commit comments

Comments
 (0)