From c2ce721d823f0efe41a555b2ddb63123f1a6c512 Mon Sep 17 00:00:00 2001 From: Siddhartha Gadgil Date: Mon, 18 Nov 2024 19:47:37 +0530 Subject: [PATCH 01/30] steps towards meta-testing; needs redesign --- Plausible/MetaTestable.lean | 586 ++++++++++++++++++++++++++++++++++++ 1 file changed, 586 insertions(+) create mode 100644 Plausible/MetaTestable.lean diff --git a/Plausible/MetaTestable.lean b/Plausible/MetaTestable.lean new file mode 100644 index 0000000..5fb8a07 --- /dev/null +++ b/Plausible/MetaTestable.lean @@ -0,0 +1,586 @@ +/- +Copyright (c) 2022 Henrik Böving. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Henrik Böving, Simon Hudon +-/ +import Lean.Elab.Tactic.Config +import Plausible.Sampleable +import Plausible.Testable +open Lean + +/-! +# `MetaTestable` Class + +MetaTestable propositions have a procedure that can generate counter-examples +together with a proof that they invalidate the proposition. + +This is a port of the Haskell QuickCheck library. + +## Creating Customized Instances + +The type classes `MetaTestable`, `SampleableExt` and `Shrinkable` are the +means by which `Plausible` creates samples and tests them. For instance, +the proposition `∀ i j : Nat, i ≤ j` has a `MetaTestable` instance because `Nat` +is sampleable and `i ≤ j` is decidable. Once `Plausible` finds the `MetaTestable` +instance, it can start using the instance to repeatedly creating samples +and checking whether they satisfy the property. Once it has found a +counter-example it will then use a `Shrinkable` instance to reduce the +example. This allows the user to create new instances and apply +`Plausible` to new situations. + +### What do I do if I'm testing a property about my newly defined type? + +Let us consider a type made for a new formalization: + +```lean +structure MyType where + x : Nat + y : Nat + h : x ≤ y + deriving Repr +``` + +How do we test a property about `MyType`? For instance, let us consider +`MetaTestable.check <| ∀ a b : MyType, a.y ≤ b.x → a.x ≤ b.y`. Writing this +property as is will give us an error because we do not have an instance +of `Shrinkable MyType` and `SampleableExt MyType`. We can define one as follows: + +```lean +instance : Shrinkable MyType where + shrink := fun ⟨x, y, _⟩ => + let proxy := Shrinkable.shrink (x, y - x) + proxy.map (fun (fst, snd) => ⟨fst, fst + snd, by omega⟩) + +instance : SampleableExt MyType := + SampleableExt.mkSelfContained do + let x ← SampleableExt.interpSample Nat + let xyDiff ← SampleableExt.interpSample Nat + return ⟨x, x + xyDiff, by omega⟩ +``` + +Again, we take advantage of the fact that other types have useful +`Shrinkable` implementations, in this case `Prod`. + +## Main definitions + +* `MetaTestable` class +* `MetaTestable.check`: a way to test a proposition using random examples + +## References + +* https://hackage.haskell.org/package/QuickCheck + +-/ + +namespace Plausible + +section Matching +open Lean Meta +/-! +# Matching propositions of specific forms +-/ +def existsProp? (e: Expr) : MetaM (Option Expr × Option Expr) := do + let u ← mkFreshLevelMVar + let α ← mkFreshExprMVar (mkSort u) + let prop := mkSort levelZero + let fmly ← mkArrow α prop + let mvar ← mkFreshExprMVar (some fmly) + let e' ← mkAppM ``Exists #[mvar] + if ← isDefEq e' e then + return (← Lean.getExprMVarAssignment? mvar.mvarId!, ← Lean.getExprMVarAssignment? α.mvarId!) + else + return (none, none) + +def orProp? (e: Expr) : MetaM (Option Expr × Option Expr) := do + let prop := mkSort levelZero + let α ← mkFreshExprMVar prop + let β ← mkFreshExprMVar prop + let e' ← mkAppM ``Or #[α, β] + if ← isDefEq e' e then + return (← Lean.getExprMVarAssignment? α.mvarId!, ← Lean.getExprMVarAssignment? β.mvarId!) + else + return (none, none) + +def andProp? (e: Expr) : MetaM (Option Expr × Option Expr) := do + let prop := mkSort levelZero + let α ← mkFreshExprMVar prop + let β ← mkFreshExprMVar prop + let e' ← mkAppM ``And #[α, β] + if ← isDefEq e' e then + return (← Lean.getExprMVarAssignment? α.mvarId!, ← Lean.getExprMVarAssignment? β.mvarId!) + else + return (none, none) + +#check mkFreshLevelMVar + +def forallProp? (e: Expr) : MetaM (Option Expr × Option Expr) := do + let u ← mkFreshLevelMVar + let α ← mkFreshExprMVar (mkSort u) + let prop := mkSort levelZero + let fmly ← mkArrow α prop + let mvar ← mkFreshExprMVar (some fmly) + let e' ← withLocalDeclD `x α fun x => do + let y ← mkAppM' mvar #[x] + mkForallFVars #[x] y + if ← isDefEq e' e then + return (← Lean.getExprMVarAssignment? mvar.mvarId!, ← Lean.getExprMVarAssignment? α.mvarId!) + else + return (none, none) + +def impProp? (e: Expr) : MetaM (Option Expr × Option Expr) := do + let prop := mkSort levelZero + let α ← mkFreshExprMVar prop + let β ← mkFreshExprMVar prop + let e' ← mkArrow α β + if ← isDefEq e' e then + return (← Lean.getExprMVarAssignment? α.mvarId!, ← Lean.getExprMVarAssignment? β.mvarId!) + else + return (none, none) + +def eqlProp? (e: Expr): MetaM (Option (Expr × Expr × Expr)) := do + let level ← mkFreshLevelMVar + let u := mkSort level + let α ← mkFreshExprMVar u + let a ← mkFreshExprMVar α + let b ← mkFreshExprMVar α + let e' ← mkAppM ``Eq #[a, b] + if ← isDefEq e' e then + let α? ← Lean.getExprMVarAssignment? α.mvarId! + let a? ← Lean.getExprMVarAssignment? a.mvarId! + let b? ← Lean.getExprMVarAssignment? b.mvarId! + let triple : Option (Expr × Expr × Expr) := do + let α ← α? + let a ← a? + let b ← b? + return (α, a, b) + return triple + else + return none + +def equality? (e: Expr): MetaM (Option (Expr × Expr × Expr)) := do + let level ← mkFreshLevelMVar + let u := mkSort level + let α ← mkFreshExprMVar u + let a ← mkFreshExprMVar α + let b ← mkFreshExprMVar α + let e' ← mkAppM ``Eq #[a, b] + let c ← mkFreshExprMVar e' + if ← isDefEq c e then + let α? ← Lean.getExprMVarAssignment? α.mvarId! + let a? ← Lean.getExprMVarAssignment? a.mvarId! + let b? ← Lean.getExprMVarAssignment? b.mvarId! + let triple : Option (Expr × Expr × Expr) := do + let α ← α? + let a ← a? + let b ← b? + return (α, a, b) + return triple + else + return none +end Matching + + +/-- Result of trying to disprove `p` -/ +inductive MetaTestResult (p : Prop) where + /-- + Succeed when we find another example satisfying `p`. + In `success h`, `h` is an optional proof of the proposition. + Without the proof, all we know is that we found one example + where `p` holds. With a proof, the one test was sufficient to + prove that `p` holds and we do not need to keep finding examples. + -/ + | success : Unit ⊕' p → MetaTestResult p + /-- + Give up when a well-formed example cannot be generated. + `gaveUp n` tells us that `n` invalid examples were tried. + -/ + | gaveUp : Nat → MetaTestResult p + /-- + There was a mismatch between `p` and the expression representing `p` + -/ + | mismatch : String → MetaTestResult p + /-- + A counter-example to `p`; the strings specify values for the relevant variables. + `failure h vs n` also carries a proof that `p` does not hold. This way, we can + guarantee that there will be no false positive. The last component, `n`, + is the number of times that the counter-example was shrunk. + -/ + | failure : ¬ p → Expr → List String → Nat → MetaTestResult p + deriving Inhabited + + +/-- `MetaTestable p` uses random examples to try to disprove `p`. -/ +class MetaTestable (p : Prop) where + run (cfg : Configuration) (minimize : Bool) (propExpr : Expr) : Gen (MetaTestResult p) + + +namespace MetaTestResult + +def toString : MetaTestResult p → String + | success (PSum.inl _) => "success (no proof)" + | success (PSum.inr _) => "success (proof)" + | gaveUp n => s!"gave {n} times" + | failure _ _ counters _ => s!"failed {counters}" + | mismatch s => s!"Mismatch: {s}" + +instance : ToString (MetaTestResult p) := ⟨toString⟩ + +/-- Applicative combinator proof carrying test results. -/ +def combine {p q : Prop} : Unit ⊕' (p → q) → Unit ⊕' p → Unit ⊕' q + | PSum.inr f, PSum.inr proof => PSum.inr <| f proof + | _, _ => PSum.inl () + +/-- Combine the test result for properties `p` and `q` to create a test for their conjunction. -/ +def and : MetaTestResult p → MetaTestResult q → MetaTestResult (p ∧ q) + | failure h pf xs n, _ => + failure (fun h2 => h h2.left) pf xs n + | _, failure h pf xs n => failure (fun h2 => h h2.right) pf xs n + | success h1, success h2 => success <| combine (combine (PSum.inr And.intro) h1) h2 + | gaveUp n, gaveUp m => gaveUp <| n + m + | gaveUp n, _ => gaveUp n + | _, gaveUp n => gaveUp n + | mismatch s, _ => mismatch s + | _, mismatch s => mismatch s + +/-- Combine the test result for properties `p` and `q` to create a test for their disjunction. -/ +def or : MetaTestResult p → MetaTestResult q → MetaTestResult (p ∨ q) + | failure h1 pf1 xs n, failure h2 pf2 ys m => + let h3 := fun h => + match h with + | Or.inl h3 => h1 h3 + | Or.inr h3 => h2 h3 + failure h3 pf1 (xs ++ ys) (n + m) + | success h, _ => success <| combine (PSum.inr Or.inl) h + | _, success h => success <| combine (PSum.inr Or.inr) h + | gaveUp n, gaveUp m => gaveUp <| n + m + | gaveUp n, _ => gaveUp n + | _, gaveUp n => gaveUp n + | mismatch s, _ => mismatch s + | _, mismatch s => mismatch s + +/-- If `q → p`, then `¬ p → ¬ q` which means that testing `p` can allow us +to find counter-examples to `q`. -/ +def imp (h : q → p) (r : MetaTestResult p) + (p : Unit ⊕' (p → q) := PSum.inl ()) : MetaTestResult q := + match r with + | failure h2 pf xs n => failure (mt h h2) pf xs n + | success h2 => success <| combine p h2 + | gaveUp n => gaveUp n + | mismatch s => mismatch s + +/-- Test `q` by testing `p` and proving the equivalence between the two. -/ +def iff (h : q ↔ p) (r : MetaTestResult p) : MetaTestResult q := + imp h.mp r (PSum.inr h.mpr) + +/-- When we assign a value to a universally quantified variable, +we record that value using this function so that our counter-examples +can be informative. -/ +def addInfo (x : String) (h : q → p) (r : MetaTestResult p) + (p : Unit ⊕' (p → q) := PSum.inl ()) : MetaTestResult q := + if let failure h2 pf xs n := r then + failure (mt h h2) pf (x :: xs) n + else + imp h r p + +/-- Add some formatting to the information recorded by `addInfo`. -/ +def addVarInfo {γ : Type _} [Repr γ] (var : String) (x : γ) (h : q → p) (r : MetaTestResult p) + (p : Unit ⊕' (p → q) := PSum.inl ()) : MetaTestResult q := + addInfo s!"{var} := {repr x}" h r p + +def isFailure : MetaTestResult p → Bool + | failure .. => true + | _ => false + +end MetaTestResult + + +namespace MetaTestable + +open MetaTestResult + +def runProp (p : Prop) [MetaTestable p] : Configuration → Bool → Expr → Gen (MetaTestResult p) := MetaTestable.run + +/-- A `dbgTrace` with special formatting -/ +def slimTrace {m : Type → Type _} [Pure m] (s : String) : m PUnit := + dbgTrace s!"[Plausible: {s}]" (fun _ => pure ()) + +instance andTestable [MetaTestable p] [MetaTestable q] : MetaTestable (p ∧ q) where + run := fun cfg min e => do + -- let pair ← andProp? e + -- match ← andProp? e with + -- | (some e₁, some e₂) => do + let xp ← runProp p cfg min e + let xq ← runProp q cfg min e + return and xp xq + -- | _ => return sorry + +instance orTestable [MetaTestable p] [MetaTestable q] : MetaTestable (p ∨ q) where + run := fun cfg min e => do + let xp ← runProp p cfg min e + -- As a little performance optimization we can just not run the second + -- test if the first succeeds + match xp with + | success (PSum.inl h) => return success (PSum.inl h) + | success (PSum.inr h) => return success (PSum.inr <| Or.inl h) + | _ => + let xq ← runProp q cfg min e + return or xp xq + +instance iffTestable [MetaTestable ((p ∧ q) ∨ (¬ p ∧ ¬ q))] : MetaTestable (p ↔ q) where + run := fun cfg min e => do + let h ← runProp ((p ∧ q) ∨ (¬ p ∧ ¬ q)) cfg min e + have := by + constructor + · intro h + simp [h, Classical.em] + · intro h + rcases h with ⟨hleft, hright⟩ | ⟨hleft, hright⟩ <;> simp [hleft, hright] + return iff this h + +variable {var : String} + +instance decGuardTestable [PrintableProp p] [Decidable p] {β : p → Prop} [∀ h, MetaTestable (β h)] : + MetaTestable (NamedBinder var <| ∀ h, β h) where + run := fun cfg min e => do + if h : p then + let res := runProp (β h) cfg min e + let s := printProp p + (fun r => addInfo s!"guard: {s}" (· <| h) r (PSum.inr <| fun q _ => q)) <$> res + else if cfg.traceDiscarded || cfg.traceSuccesses then + let res := fun _ => return gaveUp 1 + let s := printProp p + slimTrace s!"discard: Guard {s} does not hold"; res + else + return gaveUp 1 + +instance forallTypesTestable {f : Type → Prop} [MetaTestable (f Int)] : + MetaTestable (NamedBinder var <| ∀ x, f x) where + run := fun cfg min e => do + let r ← runProp (f Int) cfg min e + return addVarInfo var "Int" (· <| Int) r + +-- TODO: only in mathlib: @[pp_with_univ] +instance (priority := 100) forallTypesULiftTestable.{u} + {f : Type u → Prop} [MetaTestable (f (ULift.{u} Int))] : + MetaTestable (NamedBinder var <| ∀ x, f x) where + run cfg min e := do + let r ← runProp (f (ULift Int)) cfg min e + pure <| addVarInfo var "ULift Int" (· <| ULift Int) r + +/-- +Format the counter-examples found in a test failure. +-/ +def formatFailure (s : String) (xs : List String) (n : Nat) : String := + let counter := String.intercalate "\n" xs + let parts := [ + "\n===================", + s, + counter, + s!"({n} shrinks)", + "-------------------" + ] + String.intercalate "\n" parts + +/-- +Increase the number of shrinking steps in a test result. +-/ +def addShrinks (n : Nat) : MetaTestResult p → MetaTestResult p + | MetaTestResult.failure p pf xs m => MetaTestResult.failure p pf xs (m + n) + | p => p + +universe u in +instance {α : Type u} {m : Type u → Type _} [Pure m] : Inhabited (OptionT m α) := + ⟨(pure none : m (Option α))⟩ + +variable {α : Sort _} + +/-- Shrink a counter-example `x` by using `Shrinkable.shrink x`, picking the first +candidate that falsifies a property and recursively shrinking that one. +The process is guaranteed to terminate because `shrink x` produces +a proof that all the values it produces are smaller (according to `SizeOf`) +than `x`. -/ +partial def minimizeAux [SampleableExt α] {β : α → Prop} [∀ x, MetaTestable (β x)] (cfg : Configuration) + (var : String) (x : SampleableExt.proxy α) (n : Nat) : + OptionT Gen (Σ x, MetaTestResult (β (SampleableExt.interp x))) := do + let candidates := SampleableExt.shrink.shrink x + if cfg.traceShrinkCandidates then + slimTrace s!"Candidates for {var} := {repr x}:\n {repr candidates}" + for candidate in candidates do + if cfg.traceShrinkCandidates then + slimTrace s!"Trying {var} := {repr candidate}" + let res ← OptionT.lift <| MetaTestable.runProp (β (SampleableExt.interp candidate)) cfg true sorry + if res.isFailure then + if cfg.traceShrink then + slimTrace s!"{var} shrunk to {repr candidate} from {repr x}" + let currentStep := OptionT.lift <| return Sigma.mk candidate (addShrinks (n + 1) res) + let nextStep := minimizeAux cfg var candidate (n + 1) + return ← (nextStep <|> currentStep) + if cfg.traceShrink then + slimTrace s!"No shrinking possible for {var} := {repr x}" + failure + +/-- Once a property fails to hold on an example, look for smaller counter-examples +to show the user. -/ +def minimize [SampleableExt α] {β : α → Prop} [∀ x, MetaTestable (β x)] (cfg : Configuration) + (var : String) (x : SampleableExt.proxy α) (r : MetaTestResult (β <| SampleableExt.interp x)) : + Gen (Σ x, MetaTestResult (β <| SampleableExt.interp x)) := do + if cfg.traceShrink then + slimTrace "Shrink" + slimTrace s!"Attempting to shrink {var} := {repr x}" + let res ← OptionT.run <| minimizeAux cfg var x 0 + return res.getD ⟨x, r⟩ + +/-- Test a universal property by creating a sample of the right type and instantiating the +bound variable with it. -/ +instance varTestable [SampleableExt α] {β : α → Prop} [∀ x, MetaTestable (β x)] : + MetaTestable (NamedBinder var <| ∀ x : α, β x) where + run := fun cfg min e => do + let x ← SampleableExt.sample + if cfg.traceSuccesses || cfg.traceDiscarded then + slimTrace s!"{var} := {repr x}" + let r ← MetaTestable.runProp (β <| SampleableExt.interp x) cfg false e + let ⟨finalX, finalR⟩ ← + if isFailure r then + if cfg.traceSuccesses then + slimTrace s!"{var} := {repr x} is a failure" + if min then + minimize cfg var x r + else + pure ⟨x, r⟩ + else + pure ⟨x, r⟩ + return addVarInfo var finalX (· <| SampleableExt.interp finalX) finalR + +/-- Test a universal property about propositions -/ +instance propVarTestable {β : Prop → Prop} [∀ b : Bool, MetaTestable (β b)] : + MetaTestable (NamedBinder var <| ∀ p : Prop, β p) +where + run := fun cfg min e => + imp (fun h (b : Bool) => h b) <$> MetaTestable.runProp (NamedBinder var <| ∀ b : Bool, β b) cfg min e + +instance (priority := high) unusedVarTestable {β : Prop} [Nonempty α] [MetaTestable β] : + MetaTestable (NamedBinder var (α → β)) +where + run := fun cfg min e => do + if cfg.traceDiscarded || cfg.traceSuccesses then + slimTrace s!"{var} is unused" + let r ← MetaTestable.runProp β cfg min e + let finalR := addInfo s!"{var} is irrelevant (unused)" id r + return imp (· <| Classical.ofNonempty) finalR (PSum.inr <| fun x _ => x) + +instance (priority := 2000) subtypeVarTestable {p : α → Prop} {β : α → Prop} + [∀ x, PrintableProp (p x)] + [∀ x, MetaTestable (β x)] + [SampleableExt (Subtype p)] {var'} : + MetaTestable (NamedBinder var <| (x : α) → NamedBinder var' <| p x → β x) where + run cfg min e := + letI (x : Subtype p) : MetaTestable (β x) := + { run := fun cfg min e => do + let r ← MetaTestable.runProp (β x.val) cfg min e + return addInfo s!"guard: {printProp (p x)} (by construction)" id r (PSum.inr id) } + do + let r ← @MetaTestable.run (∀ x : Subtype p, β x.val) (@varTestable var _ _ _ _) cfg min e + have := by simp [Subtype.forall, NamedBinder] + return iff this r + +instance (priority := low) decidableTestable {p : Prop} [PrintableProp p] [Decidable p] : + MetaTestable p where + run := fun _ _ _ => + if h : p then + return success (PSum.inr h) + else + let s := printProp p + return failure h sorry [s!"issue: {s} does not hold"] 0 + +end MetaTestable + +#check PrintableProp + + +section IO +open MetaTestResult + +namespace MetaTestable +/-- Execute `cmd` and repeat every time the result is `gaveUp` (at most `n` times). -/ +def retry (cmd : Rand (MetaTestResult p)) : Nat → Rand (MetaTestResult p) + | 0 => return MetaTestResult.gaveUp 1 + | n+1 => do + let r ← cmd + match r with + | .success hp => return success hp + | .failure h pf xs n => return failure h pf xs n + | .gaveUp _ => retry cmd n + | .mismatch s => return .mismatch s + +/-- Count the number of times the test procedure gave up. -/ +def giveUp (x : Nat) : MetaTestResult p → MetaTestResult p + | success (PSum.inl ()) => gaveUp x + | success (PSum.inr p) => success <| (PSum.inr p) + | gaveUp n => gaveUp <| n + x + | MetaTestResult.failure h pf xs n => failure h pf xs n + | mismatch _ => gaveUp x + +end MetaTestable + +/-- Try `n` times to find a counter-example for `p`. -/ +def MetaTestable.runSuiteAux (p : Prop) [MetaTestable p] (cfg : Configuration) : + MetaTestResult p → Nat → Rand (MetaTestResult p) + | r, 0 => return r + | r, n+1 => do + let size := (cfg.numInst - n - 1) * cfg.maxSize / cfg.numInst + if cfg.traceSuccesses then + slimTrace s!"New sample" + slimTrace s!"Retrying up to {cfg.numRetries} times until guards hold" + let x ← retry (ReaderT.run (MetaTestable.runProp p cfg true sorry) ⟨size⟩) cfg.numRetries + match x with + | success (PSum.inl ()) => runSuiteAux p cfg r n + | gaveUp g => runSuiteAux p cfg (giveUp g r) n + | _ => return x + +/-- Try to find a counter-example of `p`. -/ +def MetaTestable.runSuite (p : Prop) [MetaTestable p] (cfg : Configuration := {}) : Rand (MetaTestResult p) := + MetaTestable.runSuiteAux p cfg (success <| PSum.inl ()) cfg.numInst + +/-- Run a test suite for `p` in `BaseIO` using the global RNG in `stdGenRef`. -/ +def MetaTestable.checkIO (p : Prop) [MetaTestable p] (cfg : Configuration := {}) : BaseIO (MetaTestResult p) := + letI : MonadLift Id BaseIO := ⟨fun f => return Id.run f⟩ + match cfg.randomSeed with + | none => runRand (MetaTestable.runSuite p cfg) + | some seed => runRandWith seed (MetaTestable.runSuite p cfg) + +end IO + +open Decorations in +/-- Run a test suite for `p` and throw an exception if `p` does not hold. -/ +def MetaTestable.check (p : Prop) (cfg : Configuration := {}) + (p' : Decorations.DecorationsOf p := by mk_decorations) [MetaTestable p'] : Lean.CoreM PUnit := do + match ← MetaTestable.checkIO p' cfg with + | MetaTestResult.success _ => if !cfg.quiet then Lean.logInfo "Unable to find a counter-example" + | MetaTestResult.gaveUp n => + if !cfg.quiet then + let msg := s!"Gave up after failing to generate values that fulfill the preconditions {n} times." + Lean.logWarning msg + | MetaTestResult.mismatch s => + if !cfg.quiet then + let msg := s!"Mismatch: {s}" + Lean.logWarning msg + | MetaTestResult.failure _ _ xs n => + let msg := "Found a counter-example!" + if cfg.quiet then + Lean.throwError msg + else + Lean.throwError <| formatFailure msg xs n + +-- #eval MetaTestable.check (∀ (x y z a : Nat) (h1 : 3 < x) (h2 : 3 < y), x - y = y - x) +-- Configuration.verbose +-- #eval MetaTestable.check (∀ x : Nat, ∀ y : Nat, x + y = y + x) Configuration.verbose +-- #eval MetaTestable.check (∀ (x : (Nat × Nat)), x.fst - x.snd - 10 = x.snd - x.fst - 10) +-- Configuration.verbose +-- #eval MetaTestable.check (∀ (x : Nat) (h : 10 < x), 5 < x) Configuration.verbose + +macro tk:"#test " e:term : command => `(command| #eval%$tk MetaTestable.check $e) + +-- #test ∀ (x : Nat) (h : 5 < x), 10 < x +-- #test ∀ (x : Nat) (h : 10 < x), 5 < x + +end Plausible From 28a7b3400a9aec944598e1cabad089c3a82d4845 Mon Sep 17 00:00:00 2001 From: Siddhartha Gadgil Date: Mon, 18 Nov 2024 21:48:23 +0530 Subject: [PATCH 02/30] monad lifts --- Plausible/MetaTestable.lean | 10 ++++++++++ 1 file changed, 10 insertions(+) diff --git a/Plausible/MetaTestable.lean b/Plausible/MetaTestable.lean index 5fb8a07..35737ba 100644 --- a/Plausible/MetaTestable.lean +++ b/Plausible/MetaTestable.lean @@ -179,6 +179,16 @@ def equality? (e: Expr): MetaM (Option (Expr × Expr × Expr)) := do return none end Matching +open Lean Meta +abbrev MRand := RandT MetaM + +instance : MonadLift Rand MRand where + monadLift := fun x s => return x.run s + +abbrev MGen (α : Type) := ReaderT (ULift Nat) MRand α + +instance : MonadLift Gen MGen where + monadLift := fun x n => x.run n.down /-- Result of trying to disprove `p` -/ inductive MetaTestResult (p : Prop) where From f4fd9ad47e97db3e153df2e7c1308fe191858768 Mon Sep 17 00:00:00 2001 From: Siddhartha Gadgil Date: Tue, 19 Nov 2024 14:23:36 +0530 Subject: [PATCH 03/30] some cases presumably correct now --- Plausible/MetaTestable.lean | 152 ++++++++++++++++++++---------------- 1 file changed, 86 insertions(+), 66 deletions(-) diff --git a/Plausible/MetaTestable.lean b/Plausible/MetaTestable.lean index 35737ba..5bd9996 100644 --- a/Plausible/MetaTestable.lean +++ b/Plausible/MetaTestable.lean @@ -190,6 +190,8 @@ abbrev MGen (α : Type) := ReaderT (ULift Nat) MRand α instance : MonadLift Gen MGen where monadLift := fun x n => x.run n.down +#synth MonadLift MetaM MRand + /-- Result of trying to disprove `p` -/ inductive MetaTestResult (p : Prop) where /-- @@ -221,7 +223,7 @@ inductive MetaTestResult (p : Prop) where /-- `MetaTestable p` uses random examples to try to disprove `p`. -/ class MetaTestable (p : Prop) where - run (cfg : Configuration) (minimize : Bool) (propExpr : Expr) : Gen (MetaTestResult p) + run (cfg : Configuration) (minimize : Bool) (propExpr : Expr) : MGen (MetaTestResult p) namespace MetaTestResult @@ -241,32 +243,55 @@ def combine {p q : Prop} : Unit ⊕' (p → q) → Unit ⊕' p → Unit ⊕' q | _, _ => PSum.inl () /-- Combine the test result for properties `p` and `q` to create a test for their conjunction. -/ -def and : MetaTestResult p → MetaTestResult q → MetaTestResult (p ∧ q) - | failure h pf xs n, _ => - failure (fun h2 => h h2.left) pf xs n - | _, failure h pf xs n => failure (fun h2 => h h2.right) pf xs n - | success h1, success h2 => success <| combine (combine (PSum.inr And.intro) h1) h2 - | gaveUp n, gaveUp m => gaveUp <| n + m - | gaveUp n, _ => gaveUp n - | _, gaveUp n => gaveUp n - | mismatch s, _ => mismatch s - | _, mismatch s => mismatch s +def and : MetaTestResult p → MetaTestResult q → Expr → MetaM (MetaTestResult (p ∧ q)) + | failure h pf xs n, _, e => do + match ← andProp? e with + | (some e₁, some e₂) => + let pf' ← withLocalDeclD `h e fun h => do + let x ← mkAppOptM ``And.left #[e₁, e₂, h] + let e' ← mkAppM' pf #[x] + mkLambdaFVars #[h] e' + return failure (fun h2 => h h2.left) pf' xs n + | (_, _) => throwError "Expected an `And` proposition" + | _, failure h pf xs n, e => do + match ← andProp? e with + | (some e₁, some e₂) => + let pf' ← withLocalDeclD `h e fun h => do + let x ← mkAppOptM ``And.right #[e₁, e₂, h] + let e' ← mkAppM' pf #[x] + mkLambdaFVars #[h] e' + return failure (fun h2 => h h2.right) pf' xs n + | (_, _) => throwError "Expected an `And` proposition" + + | success h1, success h2, _ => return success <| combine (combine (PSum.inr And.intro) h1) h2 + | gaveUp n, gaveUp m, _ => return gaveUp <| n + m + | gaveUp n, _, _ => return gaveUp n + | _, gaveUp n, _ => return gaveUp n + | mismatch s, _, _ => return mismatch s + | _, mismatch s, _ => return mismatch s + /-- Combine the test result for properties `p` and `q` to create a test for their disjunction. -/ -def or : MetaTestResult p → MetaTestResult q → MetaTestResult (p ∨ q) - | failure h1 pf1 xs n, failure h2 pf2 ys m => - let h3 := fun h => - match h with - | Or.inl h3 => h1 h3 - | Or.inr h3 => h2 h3 - failure h3 pf1 (xs ++ ys) (n + m) - | success h, _ => success <| combine (PSum.inr Or.inl) h - | _, success h => success <| combine (PSum.inr Or.inr) h - | gaveUp n, gaveUp m => gaveUp <| n + m - | gaveUp n, _ => gaveUp n - | _, gaveUp n => gaveUp n - | mismatch s, _ => mismatch s - | _, mismatch s => mismatch s +def or : MetaTestResult p → MetaTestResult q → Expr → MetaM (MetaTestResult (p ∨ q)) + | failure h1 pf1 xs n, failure h2 pf2 ys m, e => do + match ← orProp? e with + | (some α, some β) => + let motive ← withLocalDeclD `h e fun h => do + mkLambdaFVars #[h] <| mkConst ``False + let pf ← mkAppOptM ``Or.rec #[α, β, motive, pf1, pf2] + let h3 := fun h => + match h with + | Or.inl h3 => h1 h3 + | Or.inr h3 => h2 h3 + return failure h3 pf (xs ++ ys) (n + m) + | (_, _) => throwError "Expected an `Or` proposition" + | success h, _, _ => return success <| combine (PSum.inr Or.inl) h + | _, success h, _ => return success <| combine (PSum.inr Or.inr) h + | gaveUp n, gaveUp m, _ => return gaveUp <| n + m + | gaveUp n, _, _ => return gaveUp n + | _, gaveUp n, _ => return gaveUp n + | mismatch s, _, _ => return mismatch s + | _, mismatch s, _ => return mismatch s /-- If `q → p`, then `¬ p → ¬ q` which means that testing `p` can allow us to find counter-examples to `q`. -/ @@ -279,7 +304,7 @@ def imp (h : q → p) (r : MetaTestResult p) | mismatch s => mismatch s /-- Test `q` by testing `p` and proving the equivalence between the two. -/ -def iff (h : q ↔ p) (r : MetaTestResult p) : MetaTestResult q := +def iff (h : q ↔ p) (r : MetaTestResult p) : MetaTestResult q := imp h.mp r (PSum.inr h.mpr) /-- When we assign a value to a universally quantified variable, @@ -308,7 +333,7 @@ namespace MetaTestable open MetaTestResult -def runProp (p : Prop) [MetaTestable p] : Configuration → Bool → Expr → Gen (MetaTestResult p) := MetaTestable.run +def runProp (p : Prop) [MetaTestable p] : Configuration → Bool → Expr → MGen (MetaTestResult p) := MetaTestable.run /-- A `dbgTrace` with special formatting -/ def slimTrace {m : Type → Type _} [Pure m] (s : String) : m PUnit := @@ -316,13 +341,9 @@ def slimTrace {m : Type → Type _} [Pure m] (s : String) : m PUnit := instance andTestable [MetaTestable p] [MetaTestable q] : MetaTestable (p ∧ q) where run := fun cfg min e => do - -- let pair ← andProp? e - -- match ← andProp? e with - -- | (some e₁, some e₂) => do let xp ← runProp p cfg min e let xq ← runProp q cfg min e - return and xp xq - -- | _ => return sorry + and xp xq e instance orTestable [MetaTestable p] [MetaTestable q] : MetaTestable (p ∨ q) where run := fun cfg min e => do @@ -334,7 +355,7 @@ instance orTestable [MetaTestable p] [MetaTestable q] : MetaTestable (p ∨ q) w | success (PSum.inr h) => return success (PSum.inr <| Or.inl h) | _ => let xq ← runProp q cfg min e - return or xp xq + or xp xq e instance iffTestable [MetaTestable ((p ∧ q) ∨ (¬ p ∧ ¬ q))] : MetaTestable (p ↔ q) where run := fun cfg min e => do @@ -409,9 +430,9 @@ candidate that falsifies a property and recursively shrinking that one. The process is guaranteed to terminate because `shrink x` produces a proof that all the values it produces are smaller (according to `SizeOf`) than `x`. -/ -partial def minimizeAux [SampleableExt α] {β : α → Prop} [∀ x, MetaTestable (β x)] (cfg : Configuration) +partial def minimizeAux [SampleableExt α] [ToExpr (SampleableExt.proxy α)] {β : α → Prop} [∀ x, MetaTestable (β x)] (cfg : Configuration) (var : String) (x : SampleableExt.proxy α) (n : Nat) : - OptionT Gen (Σ x, MetaTestResult (β (SampleableExt.interp x))) := do + OptionT MGen (Σ x, MetaTestResult (β (SampleableExt.interp x))) := do let candidates := SampleableExt.shrink.shrink x if cfg.traceShrinkCandidates then slimTrace s!"Candidates for {var} := {repr x}:\n {repr candidates}" @@ -431,9 +452,9 @@ partial def minimizeAux [SampleableExt α] {β : α → Prop} [∀ x, MetaTestab /-- Once a property fails to hold on an example, look for smaller counter-examples to show the user. -/ -def minimize [SampleableExt α] {β : α → Prop} [∀ x, MetaTestable (β x)] (cfg : Configuration) +def minimize [SampleableExt α] [ToExpr (SampleableExt.proxy α)] {β : α → Prop} [∀ x, MetaTestable (β x)] (cfg : Configuration) (var : String) (x : SampleableExt.proxy α) (r : MetaTestResult (β <| SampleableExt.interp x)) : - Gen (Σ x, MetaTestResult (β <| SampleableExt.interp x)) := do + MGen (Σ x, MetaTestResult (β <| SampleableExt.interp x)) := do if cfg.traceShrink then slimTrace "Shrink" slimTrace s!"Attempting to shrink {var} := {repr x}" @@ -442,7 +463,7 @@ def minimize [SampleableExt α] {β : α → Prop} [∀ x, MetaTestable (β x)] /-- Test a universal property by creating a sample of the right type and instantiating the bound variable with it. -/ -instance varTestable [SampleableExt α] {β : α → Prop} [∀ x, MetaTestable (β x)] : +instance varTestable [SampleableExt α] [ToExpr (SampleableExt.proxy α)] {β : α → Prop} [∀ x, MetaTestable (β x)] : MetaTestable (NamedBinder var <| ∀ x : α, β x) where run := fun cfg min e => do let x ← SampleableExt.sample @@ -465,8 +486,8 @@ instance varTestable [SampleableExt α] {β : α → Prop} [∀ x, MetaTestable instance propVarTestable {β : Prop → Prop} [∀ b : Bool, MetaTestable (β b)] : MetaTestable (NamedBinder var <| ∀ p : Prop, β p) where - run := fun cfg min e => - imp (fun h (b : Bool) => h b) <$> MetaTestable.runProp (NamedBinder var <| ∀ b : Bool, β b) cfg min e + run := fun cfg min e => sorry -- Need to interpolate `Prop` and `Bool` + -- imp (fun h (b : Bool) => h b) <$> MetaTestable.runProp (NamedBinder var <| ∀ b : Bool, β b) cfg min e instance (priority := high) unusedVarTestable {β : Prop} [Nonempty α] [MetaTestable β] : MetaTestable (NamedBinder var (α → β)) @@ -481,17 +502,17 @@ where instance (priority := 2000) subtypeVarTestable {p : α → Prop} {β : α → Prop} [∀ x, PrintableProp (p x)] [∀ x, MetaTestable (β x)] - [SampleableExt (Subtype p)] {var'} : + [SampleableExt (Subtype p)] [ToExpr (SampleableExt.proxy (Subtype p))] {var'} : MetaTestable (NamedBinder var <| (x : α) → NamedBinder var' <| p x → β x) where - run cfg min e := - letI (x : Subtype p) : MetaTestable (β x) := - { run := fun cfg min e => do - let r ← MetaTestable.runProp (β x.val) cfg min e - return addInfo s!"guard: {printProp (p x)} (by construction)" id r (PSum.inr id) } - do - let r ← @MetaTestable.run (∀ x : Subtype p, β x.val) (@varTestable var _ _ _ _) cfg min e - have := by simp [Subtype.forall, NamedBinder] - return iff this r + run cfg min e := sorry + -- letI (x : Subtype p) : MetaTestable (β x) := + -- { run := fun cfg min e => do + -- let r ← MetaTestable.runProp (β x.val) cfg min e + -- return addInfo s!"guard: {printProp (p x)} (by construction)" id r (PSum.inr id) } + -- do + -- let r ← @MetaTestable.run (∀ x : Subtype p, β x.val) (@varTestable var _ _ _ _) cfg min e + -- have := by simp [Subtype.forall, NamedBinder] + -- return iff this r instance (priority := low) decidableTestable {p : Prop} [PrintableProp p] [Decidable p] : MetaTestable p where @@ -512,7 +533,7 @@ open MetaTestResult namespace MetaTestable /-- Execute `cmd` and repeat every time the result is `gaveUp` (at most `n` times). -/ -def retry (cmd : Rand (MetaTestResult p)) : Nat → Rand (MetaTestResult p) +def retry (cmd : MRand (MetaTestResult p)) : Nat → MRand (MetaTestResult p) | 0 => return MetaTestResult.gaveUp 1 | n+1 => do let r ← cmd @@ -533,38 +554,37 @@ def giveUp (x : Nat) : MetaTestResult p → MetaTestResult p end MetaTestable /-- Try `n` times to find a counter-example for `p`. -/ -def MetaTestable.runSuiteAux (p : Prop) [MetaTestable p] (cfg : Configuration) : - MetaTestResult p → Nat → Rand (MetaTestResult p) +def MetaTestable.runSuiteAux (p : Prop) [MetaTestable p] (propExpr: Expr) (cfg : Configuration) : + MetaTestResult p → Nat → MRand (MetaTestResult p) | r, 0 => return r | r, n+1 => do let size := (cfg.numInst - n - 1) * cfg.maxSize / cfg.numInst if cfg.traceSuccesses then slimTrace s!"New sample" slimTrace s!"Retrying up to {cfg.numRetries} times until guards hold" - let x ← retry (ReaderT.run (MetaTestable.runProp p cfg true sorry) ⟨size⟩) cfg.numRetries + let x ← retry (ReaderT.run (MetaTestable.runProp p cfg true propExpr) ⟨size⟩) cfg.numRetries match x with - | success (PSum.inl ()) => runSuiteAux p cfg r n - | gaveUp g => runSuiteAux p cfg (giveUp g r) n + | success (PSum.inl ()) => runSuiteAux p propExpr cfg r n + | gaveUp g => runSuiteAux p propExpr cfg (giveUp g r) n | _ => return x /-- Try to find a counter-example of `p`. -/ -def MetaTestable.runSuite (p : Prop) [MetaTestable p] (cfg : Configuration := {}) : Rand (MetaTestResult p) := - MetaTestable.runSuiteAux p cfg (success <| PSum.inl ()) cfg.numInst +def MetaTestable.runSuite (p : Prop) [MetaTestable p] (propExpr: Expr) (cfg : Configuration := {}) : MRand (MetaTestResult p) := + MetaTestable.runSuiteAux p propExpr cfg (success <| PSum.inl ()) cfg.numInst -/-- Run a test suite for `p` in `BaseIO` using the global RNG in `stdGenRef`. -/ -def MetaTestable.checkIO (p : Prop) [MetaTestable p] (cfg : Configuration := {}) : BaseIO (MetaTestResult p) := - letI : MonadLift Id BaseIO := ⟨fun f => return Id.run f⟩ +/-- Run a test suite for `p` in `MetaM` using the global RNG in `stdGenRef`. -/ +def MetaTestable.checkIO (p : Prop) [MetaTestable p] (propExpr: Expr) (cfg : Configuration := {}) : MetaM (MetaTestResult p) := match cfg.randomSeed with - | none => runRand (MetaTestable.runSuite p cfg) - | some seed => runRandWith seed (MetaTestable.runSuite p cfg) + | none => runRand (MetaTestable.runSuite p propExpr cfg) + | some seed => runRandWith seed (MetaTestable.runSuite p propExpr cfg) end IO open Decorations in /-- Run a test suite for `p` and throw an exception if `p` does not hold. -/ -def MetaTestable.check (p : Prop) (cfg : Configuration := {}) - (p' : Decorations.DecorationsOf p := by mk_decorations) [MetaTestable p'] : Lean.CoreM PUnit := do - match ← MetaTestable.checkIO p' cfg with +def MetaTestable.check (p : Prop)(propExpr: Expr) (cfg : Configuration := {}) + (p' : Decorations.DecorationsOf p := by mk_decorations) [MetaTestable p'] : Lean.MetaM PUnit := do + match ← MetaTestable.checkIO p' propExpr cfg with | MetaTestResult.success _ => if !cfg.quiet then Lean.logInfo "Unable to find a counter-example" | MetaTestResult.gaveUp n => if !cfg.quiet then From 6fabb20a663125747ebddd40bb48b0b34d578765 Mon Sep 17 00:00:00 2001 From: Siddhartha Gadgil Date: Tue, 19 Nov 2024 18:21:55 +0530 Subject: [PATCH 04/30] more meta code written; untested --- Plausible/MetaTestable.lean | 196 +++++++++++++++++++++++------------- 1 file changed, 128 insertions(+), 68 deletions(-) diff --git a/Plausible/MetaTestable.lean b/Plausible/MetaTestable.lean index 5bd9996..1ff8ddd 100644 --- a/Plausible/MetaTestable.lean +++ b/Plausible/MetaTestable.lean @@ -117,13 +117,13 @@ def forallProp? (e: Expr) : MetaM (Option Expr × Option Expr) := do let u ← mkFreshLevelMVar let α ← mkFreshExprMVar (mkSort u) let prop := mkSort levelZero - let fmly ← mkArrow α prop - let mvar ← mkFreshExprMVar (some fmly) + let fmlyType ← mkArrow α prop + let fmly ← mkFreshExprMVar (some fmlyType) let e' ← withLocalDeclD `x α fun x => do - let y ← mkAppM' mvar #[x] + let y ← mkAppM' fmly #[x] mkForallFVars #[x] y if ← isDefEq e' e then - return (← Lean.getExprMVarAssignment? mvar.mvarId!, ← Lean.getExprMVarAssignment? α.mvarId!) + return (← Lean.getExprMVarAssignment? fmly.mvarId!, ← Lean.getExprMVarAssignment? α.mvarId!) else return (none, none) @@ -157,6 +157,23 @@ def eqlProp? (e: Expr): MetaM (Option (Expr × Expr × Expr)) := do else return none +def iffProp? (e: Expr): MetaM (Option (Expr × Expr)) := do + let level ← mkFreshLevelMVar + let u := mkSort level + let α ← mkFreshExprMVar u + let β ← mkFreshExprMVar α + let e' ← mkAppM ``Iff #[α , β] + if ← isDefEq e' e then + let α? ← Lean.getExprMVarAssignment? α.mvarId! + let β? ← Lean.getExprMVarAssignment? β.mvarId! + let pair : Option (Expr × Expr) := do + let α ← α? + let β ← β? + return (α, β) + return pair + else + return none + def equality? (e: Expr): MetaM (Option (Expr × Expr × Expr)) := do let level ← mkFreshLevelMVar let u := mkSort level @@ -190,8 +207,6 @@ abbrev MGen (α : Type) := ReaderT (ULift Nat) MRand α instance : MonadLift Gen MGen where monadLift := fun x n => x.run n.down -#synth MonadLift MetaM MRand - /-- Result of trying to disprove `p` -/ inductive MetaTestResult (p : Prop) where /-- @@ -295,32 +310,35 @@ def or : MetaTestResult p → MetaTestResult q → Expr → MetaM (MetaTestResu /-- If `q → p`, then `¬ p → ¬ q` which means that testing `p` can allow us to find counter-examples to `q`. -/ -def imp (h : q → p) (r : MetaTestResult p) - (p : Unit ⊕' (p → q) := PSum.inl ()) : MetaTestResult q := +def imp (h : q → p) (hExp: Expr) (r : MetaTestResult p) + (p : Unit ⊕' (p → q) := PSum.inl ()) : MetaM (MetaTestResult q) := match r with - | failure h2 pf xs n => failure (mt h h2) pf xs n - | success h2 => success <| combine p h2 - | gaveUp n => gaveUp n - | mismatch s => mismatch s + | failure h2 pf xs n => do + let pf' ← mkAppM ``mt #[hExp, pf] + return failure (mt h h2) pf' xs n + | success h2 => return success <| combine p h2 + | gaveUp n => return gaveUp n + | mismatch s => return mismatch s /-- Test `q` by testing `p` and proving the equivalence between the two. -/ -def iff (h : q ↔ p) (r : MetaTestResult p) : MetaTestResult q := - imp h.mp r (PSum.inr h.mpr) +def iff (h : q ↔ p) (hExp: Expr) (r : MetaTestResult p) : MetaM (MetaTestResult q) := do + let hExp' ← mkAppM ``Iff.mp #[hExp] + imp h.mp hExp' r (PSum.inr h.mpr) /-- When we assign a value to a universally quantified variable, we record that value using this function so that our counter-examples can be informative. -/ -def addInfo (x : String) (h : q → p) (r : MetaTestResult p) - (p : Unit ⊕' (p → q) := PSum.inl ()) : MetaTestResult q := +def addInfo (x : String) (h : q → p) (hExp: Expr) (r : MetaTestResult p) + (p : Unit ⊕' (p → q) := PSum.inl ()) : (MetaM <| MetaTestResult q) := if let failure h2 pf xs n := r then - failure (mt h h2) pf (x :: xs) n + return failure (mt h h2) pf (x :: xs) n else - imp h r p + imp h hExp r p /-- Add some formatting to the information recorded by `addInfo`. -/ -def addVarInfo {γ : Type _} [Repr γ] (var : String) (x : γ) (h : q → p) (r : MetaTestResult p) - (p : Unit ⊕' (p → q) := PSum.inl ()) : MetaTestResult q := - addInfo s!"{var} := {repr x}" h r p +def addVarInfo {γ : Type _} [Repr γ] (var : String) (x : γ) (h : q → p) (hExp: Expr) (r : MetaTestResult p) + (p : Unit ⊕' (p → q) := PSum.inl ()) : MetaM (MetaTestResult q) := + addInfo s!"{var} := {repr x}" h (hExp: Expr) r p def isFailure : MetaTestResult p → Bool | failure .. => true @@ -357,16 +375,21 @@ instance orTestable [MetaTestable p] [MetaTestable q] : MetaTestable (p ∨ q) w let xq ← runProp q cfg min e or xp xq e +theorem iff_resolve (p q : Prop) : (p ↔ q) ↔ p ∧ q ∨ ¬p ∧ ¬q := by + constructor + · intro h + simp [h, Classical.em] + · intro h + rcases h with ⟨hleft, hright⟩ | ⟨hleft, hright⟩ <;> simp [hleft, hright] + instance iffTestable [MetaTestable ((p ∧ q) ∨ (¬ p ∧ ¬ q))] : MetaTestable (p ↔ q) where run := fun cfg min e => do - let h ← runProp ((p ∧ q) ∨ (¬ p ∧ ¬ q)) cfg min e - have := by - constructor - · intro h - simp [h, Classical.em] - · intro h - rcases h with ⟨hleft, hright⟩ | ⟨hleft, hright⟩ <;> simp [hleft, hright] - return iff this h + match ← iffProp? e with + | some (α, β) => do + let h ← runProp ((p ∧ q) ∨ (¬ p ∧ ¬ q)) cfg min e + let hExp ← mkAppM ``iff_resolve #[α, β] + iff (iff_resolve p q) hExp h + | none => return mismatch "Expected an `Iff` proposition" variable {var : String} @@ -376,7 +399,7 @@ instance decGuardTestable [PrintableProp p] [Decidable p] {β : p → Prop} [∀ if h : p then let res := runProp (β h) cfg min e let s := printProp p - (fun r => addInfo s!"guard: {s}" (· <| h) r (PSum.inr <| fun q _ => q)) <$> res + ← (fun r => addInfo s!"guard: {s}" (· <| h) e r (PSum.inr <| fun q _ => q)) <$> res else if cfg.traceDiscarded || cfg.traceSuccesses then let res := fun _ => return gaveUp 1 let s := printProp p @@ -388,7 +411,7 @@ instance forallTypesTestable {f : Type → Prop} [MetaTestable (f Int)] : MetaTestable (NamedBinder var <| ∀ x, f x) where run := fun cfg min e => do let r ← runProp (f Int) cfg min e - return addVarInfo var "Int" (· <| Int) r + addVarInfo var "Int" (· <| Int) e r -- TODO: only in mathlib: @[pp_with_univ] instance (priority := 100) forallTypesULiftTestable.{u} @@ -396,7 +419,7 @@ instance (priority := 100) forallTypesULiftTestable.{u} MetaTestable (NamedBinder var <| ∀ x, f x) where run cfg min e := do let r ← runProp (f (ULift Int)) cfg min e - pure <| addVarInfo var "ULift Int" (· <| ULift Int) r + addVarInfo var "ULift Int" (· <| ULift Int) e r /-- Format the counter-examples found in a test failure. @@ -430,7 +453,7 @@ candidate that falsifies a property and recursively shrinking that one. The process is guaranteed to terminate because `shrink x` produces a proof that all the values it produces are smaller (according to `SizeOf`) than `x`. -/ -partial def minimizeAux [SampleableExt α] [ToExpr (SampleableExt.proxy α)] {β : α → Prop} [∀ x, MetaTestable (β x)] (cfg : Configuration) +partial def minimizeAux [SampleableExt α] [ToExpr (SampleableExt.proxy α)] {β : α → Prop} [∀ x, MetaTestable (β x)] (αExp βExpr: Expr) (cfg : Configuration) (var : String) (x : SampleableExt.proxy α) (n : Nat) : OptionT MGen (Σ x, MetaTestResult (β (SampleableExt.interp x))) := do let candidates := SampleableExt.shrink.shrink x @@ -439,12 +462,16 @@ partial def minimizeAux [SampleableExt α] [ToExpr (SampleableExt.proxy α)] {β for candidate in candidates do if cfg.traceShrinkCandidates then slimTrace s!"Trying {var} := {repr candidate}" - let res ← OptionT.lift <| MetaTestable.runProp (β (SampleableExt.interp candidate)) cfg true sorry + let xExpr := toExpr candidate + let samp ← synthInstance <| ← mkAppM ``SampleableExt #[αExp] + let xInterp ← mkAppOptM ``SampleableExt.interp #[αExp, samp, xExpr] + let e' ← mkAppM' βExpr #[xInterp] + let res ← OptionT.lift <| MetaTestable.runProp (β (SampleableExt.interp candidate)) cfg true e' if res.isFailure then if cfg.traceShrink then slimTrace s!"{var} shrunk to {repr candidate} from {repr x}" let currentStep := OptionT.lift <| return Sigma.mk candidate (addShrinks (n + 1) res) - let nextStep := minimizeAux cfg var candidate (n + 1) + let nextStep := minimizeAux αExp βExpr cfg var candidate (n + 1) return ← (nextStep <|> currentStep) if cfg.traceShrink then slimTrace s!"No shrinking possible for {var} := {repr x}" @@ -452,13 +479,13 @@ partial def minimizeAux [SampleableExt α] [ToExpr (SampleableExt.proxy α)] {β /-- Once a property fails to hold on an example, look for smaller counter-examples to show the user. -/ -def minimize [SampleableExt α] [ToExpr (SampleableExt.proxy α)] {β : α → Prop} [∀ x, MetaTestable (β x)] (cfg : Configuration) +def minimize [SampleableExt α] [ToExpr (SampleableExt.proxy α)] {β : α → Prop} [∀ x, MetaTestable (β x)] (αExp βExpr: Expr)(cfg : Configuration) (var : String) (x : SampleableExt.proxy α) (r : MetaTestResult (β <| SampleableExt.interp x)) : MGen (Σ x, MetaTestResult (β <| SampleableExt.interp x)) := do if cfg.traceShrink then slimTrace "Shrink" slimTrace s!"Attempting to shrink {var} := {repr x}" - let res ← OptionT.run <| minimizeAux cfg var x 0 + let res ← OptionT.run <| minimizeAux αExp βExpr cfg var x 0 return res.getD ⟨x, r⟩ /-- Test a universal property by creating a sample of the right type and instantiating the @@ -466,28 +493,60 @@ bound variable with it. -/ instance varTestable [SampleableExt α] [ToExpr (SampleableExt.proxy α)] {β : α → Prop} [∀ x, MetaTestable (β x)] : MetaTestable (NamedBinder var <| ∀ x : α, β x) where run := fun cfg min e => do - let x ← SampleableExt.sample - if cfg.traceSuccesses || cfg.traceDiscarded then - slimTrace s!"{var} := {repr x}" - let r ← MetaTestable.runProp (β <| SampleableExt.interp x) cfg false e - let ⟨finalX, finalR⟩ ← - if isFailure r then - if cfg.traceSuccesses then - slimTrace s!"{var} := {repr x} is a failure" - if min then - minimize cfg var x r + match ← forallProp? e with + | (some βExp, some αExp) => do + let x ← SampleableExt.sample + let xExpr := toExpr x + let samp ← synthInstance <| ← mkAppM ``SampleableExt #[αExp] + let xInterp ← mkAppOptM ``SampleableExt.interp #[αExp, samp, xExpr] + let e' ← mkAppM' βExp #[xInterp] + if cfg.traceSuccesses || cfg.traceDiscarded then + slimTrace s!"{var} := {repr x}" + let r ← MetaTestable.runProp (β <| SampleableExt.interp x) cfg false e' + let ⟨finalX, finalR⟩ ← + if isFailure r then + if cfg.traceSuccesses then + slimTrace s!"{var} := {repr x} is a failure" + if min then + minimize αExp βExp cfg var x r + else + pure ⟨x, r⟩ else pure ⟨x, r⟩ - else - pure ⟨x, r⟩ - return addVarInfo var finalX (· <| SampleableExt.interp finalX) finalR + let xExpr := toExpr finalX + let samp ← synthInstance <| ← mkAppM ``SampleableExt #[αExp] + let xInterp ← mkAppOptM ``SampleableExt.interp #[αExp, samp, xExpr] + let e' ← mkAppM' βExp #[xInterp] + addVarInfo var finalX (· <| SampleableExt.interp finalX) e' finalR + | (_, _) => throwError "Expected a `Forall` proposition" + +abbrev ProxyExpr α [SampleableExt α] := ToExpr (SampleableExt.proxy α) + +instance : ProxyExpr Bool := (inferInstance : ToExpr Bool) +instance : ProxyExpr String := (inferInstance : ToExpr String) +instance : ProxyExpr Nat := (inferInstance : ToExpr Nat) +instance : ProxyExpr Int := (inferInstance : ToExpr Int) +instance : ProxyExpr Char := (inferInstance : ToExpr Char) +instance : ProxyExpr Unit := (inferInstance : ToExpr Unit) +instance : ProxyExpr USize := (inferInstance : ToExpr USize) +instance : ProxyExpr UInt32 := (inferInstance : ToExpr UInt32) +instance : ProxyExpr UInt64 := (inferInstance : ToExpr UInt64) +instance : ProxyExpr UInt16 := (inferInstance : ToExpr UInt16) +instance {n : Nat} : ProxyExpr (Fin n.succ) := (inferInstance : ToExpr (Fin n.succ)) + +theorem bool_to_prop_fmly (β : Prop → Prop): NamedBinder var (∀ (p : Prop), β p) → ∀ (b : Bool), β (b = true) := fun h b => h (b = true) /-- Test a universal property about propositions -/ -instance propVarTestable {β : Prop → Prop} [∀ b : Bool, MetaTestable (β b)] : +instance propVarTestable {β : Prop → Prop} [h: ∀ b : Bool, MetaTestable (β b)] : MetaTestable (NamedBinder var <| ∀ p : Prop, β p) where - run := fun cfg min e => sorry -- Need to interpolate `Prop` and `Bool` - -- imp (fun h (b : Bool) => h b) <$> MetaTestable.runProp (NamedBinder var <| ∀ b : Bool, β b) cfg min e + run := fun cfg min e => do + match ← forallProp? e with + | (some βExpr, _) => + let p ← MetaTestable.runProp (NamedBinder var <| ∀ b : Bool, β b) cfg min e + let e' ← mkAppM ``bool_to_prop_fmly #[βExpr] + imp (bool_to_prop_fmly β) e' p + | _ => throwError "Expected a `Forall` proposition" instance (priority := high) unusedVarTestable {β : Prop} [Nonempty α] [MetaTestable β] : MetaTestable (NamedBinder var (α → β)) @@ -496,37 +555,38 @@ where if cfg.traceDiscarded || cfg.traceSuccesses then slimTrace s!"{var} is unused" let r ← MetaTestable.runProp β cfg min e - let finalR := addInfo s!"{var} is irrelevant (unused)" id r - return imp (· <| Classical.ofNonempty) finalR (PSum.inr <| fun x _ => x) + let finalR ← addInfo s!"{var} is irrelevant (unused)" id e r + imp (· <| Classical.ofNonempty) e finalR (PSum.inr <| fun x _ => x) instance (priority := 2000) subtypeVarTestable {p : α → Prop} {β : α → Prop} [∀ x, PrintableProp (p x)] [∀ x, MetaTestable (β x)] [SampleableExt (Subtype p)] [ToExpr (SampleableExt.proxy (Subtype p))] {var'} : MetaTestable (NamedBinder var <| (x : α) → NamedBinder var' <| p x → β x) where - run cfg min e := sorry - -- letI (x : Subtype p) : MetaTestable (β x) := - -- { run := fun cfg min e => do - -- let r ← MetaTestable.runProp (β x.val) cfg min e - -- return addInfo s!"guard: {printProp (p x)} (by construction)" id r (PSum.inr id) } - -- do - -- let r ← @MetaTestable.run (∀ x : Subtype p, β x.val) (@varTestable var _ _ _ _) cfg min e - -- have := by simp [Subtype.forall, NamedBinder] - -- return iff this r + run cfg min e := + letI (x : Subtype p) : MetaTestable (β x) := + { run := fun cfg min e => do + let r ← MetaTestable.runProp (β x.val) cfg min e + addInfo s!"guard: {printProp (p x)} (by construction)" id e r (PSum.inr id) } + do + let r ← @MetaTestable.run (∀ x : Subtype p, β x.val) (@varTestable var _ _ _ _ _) cfg min e + have := by simp [Subtype.forall, NamedBinder] + iff this e r instance (priority := low) decidableTestable {p : Prop} [PrintableProp p] [Decidable p] : MetaTestable p where - run := fun _ _ _ => + run := fun _ _ e _ => do if h : p then return success (PSum.inr h) else let s := printProp p - return failure h sorry [s!"issue: {s} does not hold"] 0 + let inst ← synthInstance <| ← mkAppM ``Decidable #[e] + let falseRefl ← mkAppM ``Eq.refl #[mkConst ``false] + let e' ← mkAppM ``of_decide_eq_false #[e, inst, falseRefl] + return failure h e' [s!"issue: {s} does not hold"] 0 end MetaTestable -#check PrintableProp - section IO open MetaTestResult From ed3ed3bd44f4cbdd6be35432f3944f5a49d0e1d3 Mon Sep 17 00:00:00 2001 From: Siddhartha Gadgil Date: Wed, 20 Nov 2024 08:58:09 +0530 Subject: [PATCH 05/30] removed `.mistmatch`, throwing errors instead --- Plausible/MetaTestable.lean | 30 ++++++------------------------ 1 file changed, 6 insertions(+), 24 deletions(-) diff --git a/Plausible/MetaTestable.lean b/Plausible/MetaTestable.lean index 1ff8ddd..a146724 100644 --- a/Plausible/MetaTestable.lean +++ b/Plausible/MetaTestable.lean @@ -111,8 +111,6 @@ def andProp? (e: Expr) : MetaM (Option Expr × Option Expr) := do else return (none, none) -#check mkFreshLevelMVar - def forallProp? (e: Expr) : MetaM (Option Expr × Option Expr) := do let u ← mkFreshLevelMVar let α ← mkFreshExprMVar (mkSort u) @@ -223,10 +221,6 @@ inductive MetaTestResult (p : Prop) where -/ | gaveUp : Nat → MetaTestResult p /-- - There was a mismatch between `p` and the expression representing `p` - -/ - | mismatch : String → MetaTestResult p - /-- A counter-example to `p`; the strings specify values for the relevant variables. `failure h vs n` also carries a proof that `p` does not hold. This way, we can guarantee that there will be no false positive. The last component, `n`, @@ -248,7 +242,6 @@ def toString : MetaTestResult p → String | success (PSum.inr _) => "success (proof)" | gaveUp n => s!"gave {n} times" | failure _ _ counters _ => s!"failed {counters}" - | mismatch s => s!"Mismatch: {s}" instance : ToString (MetaTestResult p) := ⟨toString⟩ @@ -267,7 +260,7 @@ def and : MetaTestResult p → MetaTestResult q → Expr → MetaM (MetaTestRes let e' ← mkAppM' pf #[x] mkLambdaFVars #[h] e' return failure (fun h2 => h h2.left) pf' xs n - | (_, _) => throwError "Expected an `And` proposition" + | (_, _) => throwError m!"Expected an `And` proposition, got {← ppExpr e}" | _, failure h pf xs n, e => do match ← andProp? e with | (some e₁, some e₂) => @@ -276,14 +269,12 @@ def and : MetaTestResult p → MetaTestResult q → Expr → MetaM (MetaTestRes let e' ← mkAppM' pf #[x] mkLambdaFVars #[h] e' return failure (fun h2 => h h2.right) pf' xs n - | (_, _) => throwError "Expected an `And` proposition" + | (_, _) => throwError m!"Expected an `And` proposition, got {← ppExpr e}" | success h1, success h2, _ => return success <| combine (combine (PSum.inr And.intro) h1) h2 | gaveUp n, gaveUp m, _ => return gaveUp <| n + m | gaveUp n, _, _ => return gaveUp n | _, gaveUp n, _ => return gaveUp n - | mismatch s, _, _ => return mismatch s - | _, mismatch s, _ => return mismatch s /-- Combine the test result for properties `p` and `q` to create a test for their disjunction. -/ @@ -299,14 +290,12 @@ def or : MetaTestResult p → MetaTestResult q → Expr → MetaM (MetaTestResu | Or.inl h3 => h1 h3 | Or.inr h3 => h2 h3 return failure h3 pf (xs ++ ys) (n + m) - | (_, _) => throwError "Expected an `Or` proposition" + | (_, _) => throwError m!"Expected an `Or` proposition, got {← ppExpr e}" | success h, _, _ => return success <| combine (PSum.inr Or.inl) h | _, success h, _ => return success <| combine (PSum.inr Or.inr) h | gaveUp n, gaveUp m, _ => return gaveUp <| n + m | gaveUp n, _, _ => return gaveUp n | _, gaveUp n, _ => return gaveUp n - | mismatch s, _, _ => return mismatch s - | _, mismatch s, _ => return mismatch s /-- If `q → p`, then `¬ p → ¬ q` which means that testing `p` can allow us to find counter-examples to `q`. -/ @@ -318,7 +307,6 @@ def imp (h : q → p) (hExp: Expr) (r : MetaTestResult p) return failure (mt h h2) pf' xs n | success h2 => return success <| combine p h2 | gaveUp n => return gaveUp n - | mismatch s => return mismatch s /-- Test `q` by testing `p` and proving the equivalence between the two. -/ def iff (h : q ↔ p) (hExp: Expr) (r : MetaTestResult p) : MetaM (MetaTestResult q) := do @@ -389,7 +377,7 @@ instance iffTestable [MetaTestable ((p ∧ q) ∨ (¬ p ∧ ¬ q))] : MetaTestab let h ← runProp ((p ∧ q) ∨ (¬ p ∧ ¬ q)) cfg min e let hExp ← mkAppM ``iff_resolve #[α, β] iff (iff_resolve p q) hExp h - | none => return mismatch "Expected an `Iff` proposition" + | none => throwError m!"Expected an `Iff` proposition, got {← ppExpr e}" variable {var : String} @@ -518,7 +506,7 @@ instance varTestable [SampleableExt α] [ToExpr (SampleableExt.proxy α)] {β : let xInterp ← mkAppOptM ``SampleableExt.interp #[αExp, samp, xExpr] let e' ← mkAppM' βExp #[xInterp] addVarInfo var finalX (· <| SampleableExt.interp finalX) e' finalR - | (_, _) => throwError "Expected a `Forall` proposition" + | (_, _) => throwError m!"Expected a `Forall` proposition, got {← ppExpr e}" abbrev ProxyExpr α [SampleableExt α] := ToExpr (SampleableExt.proxy α) @@ -546,7 +534,7 @@ where let p ← MetaTestable.runProp (NamedBinder var <| ∀ b : Bool, β b) cfg min e let e' ← mkAppM ``bool_to_prop_fmly #[βExpr] imp (bool_to_prop_fmly β) e' p - | _ => throwError "Expected a `Forall` proposition" + | _ => throwError m!"Expected a `Forall` proposition, got {← ppExpr e}" instance (priority := high) unusedVarTestable {β : Prop} [Nonempty α] [MetaTestable β] : MetaTestable (NamedBinder var (α → β)) @@ -601,7 +589,6 @@ def retry (cmd : MRand (MetaTestResult p)) : Nat → MRand (MetaTestResult p) | .success hp => return success hp | .failure h pf xs n => return failure h pf xs n | .gaveUp _ => retry cmd n - | .mismatch s => return .mismatch s /-- Count the number of times the test procedure gave up. -/ def giveUp (x : Nat) : MetaTestResult p → MetaTestResult p @@ -609,7 +596,6 @@ def giveUp (x : Nat) : MetaTestResult p → MetaTestResult p | success (PSum.inr p) => success <| (PSum.inr p) | gaveUp n => gaveUp <| n + x | MetaTestResult.failure h pf xs n => failure h pf xs n - | mismatch _ => gaveUp x end MetaTestable @@ -650,10 +636,6 @@ def MetaTestable.check (p : Prop)(propExpr: Expr) (cfg : Configuration := {}) if !cfg.quiet then let msg := s!"Gave up after failing to generate values that fulfill the preconditions {n} times." Lean.logWarning msg - | MetaTestResult.mismatch s => - if !cfg.quiet then - let msg := s!"Mismatch: {s}" - Lean.logWarning msg | MetaTestResult.failure _ _ xs n => let msg := "Found a counter-example!" if cfg.quiet then From 3939224cdd0ad6ef35ef5abb4216124818e8d0e3 Mon Sep 17 00:00:00 2001 From: Siddhartha Gadgil Date: Wed, 20 Nov 2024 12:07:39 +0530 Subject: [PATCH 06/30] tests for matching --- Plausible/MetaTestable.lean | 42 +++++++++++++++++++++++++++++++++---- 1 file changed, 38 insertions(+), 4 deletions(-) diff --git a/Plausible/MetaTestable.lean b/Plausible/MetaTestable.lean index a146724..3820590 100644 --- a/Plausible/MetaTestable.lean +++ b/Plausible/MetaTestable.lean @@ -156,10 +156,9 @@ def eqlProp? (e: Expr): MetaM (Option (Expr × Expr × Expr)) := do return none def iffProp? (e: Expr): MetaM (Option (Expr × Expr)) := do - let level ← mkFreshLevelMVar - let u := mkSort level - let α ← mkFreshExprMVar u - let β ← mkFreshExprMVar α + let prop := mkSort levelZero + let α ← mkFreshExprMVar prop + let β ← mkFreshExprMVar prop let e' ← mkAppM ``Iff #[α , β] if ← isDefEq e' e then let α? ← Lean.getExprMVarAssignment? α.mvarId! @@ -192,6 +191,41 @@ def equality? (e: Expr): MetaM (Option (Expr × Expr × Expr)) := do return triple else return none + +-- for testing +open Lean Elab Term in +elab "#decompose_prop" t:term : command => + Command.liftTermElabM do + let e ← elabType t + match ← orProp? e with + | (some α, some β) => logInfo s!"Or: {← ppExpr α}; {← ppExpr β}" + | _ => pure () + match ← andProp? e with + | (some α, some β) => logInfo s!"And: {← ppExpr α}; {← ppExpr β}" + | _ => pure () + match ← existsProp? e with + | (some α, some β) => logInfo s!"Exists: {← ppExpr α}; domain {← ppExpr β}" + | _ => pure () + match ← forallProp? e with + | (some α, some β) => logInfo s!"Forall: {← ppExpr α}; domain {← ppExpr β}" + | _ => pure () + match ← impProp? e with + | (some α, some β) => logInfo s!"Imp: {← ppExpr α}; {← ppExpr β}" + | _ => pure () + match ← eqlProp? e with + | some (α, a, b) => logInfo s!"Eq: {← ppExpr α}; {← ppExpr a} and {← ppExpr b}" + | _ => pure () + match ← iffProp? e with + | some (α, β) => logInfo s!"Iff: {← ppExpr α}; {← ppExpr β}" + | _ => pure () + +#decompose_prop ∀ (n: Nat), n = 0 ∨ n ≠ 0 +#decompose_prop NamedBinder "blah" <| ∀ (n: Nat), n = 0 ∨ n ≠ 0 +#decompose_prop 1 = 0 ∨ 2 ≠ 0 +#decompose_prop 1 = 0 ∧ 2 ≠ 0 +#decompose_prop ∃ (n: Nat), n = 0 ∨ n ≠ 0 + + end Matching open Lean Meta From 814db0ee9d1e05dd70d702a84a7e50dd0b07a920 Mon Sep 17 00:00:00 2001 From: Siddhartha Gadgil Date: Wed, 20 Nov 2024 12:21:43 +0530 Subject: [PATCH 07/30] more instances --- Plausible/MetaTestable.lean | 38 ++++++++++++++++++------------------- 1 file changed, 18 insertions(+), 20 deletions(-) diff --git a/Plausible/MetaTestable.lean b/Plausible/MetaTestable.lean index a146724..27e5243 100644 --- a/Plausible/MetaTestable.lean +++ b/Plausible/MetaTestable.lean @@ -339,6 +339,9 @@ namespace MetaTestable open MetaTestResult +abbrev ProxyExpr α [SampleableExt α] := ToExpr (SampleableExt.proxy α) + + def runProp (p : Prop) [MetaTestable p] : Configuration → Bool → Expr → MGen (MetaTestResult p) := MetaTestable.run /-- A `dbgTrace` with special formatting -/ @@ -409,20 +412,6 @@ instance (priority := 100) forallTypesULiftTestable.{u} let r ← runProp (f (ULift Int)) cfg min e addVarInfo var "ULift Int" (· <| ULift Int) e r -/-- -Format the counter-examples found in a test failure. --/ -def formatFailure (s : String) (xs : List String) (n : Nat) : String := - let counter := String.intercalate "\n" xs - let parts := [ - "\n===================", - s, - counter, - s!"({n} shrinks)", - "-------------------" - ] - String.intercalate "\n" parts - /-- Increase the number of shrinking steps in a test result. -/ @@ -441,7 +430,7 @@ candidate that falsifies a property and recursively shrinking that one. The process is guaranteed to terminate because `shrink x` produces a proof that all the values it produces are smaller (according to `SizeOf`) than `x`. -/ -partial def minimizeAux [SampleableExt α] [ToExpr (SampleableExt.proxy α)] {β : α → Prop} [∀ x, MetaTestable (β x)] (αExp βExpr: Expr) (cfg : Configuration) +partial def minimizeAux [SampleableExt α] [ProxyExpr α] {β : α → Prop} [∀ x, MetaTestable (β x)] (αExp βExpr: Expr) (cfg : Configuration) (var : String) (x : SampleableExt.proxy α) (n : Nat) : OptionT MGen (Σ x, MetaTestResult (β (SampleableExt.interp x))) := do let candidates := SampleableExt.shrink.shrink x @@ -467,7 +456,7 @@ partial def minimizeAux [SampleableExt α] [ToExpr (SampleableExt.proxy α)] {β /-- Once a property fails to hold on an example, look for smaller counter-examples to show the user. -/ -def minimize [SampleableExt α] [ToExpr (SampleableExt.proxy α)] {β : α → Prop} [∀ x, MetaTestable (β x)] (αExp βExpr: Expr)(cfg : Configuration) +def minimize [SampleableExt α] [ProxyExpr α] {β : α → Prop} [∀ x, MetaTestable (β x)] (αExp βExpr: Expr)(cfg : Configuration) (var : String) (x : SampleableExt.proxy α) (r : MetaTestResult (β <| SampleableExt.interp x)) : MGen (Σ x, MetaTestResult (β <| SampleableExt.interp x)) := do if cfg.traceShrink then @@ -478,7 +467,7 @@ def minimize [SampleableExt α] [ToExpr (SampleableExt.proxy α)] {β : α → P /-- Test a universal property by creating a sample of the right type and instantiating the bound variable with it. -/ -instance varTestable [SampleableExt α] [ToExpr (SampleableExt.proxy α)] {β : α → Prop} [∀ x, MetaTestable (β x)] : +instance varTestable [SampleableExt α] [ProxyExpr α] {β : α → Prop} [∀ x, MetaTestable (β x)] : MetaTestable (NamedBinder var <| ∀ x : α, β x) where run := fun cfg min e => do match ← forallProp? e with @@ -508,7 +497,6 @@ instance varTestable [SampleableExt α] [ToExpr (SampleableExt.proxy α)] {β : addVarInfo var finalX (· <| SampleableExt.interp finalX) e' finalR | (_, _) => throwError m!"Expected a `Forall` proposition, got {← ppExpr e}" -abbrev ProxyExpr α [SampleableExt α] := ToExpr (SampleableExt.proxy α) instance : ProxyExpr Bool := (inferInstance : ToExpr Bool) instance : ProxyExpr String := (inferInstance : ToExpr String) @@ -522,6 +510,16 @@ instance : ProxyExpr UInt64 := (inferInstance : ToExpr UInt64) instance : ProxyExpr UInt16 := (inferInstance : ToExpr UInt16) instance {n : Nat} : ProxyExpr (Fin n.succ) := (inferInstance : ToExpr (Fin n.succ)) +instance : ProxyExpr Prop := (inferInstance : ToExpr Bool) + +instance [SampleableExt α][ProxyExpr α] : ProxyExpr (List α) := + (inferInstance : ToExpr (List (SampleableExt.proxy α))) +instance [SampleableExt α][ProxyExpr α] : ProxyExpr (Array α) := + (inferInstance : ToExpr (Array (SampleableExt.proxy α))) +instance [SampleableExt α][SampleableExt β] [ProxyExpr α][ProxyExpr β] : ProxyExpr (Prod α β) := + (inferInstance : ToExpr (Prod (SampleableExt.proxy α) (SampleableExt.proxy β) )) + + theorem bool_to_prop_fmly (β : Prop → Prop): NamedBinder var (∀ (p : Prop), β p) → ∀ (b : Bool), β (b = true) := fun h b => h (b = true) /-- Test a universal property about propositions -/ @@ -549,7 +547,7 @@ where instance (priority := 2000) subtypeVarTestable {p : α → Prop} {β : α → Prop} [∀ x, PrintableProp (p x)] [∀ x, MetaTestable (β x)] - [SampleableExt (Subtype p)] [ToExpr (SampleableExt.proxy (Subtype p))] {var'} : + [SampleableExt (Subtype p)] [ProxyExpr (Subtype p)] {var'} : MetaTestable (NamedBinder var <| (x : α) → NamedBinder var' <| p x → β x) where run cfg min e := letI (x : Subtype p) : MetaTestable (β x) := @@ -641,7 +639,7 @@ def MetaTestable.check (p : Prop)(propExpr: Expr) (cfg : Configuration := {}) if cfg.quiet then Lean.throwError msg else - Lean.throwError <| formatFailure msg xs n + Lean.throwError <| Testable.formatFailure msg xs n -- #eval MetaTestable.check (∀ (x y z a : Nat) (h1 : 3 < x) (h2 : 3 < y), x - y = y - x) -- Configuration.verbose From 699bd8ca6376b2d769e6d4f2a495a81d82967ff1 Mon Sep 17 00:00:00 2001 From: Siddhartha Gadgil Date: Wed, 20 Nov 2024 19:01:09 +0530 Subject: [PATCH 08/30] ported test, fiddly expressions --- Plausible/MetaTestable.lean | 52 ++++++-------------- Test/MetaTestable.lean | 95 +++++++++++++++++++++++++++++++++++++ 2 files changed, 110 insertions(+), 37 deletions(-) create mode 100644 Test/MetaTestable.lean diff --git a/Plausible/MetaTestable.lean b/Plausible/MetaTestable.lean index 31d407b..a6a8180 100644 --- a/Plausible/MetaTestable.lean +++ b/Plausible/MetaTestable.lean @@ -192,39 +192,6 @@ def equality? (e: Expr): MetaM (Option (Expr × Expr × Expr)) := do else return none --- for testing -open Lean Elab Term in -elab "#decompose_prop" t:term : command => - Command.liftTermElabM do - let e ← elabType t - match ← orProp? e with - | (some α, some β) => logInfo s!"Or: {← ppExpr α}; {← ppExpr β}" - | _ => pure () - match ← andProp? e with - | (some α, some β) => logInfo s!"And: {← ppExpr α}; {← ppExpr β}" - | _ => pure () - match ← existsProp? e with - | (some α, some β) => logInfo s!"Exists: {← ppExpr α}; domain {← ppExpr β}" - | _ => pure () - match ← forallProp? e with - | (some α, some β) => logInfo s!"Forall: {← ppExpr α}; domain {← ppExpr β}" - | _ => pure () - match ← impProp? e with - | (some α, some β) => logInfo s!"Imp: {← ppExpr α}; {← ppExpr β}" - | _ => pure () - match ← eqlProp? e with - | some (α, a, b) => logInfo s!"Eq: {← ppExpr α}; {← ppExpr a} and {← ppExpr b}" - | _ => pure () - match ← iffProp? e with - | some (α, β) => logInfo s!"Iff: {← ppExpr α}; {← ppExpr β}" - | _ => pure () - -#decompose_prop ∀ (n: Nat), n = 0 ∨ n ≠ 0 -#decompose_prop NamedBinder "blah" <| ∀ (n: Nat), n = 0 ∨ n ≠ 0 -#decompose_prop 1 = 0 ∨ 2 ≠ 0 -#decompose_prop 1 = 0 ∧ 2 ≠ 0 -#decompose_prop ∃ (n: Nat), n = 0 ∨ n ≠ 0 - end Matching @@ -284,6 +251,12 @@ def combine {p q : Prop} : Unit ⊕' (p → q) → Unit ⊕' p → Unit ⊕' q | PSum.inr f, PSum.inr proof => PSum.inr <| f proof | _, _ => PSum.inl () +def checkDisproof (pf prop: Expr) : MetaM Unit := do + let negProp ← mkAppM ``Not #[prop] + let pfType ← inferType pf + unless ← isDefEq pfType negProp do + throwError m!"Expected a proof of {negProp}, got proof of {← ppExpr pfType}" + /-- Combine the test result for properties `p` and `q` to create a test for their conjunction. -/ def and : MetaTestResult p → MetaTestResult q → Expr → MetaM (MetaTestResult (p ∧ q)) | failure h pf xs n, _, e => do @@ -293,6 +266,7 @@ def and : MetaTestResult p → MetaTestResult q → Expr → MetaM (MetaTestRes let x ← mkAppOptM ``And.left #[e₁, e₂, h] let e' ← mkAppM' pf #[x] mkLambdaFVars #[h] e' + checkDisproof pf' e return failure (fun h2 => h h2.left) pf' xs n | (_, _) => throwError m!"Expected an `And` proposition, got {← ppExpr e}" | _, failure h pf xs n, e => do @@ -302,6 +276,7 @@ def and : MetaTestResult p → MetaTestResult q → Expr → MetaM (MetaTestRes let x ← mkAppOptM ``And.right #[e₁, e₂, h] let e' ← mkAppM' pf #[x] mkLambdaFVars #[h] e' + checkDisproof pf' e return failure (fun h2 => h h2.right) pf' xs n | (_, _) => throwError m!"Expected an `And` proposition, got {← ppExpr e}" @@ -323,6 +298,7 @@ def or : MetaTestResult p → MetaTestResult q → Expr → MetaM (MetaTestResu match h with | Or.inl h3 => h1 h3 | Or.inr h3 => h2 h3 + checkDisproof pf e return failure h3 pf (xs ++ ys) (n + m) | (_, _) => throwError m!"Expected an `Or` proposition, got {← ppExpr e}" | success h, _, _ => return success <| combine (PSum.inr Or.inl) h @@ -351,9 +327,10 @@ def iff (h : q ↔ p) (hExp: Expr) (r : MetaTestResult p) : MetaM (MetaTestResul we record that value using this function so that our counter-examples can be informative. -/ def addInfo (x : String) (h : q → p) (hExp: Expr) (r : MetaTestResult p) - (p : Unit ⊕' (p → q) := PSum.inl ()) : (MetaM <| MetaTestResult q) := + (p : Unit ⊕' (p → q) := PSum.inl ()) : (MetaM <| MetaTestResult q) := do if let failure h2 pf xs n := r then - return failure (mt h h2) pf (x :: xs) n + let pf' ← mkAppM ``mt #[hExp, pf] + return failure (mt h h2) pf' (x :: xs) n else imp h hExp r p @@ -602,8 +579,9 @@ instance (priority := low) decidableTestable {p : Prop} [PrintableProp p] [Decid let s := printProp p let inst ← synthInstance <| ← mkAppM ``Decidable #[e] let falseRefl ← mkAppM ``Eq.refl #[mkConst ``false] - let e' ← mkAppM ``of_decide_eq_false #[e, inst, falseRefl] - return failure h e' [s!"issue: {s} does not hold"] 0 + let pf' ← mkAppOptM ``of_decide_eq_false #[e, inst, falseRefl] + checkDisproof pf' e + return failure h pf' [s!"issue: {s} does not hold"] 0 end MetaTestable diff --git a/Test/MetaTestable.lean b/Test/MetaTestable.lean new file mode 100644 index 0000000..be02cb4 --- /dev/null +++ b/Test/MetaTestable.lean @@ -0,0 +1,95 @@ +import Lean +import Plausible.MetaTestable + +open Plausible Plausible.MetaTestable Lean Meta Elab Term + +open Lean Elab Term in +elab "#decompose_prop" t:term : command => + Command.liftTermElabM do + let e ← elabType t + match ← orProp? e with + | (some α, some β) => logInfo s!"Or: {← ppExpr α}; {← ppExpr β}" + | _ => pure () + match ← andProp? e with + | (some α, some β) => logInfo s!"And: {← ppExpr α}; {← ppExpr β}" + | _ => pure () + match ← existsProp? e with + | (some α, some β) => logInfo s!"Exists: {← ppExpr α}; domain {← ppExpr β}" + | _ => pure () + match ← forallProp? e with + | (some α, some β) => logInfo s!"Forall: {← ppExpr α}; domain {← ppExpr β}" + | _ => pure () + match ← impProp? e with + | (some α, some β) => logInfo s!"Imp: {← ppExpr α}; {← ppExpr β}" + | _ => pure () + match ← eqlProp? e with + | some (α, a, b) => logInfo s!"Eq: {← ppExpr α}; {← ppExpr a} and {← ppExpr b}" + | _ => pure () + match ← iffProp? e with + | some (α, β) => logInfo s!"Iff: {← ppExpr α}; {← ppExpr β}" + | _ => pure () + +/-- info: Forall: fun x => x = 0 ∨ x ≠ 0; domain Nat -/ +#guard_msgs in +#decompose_prop ∀ (n: Nat), n = 0 ∨ n ≠ 0 + +/-- info: Forall: fun x => x = 0 ∨ x ≠ 0; domain Nat -/ +#guard_msgs in +#decompose_prop NamedBinder "blah" <| ∀ (n: Nat), n = 0 ∨ n ≠ 0 + +/-- info: Or: 1 = 0; 2 ≠ 0 -/ +#guard_msgs in +#decompose_prop 1 = 0 ∨ 2 ≠ 0 + +/-- info: And: 1 = 0; 2 ≠ 0 -/ +#guard_msgs in +#decompose_prop 1 = 0 ∧ 2 ≠ 0 + +/-- info: Exists: fun n => n = 0 ∨ n ≠ 0; domain Nat -/ +#guard_msgs in +#decompose_prop ∃ (n: Nat), n = 0 ∨ n ≠ 0 + +#check List.nil + +#check MetaTestResult.failure + +elab "mk_failure%" prop:term ";" pf:term ";" : term => do + let prop ← elabType prop + let notProp ← mkAppM ``Not #[prop] + let pf ← elabTerm pf notProp + let lst ← mkAppOptM ``List.nil #[mkConst ``String] + let pfExpr := Lean.Expr.letE + `pf_var + (notProp) + (pf) + (Lean.Expr.app + (Lean.Expr.const `Lean.Expr.fvar []) + (Lean.Expr.app + (Lean.Expr.const `Lean.FVarId.mk []) + (Lean.Expr.app (Lean.Expr.const `Lean.Name.mkStr1 []) (Lean.Expr.lit (Lean.Literal.strVal "pf_var"))))) + false + logInfo s!"{repr pfExpr}" + mkAppOptM ``MetaTestResult.failure #[prop, pf, pfExpr, lst, mkConst ``Nat.zero] + +set_option pp.proofs true in +#check mk_failure% False ; fun (x: False) ↦ x ; + +elab "#expr" e:term : command => + Command.liftTermElabM do + let e ← elabTerm e none + logInfo s!"{repr e}" + logInfo s!"{← reduce e}" + +#check Expr.fvar {name := `n} + +/- +Lean.Expr.fvar (Lean.FVarId.mk (Lean.Name.mkStr1 "n")) +-/ +#expr ( + let pf_var : Nat := 0 + Expr.fvar {name := `pf_var}) +#expr Lean.mkConst ``Nat + +#check Lean.Expr.fvar (Lean.FVarId.mk (Lean.Name.mkStr1 "n")) + +#check Lean.Expr.mvar (Lean.MVarId.mk (Lean.Name.mkStr1 "n")) From e1974f8fe7ae84d476492ce187d803128bd1abc1 Mon Sep 17 00:00:00 2001 From: Siddhartha Gadgil Date: Thu, 21 Nov 2024 09:32:43 +0530 Subject: [PATCH 09/30] test code including some buggy code --- Test/MetaTestable.lean | 56 +++++++++++++++++++++++++----------------- 1 file changed, 33 insertions(+), 23 deletions(-) diff --git a/Test/MetaTestable.lean b/Test/MetaTestable.lean index be02cb4..b66d655 100644 --- a/Test/MetaTestable.lean +++ b/Test/MetaTestable.lean @@ -53,43 +53,53 @@ elab "#decompose_prop" t:term : command => #check MetaTestResult.failure +-- Incorrect elab "mk_failure%" prop:term ";" pf:term ";" : term => do let prop ← elabType prop let notProp ← mkAppM ``Not #[prop] + logInfo s!"{repr notProp}" let pf ← elabTerm pf notProp + logInfo s!"{repr pf}" let lst ← mkAppOptM ``List.nil #[mkConst ``String] let pfExpr := Lean.Expr.letE `pf_var - (notProp) - (pf) + notProp + pf (Lean.Expr.app - (Lean.Expr.const `Lean.Expr.fvar []) - (Lean.Expr.app - (Lean.Expr.const `Lean.FVarId.mk []) - (Lean.Expr.app (Lean.Expr.const `Lean.Name.mkStr1 []) (Lean.Expr.lit (Lean.Literal.strVal "pf_var"))))) - false + (Lean.Expr.const `Lean.Expr.bvar []) + (Lean.Expr.const ``Nat.zero [])) + true logInfo s!"{repr pfExpr}" mkAppOptM ``MetaTestResult.failure #[prop, pf, pfExpr, lst, mkConst ``Nat.zero] -set_option pp.proofs true in -#check mk_failure% False ; fun (x: False) ↦ x ; +#check Lean.Expr.bvar -elab "#expr" e:term : command => - Command.liftTermElabM do - let e ← elabTerm e none - logInfo s!"{repr e}" - logInfo s!"{← reduce e}" +def eg_fail_0 : MetaTestResult False := + mk_failure% False ; fun (x: False) ↦ x ; + +-- Incorrect; checks but extracting expression fails. +#check eg_fail_0 + +def eg_fail : MetaTestResult False := + @MetaTestResult.failure False (fun (x: False) ↦ x) + (Lean.Expr.lam `x (Lean.Expr.const `False []) (Lean.Expr.bvar 0) (Lean.BinderInfo.default)) [] 0 -#check Expr.fvar {name := `n} +def disproofExpr {p: Prop} : MetaTestResult p → MetaM Lean.Expr + | MetaTestResult.failure _ pfExpr _ _ => do + return pfExpr + | _ => + throwError "disproofExpr: expected failure" + +elab "disproof_expr_eg%" : term => do + disproofExpr eg_fail /- -Lean.Expr.fvar (Lean.FVarId.mk (Lean.Name.mkStr1 "n")) +fun x => x : False → False -/ -#expr ( - let pf_var : Nat := 0 - Expr.fvar {name := `pf_var}) -#expr Lean.mkConst ``Nat - -#check Lean.Expr.fvar (Lean.FVarId.mk (Lean.Name.mkStr1 "n")) +#check disproof_expr_eg% -#check Lean.Expr.mvar (Lean.MVarId.mk (Lean.Name.mkStr1 "n")) +elab "#expr" e:term : command => + Command.liftTermElabM do + let e ← elabTerm e none + logInfo s!"{repr e}" + logInfo s!"{← reduce e}" From e7c235db98b8ce987d41569035ccdb211b44f239 Mon Sep 17 00:00:00 2001 From: Siddhartha Gadgil Date: Thu, 21 Nov 2024 19:52:39 +0530 Subject: [PATCH 10/30] fixes; first case working (bug in second case) --- Plausible/MetaTestable.lean | 71 +++++++++++++++++++-------- Test/MetaTestable.lean | 97 ++++++++++++++++++++++++++----------- 2 files changed, 120 insertions(+), 48 deletions(-) diff --git a/Plausible/MetaTestable.lean b/Plausible/MetaTestable.lean index a6a8180..2d480f9 100644 --- a/Plausible/MetaTestable.lean +++ b/Plausible/MetaTestable.lean @@ -126,8 +126,9 @@ def forallProp? (e: Expr) : MetaM (Option Expr × Option Expr) := do return (none, none) def impProp? (e: Expr) : MetaM (Option Expr × Option Expr) := do + let u ← mkFreshLevelMVar + let α ← mkFreshExprMVar (mkSort u) let prop := mkSort levelZero - let α ← mkFreshExprMVar prop let β ← mkFreshExprMVar prop let e' ← mkArrow α β if ← isDefEq e' e then @@ -313,6 +314,7 @@ def imp (h : q → p) (hExp: Expr) (r : MetaTestResult p) (p : Unit ⊕' (p → q) := PSum.inl ()) : MetaM (MetaTestResult q) := match r with | failure h2 pf xs n => do + logInfo s!"Implication: {hExp}; type {← inferType hExp}" let pf' ← mkAppM ``mt #[hExp, pf] return failure (mt h h2) pf' xs n | success h2 => return success <| combine p h2 @@ -329,14 +331,17 @@ can be informative. -/ def addInfo (x : String) (h : q → p) (hExp: Expr) (r : MetaTestResult p) (p : Unit ⊕' (p → q) := PSum.inl ()) : (MetaM <| MetaTestResult q) := do if let failure h2 pf xs n := r then + logInfo s!"Adding info: {x} := {repr x}; type {← inferType hExp}" let pf' ← mkAppM ``mt #[hExp, pf] + logInfo s!"Proof: {pf'}" return failure (mt h h2) pf' (x :: xs) n else imp h hExp r p /-- Add some formatting to the information recorded by `addInfo`. -/ def addVarInfo {γ : Type _} [Repr γ] (var : String) (x : γ) (h : q → p) (hExp: Expr) (r : MetaTestResult p) - (p : Unit ⊕' (p → q) := PSum.inl ()) : MetaM (MetaTestResult q) := + (p : Unit ⊕' (p → q) := PSum.inl ()) : MetaM (MetaTestResult q) := do + logInfo "Adding var info" addInfo s!"{var} := {repr x}" h (hExp: Expr) r p def isFailure : MetaTestResult p → Bool @@ -353,7 +358,9 @@ open MetaTestResult abbrev ProxyExpr α [SampleableExt α] := ToExpr (SampleableExt.proxy α) -def runProp (p : Prop) [MetaTestable p] : Configuration → Bool → Expr → MGen (MetaTestResult p) := MetaTestable.run +def runProp (p : Prop) [MetaTestable p] : Configuration → Bool → Expr → MGen (MetaTestResult p) := fun cfg b e => do + logInfo s!"Testing {e}" + MetaTestable.run cfg b e /-- A `dbgTrace` with special formatting -/ def slimTrace {m : Type → Type _} [Pure m] (s : String) : m PUnit := @@ -399,9 +406,14 @@ instance decGuardTestable [PrintableProp p] [Decidable p] {β : p → Prop} [∀ MetaTestable (NamedBinder var <| ∀ h, β h) where run := fun cfg min e => do if h : p then - let res := runProp (β h) cfg min e - let s := printProp p - ← (fun r => addInfo s!"guard: {s}" (· <| h) e r (PSum.inr <| fun q _ => q)) <$> res + match ← forallProp? e with + | (some βExp, some pExp) => do + let yExp ← mkAppM' βExp #[pExp] + let res := runProp (β h) cfg min yExp + let h' := (· <| h) + let s := printProp p + ← (fun r => addInfo s!"guard: {s}" h' e r (PSum.inr <| fun q _ => q)) <$> res + | _ => throwError m!"Expected a `Forall` proposition, got {← ppExpr e}" else if cfg.traceDiscarded || cfg.traceSuccesses then let res := fun _ => return gaveUp 1 let s := printProp p @@ -454,6 +466,7 @@ partial def minimizeAux [SampleableExt α] [ProxyExpr α] {β : α → Prop} [ let samp ← synthInstance <| ← mkAppM ``SampleableExt #[αExp] let xInterp ← mkAppOptM ``SampleableExt.interp #[αExp, samp, xExpr] let e' ← mkAppM' βExpr #[xInterp] + logInfo "Minimize" let res ← OptionT.lift <| MetaTestable.runProp (β (SampleableExt.interp candidate)) cfg true e' if res.isFailure then if cfg.traceShrink then @@ -490,6 +503,7 @@ instance varTestable [SampleableExt α] [ProxyExpr α] {β : α → Prop} [∀ x let e' ← mkAppM' βExp #[xInterp] if cfg.traceSuccesses || cfg.traceDiscarded then slimTrace s!"{var} := {repr x}" + logInfo s!"variable testable: outer prop: {e}, inner prop: {e'}" let r ← MetaTestable.runProp (β <| SampleableExt.interp x) cfg false e' let ⟨finalX, finalR⟩ ← if isFailure r then @@ -540,6 +554,7 @@ where run := fun cfg min e => do match ← forallProp? e with | (some βExpr, _) => + logInfo s!"prop variable testable: outer prop: {e}, inner prop: {βExpr}" let p ← MetaTestable.runProp (NamedBinder var <| ∀ b : Bool, β b) cfg min e let e' ← mkAppM ``bool_to_prop_fmly #[βExpr] imp (bool_to_prop_fmly β) e' p @@ -551,9 +566,20 @@ where run := fun cfg min e => do if cfg.traceDiscarded || cfg.traceSuccesses then slimTrace s!"{var} is unused" - let r ← MetaTestable.runProp β cfg min e - let finalR ← addInfo s!"{var} is irrelevant (unused)" id e r - imp (· <| Classical.ofNonempty) e finalR (PSum.inr <| fun x _ => x) + logInfo s!"unused variable testable: outer prop: {e}" + match ← impProp? e with + | (some aExp, some e') => + let r ← MetaTestable.runProp β cfg min e' + let hExp ← mkAppOptM ``id #[e'] + let finalR ← addInfo s!"{var} is irrelevant (unused)" id hExp r + let h := (· <| Classical.ofNonempty) + let nInst ← synthInstance <| ← mkAppM ``Nonempty #[aExp] + logInfo "building h expression" + let hExp ← withLocalDeclD `h e fun h => do + mkLambdaFVars #[h] <| mkApp h nInst + logInfo m!"Calling imp; hExpr of type {← inferType hExp}" + imp h hExp finalR (PSum.inr <| fun x _ => x) + | _ => throwError m!"Expected an `Imp` proposition, got {← ppExpr e}" instance (priority := 2000) subtypeVarTestable {p : α → Prop} {β : α → Prop} [∀ x, PrintableProp (p x)] @@ -579,14 +605,17 @@ instance (priority := low) decidableTestable {p : Prop} [PrintableProp p] [Decid let s := printProp p let inst ← synthInstance <| ← mkAppM ``Decidable #[e] let falseRefl ← mkAppM ``Eq.refl #[mkConst ``false] + logInfo "Decidable used; building proof" let pf' ← mkAppOptM ``of_decide_eq_false #[e, inst, falseRefl] + logInfo "Decidable proof built" checkDisproof pf' e + logInfo "Decidable proof checked" return failure h pf' [s!"issue: {s} does not hold"] 0 end MetaTestable -section IO +section Meta open MetaTestResult namespace MetaTestable @@ -629,29 +658,33 @@ def MetaTestable.runSuite (p : Prop) [MetaTestable p] (propExpr: Expr) (cfg : Co MetaTestable.runSuiteAux p propExpr cfg (success <| PSum.inl ()) cfg.numInst /-- Run a test suite for `p` in `MetaM` using the global RNG in `stdGenRef`. -/ -def MetaTestable.checkIO (p : Prop) [MetaTestable p] (propExpr: Expr) (cfg : Configuration := {}) : MetaM (MetaTestResult p) := +def MetaTestable.checkM (p : Prop) [MetaTestable p] (cfg : Configuration := {}) (propExpr: Expr) : MetaM (MetaTestResult p) := match cfg.randomSeed with | none => runRand (MetaTestable.runSuite p propExpr cfg) | some seed => runRandWith seed (MetaTestable.runSuite p propExpr cfg) -end IO +end Meta open Decorations in /-- Run a test suite for `p` and throw an exception if `p` does not hold. -/ -def MetaTestable.check (p : Prop)(propExpr: Expr) (cfg : Configuration := {}) - (p' : Decorations.DecorationsOf p := by mk_decorations) [MetaTestable p'] : Lean.MetaM PUnit := do - match ← MetaTestable.checkIO p' propExpr cfg with - | MetaTestResult.success _ => if !cfg.quiet then Lean.logInfo "Unable to find a counter-example" +def MetaTestable.check (p : Prop) (cfg : Configuration := {}) + (p' : Decorations.DecorationsOf p := by mk_decorations) [MetaTestable p'](propExpr: Expr) : Lean.MetaM (Option Expr) := do + match ← MetaTestable.checkM p' cfg propExpr with + | MetaTestResult.success _ => + if !cfg.quiet then Lean.logInfo "Unable to find a counter-example" + return none | MetaTestResult.gaveUp n => if !cfg.quiet then let msg := s!"Gave up after failing to generate values that fulfill the preconditions {n} times." Lean.logWarning msg - | MetaTestResult.failure _ _ xs n => + return none + | MetaTestResult.failure _ pf xs n => let msg := "Found a counter-example!" if cfg.quiet then - Lean.throwError msg + Lean.logWarning msg else - Lean.throwError <| Testable.formatFailure msg xs n + Lean.logWarning <| Testable.formatFailure msg xs n + return some pf -- #eval MetaTestable.check (∀ (x y z a : Nat) (h1 : 3 < x) (h2 : 3 < y), x - y = y - x) -- Configuration.verbose diff --git a/Test/MetaTestable.lean b/Test/MetaTestable.lean index b66d655..8ce8a5e 100644 --- a/Test/MetaTestable.lean +++ b/Test/MetaTestable.lean @@ -1,7 +1,8 @@ import Lean import Plausible.MetaTestable -open Plausible Plausible.MetaTestable Lean Meta Elab Term +open Plausible Plausible.MetaTestable +open Lean Meta Elab Term Tactic open Lean Elab Term in elab "#decompose_prop" t:term : command => @@ -20,7 +21,7 @@ elab "#decompose_prop" t:term : command => | (some α, some β) => logInfo s!"Forall: {← ppExpr α}; domain {← ppExpr β}" | _ => pure () match ← impProp? e with - | (some α, some β) => logInfo s!"Imp: {← ppExpr α}; {← ppExpr β}" + | (some α, some β) => logInfo s!"Implies: {← ppExpr α}; {← ppExpr β}" | _ => pure () match ← eqlProp? e with | some (α, a, b) => logInfo s!"Eq: {← ppExpr α}; {← ppExpr a} and {← ppExpr b}" @@ -49,36 +50,21 @@ elab "#decompose_prop" t:term : command => #guard_msgs in #decompose_prop ∃ (n: Nat), n = 0 ∨ n ≠ 0 -#check List.nil +/-- +info: Forall: fun x => 2 ≠ 0; domain 1 = 0 +--- +info: Implies: 1 = 0; 2 ≠ 0 +-/ +#guard_msgs in +#decompose_prop 1 = 0 → 2 ≠ 0 + + +/-- info: Iff: 1 = 0; 2 ≠ 0 -/ +#guard_msgs in +#decompose_prop 1 = 0 ↔ 2 ≠ 0 #check MetaTestResult.failure --- Incorrect -elab "mk_failure%" prop:term ";" pf:term ";" : term => do - let prop ← elabType prop - let notProp ← mkAppM ``Not #[prop] - logInfo s!"{repr notProp}" - let pf ← elabTerm pf notProp - logInfo s!"{repr pf}" - let lst ← mkAppOptM ``List.nil #[mkConst ``String] - let pfExpr := Lean.Expr.letE - `pf_var - notProp - pf - (Lean.Expr.app - (Lean.Expr.const `Lean.Expr.bvar []) - (Lean.Expr.const ``Nat.zero [])) - true - logInfo s!"{repr pfExpr}" - mkAppOptM ``MetaTestResult.failure #[prop, pf, pfExpr, lst, mkConst ``Nat.zero] - -#check Lean.Expr.bvar - -def eg_fail_0 : MetaTestResult False := - mk_failure% False ; fun (x: False) ↦ x ; - --- Incorrect; checks but extracting expression fails. -#check eg_fail_0 def eg_fail : MetaTestResult False := @MetaTestResult.failure False (fun (x: False) ↦ x) @@ -103,3 +89,56 @@ elab "#expr" e:term : command => let e ← elabTerm e none logInfo s!"{repr e}" logInfo s!"{← reduce e}" + +elab "expr%" e:term : term => do + let e ← elabTerm e none + logInfo s!"{repr e}" + logInfo s!"{← reduce e}" + return e +-- + +structure MyType where + x : Nat + y : Nat + h : x ≤ y + deriving Repr + +#check fun (x y : Nat) (h : x ≤ y) => expr% (MyType.mk x y h) + +instance : Shrinkable MyType where + shrink := fun ⟨x, y, _⟩ => + let proxy := Shrinkable.shrink (x, y - x) + proxy.map (fun (fst, snd) => ⟨fst, fst + snd, by omega⟩) + +instance : SampleableExt MyType := + SampleableExt.mkSelfContained do + let x ← SampleableExt.interpSample Nat + let xyDiff ← SampleableExt.interpSample Nat + return ⟨x, x + xyDiff, by omega⟩ + +set_option diagnostics true in +-- #eval MetaTestable.check <| ∀ a b : MyType, a.y ≤ b.x → a.x ≤ b.y + +open Decorations Lean Elab Tactic + +elab "mk_decorations" : tactic => do + let goal ← getMainGoal + let goalType ← goal.getType + if let .app (.const ``Decorations.DecorationsOf _) body := goalType then + closeMainGoal `mk_decorations (← Decorations.addDecorations body) + + +def de : Decorations.DecorationsOf (∀ a b : Nat, a ≤ b) := by mk_decorations + +#print de + +#synth MetaTestable <| (1: Nat) = 0 + +#synth Testable (NamedBinder "a" (∀ (a : Nat), a ≤ 1)) + +#synth MetaTestable (NamedBinder "a" (∀ (a : Nat), a ≤ 1)) + +#synth MetaTestable (NamedBinder "a" (∀ (a : Nat), NamedBinder "b" (∀ (b : Nat), a ≤ b))) + + +#eval MetaTestable.check (∀ (a : Nat), False) (propExpr := Lean.Expr.forallE `a (Lean.Expr.const `Nat []) (Lean.Expr.const `False []) (Lean.BinderInfo.default)) From 01760653807f48450f2e6cadd2c96516d93a78a0 Mon Sep 17 00:00:00 2001 From: Siddhartha Gadgil Date: Thu, 21 Nov 2024 21:26:32 +0530 Subject: [PATCH 11/30] fixed another case --- Plausible/MetaTestable.lean | 22 ++++++++++++++++++---- Test/MetaTestable.lean | 16 ++++++++++++++++ 2 files changed, 34 insertions(+), 4 deletions(-) diff --git a/Plausible/MetaTestable.lean b/Plausible/MetaTestable.lean index 2d480f9..6ac25bf 100644 --- a/Plausible/MetaTestable.lean +++ b/Plausible/MetaTestable.lean @@ -463,7 +463,10 @@ partial def minimizeAux [SampleableExt α] [ProxyExpr α] {β : α → Prop} [ if cfg.traceShrinkCandidates then slimTrace s!"Trying {var} := {repr candidate}" let xExpr := toExpr candidate - let samp ← synthInstance <| ← mkAppM ``SampleableExt #[αExp] + let .sort u := ← inferType αExp | throwError m!"Expected a sort, got {αExp}" + let v ← mkFreshLevelMVar + let instType := mkApp (mkConst ``SampleableExt [u, v]) αExp + let samp ← synthInstance instType let xInterp ← mkAppOptM ``SampleableExt.interp #[αExp, samp, xExpr] let e' ← mkAppM' βExpr #[xInterp] logInfo "Minimize" @@ -498,7 +501,12 @@ instance varTestable [SampleableExt α] [ProxyExpr α] {β : α → Prop} [∀ x | (some βExp, some αExp) => do let x ← SampleableExt.sample let xExpr := toExpr x - let samp ← synthInstance <| ← mkAppM ``SampleableExt #[αExp] + let αExp ← instantiateMVars αExp + let .sort u := ← inferType αExp | throwError m!"Expected a sort, got {αExp}" + let v ← mkFreshLevelMVar + let instType := mkApp (mkConst ``SampleableExt [u, v]) αExp + let samp ← synthInstance instType + -- logInfo s!"Instance: {samp}" let xInterp ← mkAppOptM ``SampleableExt.interp #[αExp, samp, xExpr] let e' ← mkAppM' βExp #[xInterp] if cfg.traceSuccesses || cfg.traceDiscarded then @@ -516,10 +524,16 @@ instance varTestable [SampleableExt α] [ProxyExpr α] {β : α → Prop} [∀ x else pure ⟨x, r⟩ let xExpr := toExpr finalX - let samp ← synthInstance <| ← mkAppM ``SampleableExt #[αExp] + let .sort u := ← inferType αExp | throwError m!"Expected a sort, got {αExp}" + let v ← mkFreshLevelMVar + let instType := mkApp (mkConst ``SampleableExt [u, v]) αExp + let samp ← synthInstance instType let xInterp ← mkAppOptM ``SampleableExt.interp #[αExp, samp, xExpr] + let h := (· <| SampleableExt.interp finalX) let e' ← mkAppM' βExp #[xInterp] - addVarInfo var finalX (· <| SampleableExt.interp finalX) e' finalR + let hExpr ← withLocalDeclD `h e fun h => do + mkLambdaFVars #[h] (mkApp h e') + addVarInfo var finalX h hExpr finalR | (_, _) => throwError m!"Expected a `Forall` proposition, got {← ppExpr e}" diff --git a/Test/MetaTestable.lean b/Test/MetaTestable.lean index 8ce8a5e..af33b79 100644 --- a/Test/MetaTestable.lean +++ b/Test/MetaTestable.lean @@ -142,3 +142,19 @@ def de : Decorations.DecorationsOf (∀ a b : Nat, a ≤ b) := by mk_decorations #eval MetaTestable.check (∀ (a : Nat), False) (propExpr := Lean.Expr.forallE `a (Lean.Expr.const `Nat []) (Lean.Expr.const `False []) (Lean.BinderInfo.default)) + +/- +AppBuilder for 'mkAppM', result contains metavariables + SampleableExt Nat +-/ +set_option pp.universes true +#eval MetaTestable.check (∀ (a : Nat), a < 1) (propExpr := Lean.Expr.forallE `a (Lean.Expr.const `Nat []) (Lean.Expr.const `False []) (Lean.BinderInfo.default)) + +set_option pp.all true in +-- #eval MetaTestable.check (∀ (a b : Nat), a < b) (propExpr := Lean.Expr.forallE `a (Lean.Expr.const `Nat []) (Lean.Expr.const `False []) (Lean.BinderInfo.default)) + +def samp := SampleableExt Nat +#print samp + +#synth SampleableExt Nat +#check Nat.sampleableExt From 189afce5735a8b1386e4ef2d7dbe8118763270a6 Mon Sep 17 00:00:00 2001 From: Siddhartha Gadgil Date: Fri, 22 Nov 2024 11:09:08 +0530 Subject: [PATCH 12/30] bug fixed, more tests --- Plausible/MetaTestable.lean | 21 +++++++++++++--- Test/MetaTestable.lean | 50 +++++++++++++++++++++++++++++++------ 2 files changed, 60 insertions(+), 11 deletions(-) diff --git a/Plausible/MetaTestable.lean b/Plausible/MetaTestable.lean index 6ac25bf..ac8723d 100644 --- a/Plausible/MetaTestable.lean +++ b/Plausible/MetaTestable.lean @@ -3,6 +3,7 @@ Copyright (c) 2022 Henrik Böving. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Henrik Böving, Simon Hudon -/ +import Lean.Meta import Lean.Elab.Tactic.Config import Plausible.Sampleable import Plausible.Testable @@ -332,6 +333,8 @@ def addInfo (x : String) (h : q → p) (hExp: Expr) (r : MetaTestResult p) (p : Unit ⊕' (p → q) := PSum.inl ()) : (MetaM <| MetaTestResult q) := do if let failure h2 pf xs n := r then logInfo s!"Adding info: {x} := {repr x}; type {← inferType hExp}" + logInfo s!"proof type: {← ppExpr <| ← inferType pf}" + logInfo s!"hExp type: {← ppExpr <| ← inferType hExp}" let pf' ← mkAppM ``mt #[hExp, pf] logInfo s!"Proof: {pf'}" return failure (mt h h2) pf' (x :: xs) n @@ -411,8 +414,17 @@ instance decGuardTestable [PrintableProp p] [Decidable p] {β : p → Prop} [∀ let yExp ← mkAppM' βExp #[pExp] let res := runProp (β h) cfg min yExp let h' := (· <| h) + let decInstType ← mkAppM ``Decidable #[pExp] + let inst ← synthInstance decInstType + let falseRefl ← mkAppM ``Eq.refl #[mkConst ``false] + logInfo "Decidable used; building proof" + let pf ← mkAppOptM ``of_decide_eq_true #[pExp, inst, falseRefl] + let cod := mkApp βExp pf + let hExp ← withLocalDeclD `x (← mkArrow pExp cod) fun x => do + let y ← mkAppM' x #[cod] + mkLambdaFVars #[x] y let s := printProp p - ← (fun r => addInfo s!"guard: {s}" h' e r (PSum.inr <| fun q _ => q)) <$> res + ← (fun r => addInfo s!"guard: {s}" h' hExp r (PSum.inr <| fun q _ => q)) <$> res | _ => throwError m!"Expected a `Forall` proposition, got {← ppExpr e}" else if cfg.traceDiscarded || cfg.traceSuccesses then let res := fun _ => return gaveUp 1 @@ -492,6 +504,7 @@ def minimize [SampleableExt α] [ProxyExpr α] {β : α → Prop} [∀ x, MetaTe let res ← OptionT.run <| minimizeAux αExp βExpr cfg var x 0 return res.getD ⟨x, r⟩ +open Lean Meta Elab Term Tactic in /-- Test a universal property by creating a sample of the right type and instantiating the bound variable with it. -/ instance varTestable [SampleableExt α] [ProxyExpr α] {β : α → Prop} [∀ x, MetaTestable (β x)] : @@ -509,6 +522,7 @@ instance varTestable [SampleableExt α] [ProxyExpr α] {β : α → Prop} [∀ x -- logInfo s!"Instance: {samp}" let xInterp ← mkAppOptM ``SampleableExt.interp #[αExp, samp, xExpr] let e' ← mkAppM' βExp #[xInterp] + let (e', _) ← dsimp e' {} if cfg.traceSuccesses || cfg.traceDiscarded then slimTrace s!"{var} := {repr x}" logInfo s!"variable testable: outer prop: {e}, inner prop: {e'}" @@ -531,8 +545,9 @@ instance varTestable [SampleableExt α] [ProxyExpr α] {β : α → Prop} [∀ x let xInterp ← mkAppOptM ``SampleableExt.interp #[αExp, samp, xExpr] let h := (· <| SampleableExt.interp finalX) let e' ← mkAppM' βExp #[xInterp] - let hExpr ← withLocalDeclD `h e fun h => do - mkLambdaFVars #[h] (mkApp h e') + logInfo s!"Adding var info from varTestable, goal : {← ppExpr e}" + let hExpr ← withLocalDeclD `x e fun x => do + mkLambdaFVars #[x] (mkApp x xInterp) addVarInfo var finalX h hExpr finalR | (_, _) => throwError m!"Expected a `Forall` proposition, got {← ppExpr e}" diff --git a/Test/MetaTestable.lean b/Test/MetaTestable.lean index af33b79..5b82b65 100644 --- a/Test/MetaTestable.lean +++ b/Test/MetaTestable.lean @@ -147,14 +147,48 @@ def de : Decorations.DecorationsOf (∀ a b : Nat, a ≤ b) := by mk_decorations AppBuilder for 'mkAppM', result contains metavariables SampleableExt Nat -/ -set_option pp.universes true +set_option pp.universes true in #eval MetaTestable.check (∀ (a : Nat), a < 1) (propExpr := Lean.Expr.forallE `a (Lean.Expr.const `Nat []) (Lean.Expr.const `False []) (Lean.BinderInfo.default)) -set_option pp.all true in --- #eval MetaTestable.check (∀ (a b : Nat), a < b) (propExpr := Lean.Expr.forallE `a (Lean.Expr.const `Nat []) (Lean.Expr.const `False []) (Lean.BinderInfo.default)) -def samp := SampleableExt Nat -#print samp - -#synth SampleableExt Nat -#check Nat.sampleableExt +#eval MetaTestable.check (∀ (a b : Nat), a < b) (propExpr := (Lean.Expr.forallE + `a + (Lean.Expr.const `Nat []) + (Lean.Expr.forallE + `b + (Lean.Expr.const `Nat []) + (Lean.Expr.app + (Lean.Expr.app + (Lean.Expr.app + (Lean.Expr.app + (Lean.Expr.const `LT.lt [Level.succ Level.zero]) + (Lean.Expr.const `Nat [])) + (Lean.Expr.const `instLTNat [])) + (Lean.Expr.bvar 1)) + (Lean.Expr.bvar 0)) + (Lean.BinderInfo.default)) + (Lean.BinderInfo.default))) + + +#eval MetaTestable.check (∀ (a b : Nat), a < 0) (propExpr := (Lean.Expr.forallE + `a + (Lean.Expr.const `Nat []) + (Lean.Expr.forallE + `b + (Lean.Expr.const `Nat []) + (Lean.Expr.app + (Lean.Expr.app + (Lean.Expr.app + (Lean.Expr.app + (Lean.Expr.const `LT.lt [Level.succ Level.zero]) + (Lean.Expr.const `Nat [])) + (Lean.Expr.const `instLTNat [])) + (Lean.Expr.bvar 1)) + (Lean.Expr.lit (Lean.Literal.natVal 0))) + (Lean.BinderInfo.default)) + (Lean.BinderInfo.default))) + +#expr ∀ (a b : Nat), a < b +#expr ∀ (a b : Nat), a < 0 + +#check Lean.Expr.lit (Lean.Literal.natVal 0) From 6384d06050b333d04b6976edb79b6a95b3405161 Mon Sep 17 00:00:00 2001 From: Siddhartha Gadgil Date: Fri, 22 Nov 2024 14:56:52 +0530 Subject: [PATCH 13/30] function, elaborator for disproof; bugs remain: `and`, `or`, maybe others --- Plausible/MetaTestable.lean | 234 ++++++++++++++++++------------------ Test/MetaTestable.lean | 16 ++- 2 files changed, 133 insertions(+), 117 deletions(-) diff --git a/Plausible/MetaTestable.lean b/Plausible/MetaTestable.lean index ac8723d..3694b44 100644 --- a/Plausible/MetaTestable.lean +++ b/Plausible/MetaTestable.lean @@ -262,27 +262,24 @@ def checkDisproof (pf prop: Expr) : MetaM Unit := do /-- Combine the test result for properties `p` and `q` to create a test for their conjunction. -/ def and : MetaTestResult p → MetaTestResult q → Expr → MetaM (MetaTestResult (p ∧ q)) | failure h pf xs n, _, e => do - match ← andProp? e with - | (some e₁, some e₂) => - let pf' ← withLocalDeclD `h e fun h => do - let x ← mkAppOptM ``And.left #[e₁, e₂, h] - let e' ← mkAppM' pf #[x] - mkLambdaFVars #[h] e' - checkDisproof pf' e - return failure (fun h2 => h h2.left) pf' xs n - | (_, _) => throwError m!"Expected an `And` proposition, got {← ppExpr e}" + let (some e₁, some e₂) ← andProp? e | throwError m!"Expected an `And` proposition, got {← ppExpr e}" + let pf' ← withLocalDeclD `h e fun h => do + let x ← mkAppOptM ``And.left #[e₁, e₂, h] + let e' ← mkAppM' pf #[x] + mkLambdaFVars #[h] e' + checkDisproof pf' e + return failure (fun h2 => h h2.left) pf' xs n | _, failure h pf xs n, e => do - match ← andProp? e with - | (some e₁, some e₂) => - let pf' ← withLocalDeclD `h e fun h => do - let x ← mkAppOptM ``And.right #[e₁, e₂, h] - let e' ← mkAppM' pf #[x] - mkLambdaFVars #[h] e' - checkDisproof pf' e - return failure (fun h2 => h h2.right) pf' xs n - | (_, _) => throwError m!"Expected an `And` proposition, got {← ppExpr e}" - - | success h1, success h2, _ => return success <| combine (combine (PSum.inr And.intro) h1) h2 + let (some e₁, some e₂) ← andProp? e | throwError m!"Expected an `And` proposition, got {← ppExpr e}" + + let pf' ← withLocalDeclD `h e fun h => do + let x ← mkAppOptM ``And.right #[e₁, e₂, h] + let e' ← mkAppM' pf #[x] + mkLambdaFVars #[h] e' + checkDisproof pf' e + return failure (fun h2 => h h2.right) pf' xs n + | success h1, success h2, _ => + return success <| combine (combine (PSum.inr And.intro) h1) h2 | gaveUp n, gaveUp m, _ => return gaveUp <| n + m | gaveUp n, _, _ => return gaveUp n | _, gaveUp n, _ => return gaveUp n @@ -291,18 +288,16 @@ def and : MetaTestResult p → MetaTestResult q → Expr → MetaM (MetaTestRes /-- Combine the test result for properties `p` and `q` to create a test for their disjunction. -/ def or : MetaTestResult p → MetaTestResult q → Expr → MetaM (MetaTestResult (p ∨ q)) | failure h1 pf1 xs n, failure h2 pf2 ys m, e => do - match ← orProp? e with - | (some α, some β) => - let motive ← withLocalDeclD `h e fun h => do - mkLambdaFVars #[h] <| mkConst ``False - let pf ← mkAppOptM ``Or.rec #[α, β, motive, pf1, pf2] - let h3 := fun h => - match h with - | Or.inl h3 => h1 h3 - | Or.inr h3 => h2 h3 - checkDisproof pf e - return failure h3 pf (xs ++ ys) (n + m) - | (_, _) => throwError m!"Expected an `Or` proposition, got {← ppExpr e}" + let (some α, some β) ← orProp? e | throwError m!"Expected an `Or` proposition, got {← ppExpr e}" + let motive ← withLocalDeclD `h e fun h => do + mkLambdaFVars #[h] <| mkConst ``False + let pf ← mkAppOptM ``Or.rec #[α, β, motive, pf1, pf2] + let h3 := fun h => + match h with + | Or.inl h3 => h1 h3 + | Or.inr h3 => h2 h3 + checkDisproof pf e + return failure h3 pf (xs ++ ys) (n + m) | success h, _, _ => return success <| combine (PSum.inr Or.inl) h | _, success h, _ => return success <| combine (PSum.inr Or.inr) h | gaveUp n, gaveUp m, _ => return gaveUp <| n + m @@ -396,12 +391,10 @@ theorem iff_resolve (p q : Prop) : (p ↔ q) ↔ p ∧ q ∨ ¬p ∧ ¬q := by instance iffTestable [MetaTestable ((p ∧ q) ∨ (¬ p ∧ ¬ q))] : MetaTestable (p ↔ q) where run := fun cfg min e => do - match ← iffProp? e with - | some (α, β) => do - let h ← runProp ((p ∧ q) ∨ (¬ p ∧ ¬ q)) cfg min e - let hExp ← mkAppM ``iff_resolve #[α, β] - iff (iff_resolve p q) hExp h - | none => throwError m!"Expected an `Iff` proposition, got {← ppExpr e}" + let some (α, β) ← iffProp? e | throwError m!"Expected an `Iff` proposition, got {← ppExpr e}" + let h ← runProp ((p ∧ q) ∨ (¬ p ∧ ¬ q)) cfg min e + let hExp ← mkAppM ``iff_resolve #[α, β] + iff (iff_resolve p q) hExp h variable {var : String} @@ -409,23 +402,21 @@ instance decGuardTestable [PrintableProp p] [Decidable p] {β : p → Prop} [∀ MetaTestable (NamedBinder var <| ∀ h, β h) where run := fun cfg min e => do if h : p then - match ← forallProp? e with - | (some βExp, some pExp) => do - let yExp ← mkAppM' βExp #[pExp] - let res := runProp (β h) cfg min yExp - let h' := (· <| h) - let decInstType ← mkAppM ``Decidable #[pExp] - let inst ← synthInstance decInstType - let falseRefl ← mkAppM ``Eq.refl #[mkConst ``false] - logInfo "Decidable used; building proof" - let pf ← mkAppOptM ``of_decide_eq_true #[pExp, inst, falseRefl] - let cod := mkApp βExp pf - let hExp ← withLocalDeclD `x (← mkArrow pExp cod) fun x => do - let y ← mkAppM' x #[cod] - mkLambdaFVars #[x] y - let s := printProp p - ← (fun r => addInfo s!"guard: {s}" h' hExp r (PSum.inr <| fun q _ => q)) <$> res - | _ => throwError m!"Expected a `Forall` proposition, got {← ppExpr e}" + let (some βExp, some pExp) ← forallProp? e | throwError m!"Expected a `Forall` proposition, got {← ppExpr e}" + let yExp ← mkAppM' βExp #[pExp] + let res := runProp (β h) cfg min yExp + let h' := (· <| h) + let decInstType ← mkAppM ``Decidable #[pExp] + let inst ← synthInstance decInstType + let falseRefl ← mkAppM ``Eq.refl #[mkConst ``false] + logInfo "Decidable used; building proof" + let pf ← mkAppOptM ``of_decide_eq_true #[pExp, inst, falseRefl] + let cod := mkApp βExp pf + let hExp ← withLocalDeclD `x (← mkArrow pExp cod) fun x => do + let y ← mkAppM' x #[cod] + mkLambdaFVars #[x] y + let s := printProp p + ← (fun r => addInfo s!"guard: {s}" h' hExp r (PSum.inr <| fun q _ => q)) <$> res else if cfg.traceDiscarded || cfg.traceSuccesses then let res := fun _ => return gaveUp 1 let s := printProp p @@ -510,46 +501,44 @@ bound variable with it. -/ instance varTestable [SampleableExt α] [ProxyExpr α] {β : α → Prop} [∀ x, MetaTestable (β x)] : MetaTestable (NamedBinder var <| ∀ x : α, β x) where run := fun cfg min e => do - match ← forallProp? e with - | (some βExp, some αExp) => do - let x ← SampleableExt.sample - let xExpr := toExpr x - let αExp ← instantiateMVars αExp - let .sort u := ← inferType αExp | throwError m!"Expected a sort, got {αExp}" - let v ← mkFreshLevelMVar - let instType := mkApp (mkConst ``SampleableExt [u, v]) αExp - let samp ← synthInstance instType - -- logInfo s!"Instance: {samp}" - let xInterp ← mkAppOptM ``SampleableExt.interp #[αExp, samp, xExpr] - let e' ← mkAppM' βExp #[xInterp] - let (e', _) ← dsimp e' {} - if cfg.traceSuccesses || cfg.traceDiscarded then - slimTrace s!"{var} := {repr x}" - logInfo s!"variable testable: outer prop: {e}, inner prop: {e'}" - let r ← MetaTestable.runProp (β <| SampleableExt.interp x) cfg false e' - let ⟨finalX, finalR⟩ ← - if isFailure r then - if cfg.traceSuccesses then - slimTrace s!"{var} := {repr x} is a failure" - if min then - minimize αExp βExp cfg var x r - else - pure ⟨x, r⟩ + let (some βExp, some αExp) ← forallProp? e | throwError m!"Expected a `Forall` proposition, got {← ppExpr e}" + let x ← SampleableExt.sample + let xExpr := toExpr x + let αExp ← instantiateMVars αExp + let .sort u := ← inferType αExp | throwError m!"Expected a sort, got {αExp}" + let v ← mkFreshLevelMVar + let instType := mkApp (mkConst ``SampleableExt [u, v]) αExp + let samp ← synthInstance instType + -- logInfo s!"Instance: {samp}" + let xInterp ← mkAppOptM ``SampleableExt.interp #[αExp, samp, xExpr] + let e' ← mkAppM' βExp #[xInterp] + let (e', _) ← dsimp e' {} + if cfg.traceSuccesses || cfg.traceDiscarded then + slimTrace s!"{var} := {repr x}" + logInfo s!"variable testable: outer prop: {e}, inner prop: {e'}" + let r ← MetaTestable.runProp (β <| SampleableExt.interp x) cfg false e' + let ⟨finalX, finalR⟩ ← + if isFailure r then + if cfg.traceSuccesses then + slimTrace s!"{var} := {repr x} is a failure" + if min then + minimize αExp βExp cfg var x r else pure ⟨x, r⟩ - let xExpr := toExpr finalX - let .sort u := ← inferType αExp | throwError m!"Expected a sort, got {αExp}" - let v ← mkFreshLevelMVar - let instType := mkApp (mkConst ``SampleableExt [u, v]) αExp - let samp ← synthInstance instType - let xInterp ← mkAppOptM ``SampleableExt.interp #[αExp, samp, xExpr] - let h := (· <| SampleableExt.interp finalX) - let e' ← mkAppM' βExp #[xInterp] - logInfo s!"Adding var info from varTestable, goal : {← ppExpr e}" - let hExpr ← withLocalDeclD `x e fun x => do - mkLambdaFVars #[x] (mkApp x xInterp) - addVarInfo var finalX h hExpr finalR - | (_, _) => throwError m!"Expected a `Forall` proposition, got {← ppExpr e}" + else + pure ⟨x, r⟩ + let xExpr := toExpr finalX + let .sort u ← inferType αExp | throwError m!"Expected a sort, got {αExp}" + let v ← mkFreshLevelMVar + let instType := mkApp (mkConst ``SampleableExt [u, v]) αExp + let samp ← synthInstance instType + let xInterp ← mkAppOptM ``SampleableExt.interp #[αExp, samp, xExpr] + let h := (· <| SampleableExt.interp finalX) + logInfo s!"Adding var info from varTestable, goal : {← ppExpr e}" + let hExpr ← withLocalDeclD `x e fun x => do + mkLambdaFVars #[x] (mkApp x xInterp) + addVarInfo var finalX h hExpr finalR + instance : ProxyExpr Bool := (inferInstance : ToExpr Bool) @@ -581,13 +570,11 @@ instance propVarTestable {β : Prop → Prop} [h: ∀ b : Bool, MetaTestable (β MetaTestable (NamedBinder var <| ∀ p : Prop, β p) where run := fun cfg min e => do - match ← forallProp? e with - | (some βExpr, _) => - logInfo s!"prop variable testable: outer prop: {e}, inner prop: {βExpr}" - let p ← MetaTestable.runProp (NamedBinder var <| ∀ b : Bool, β b) cfg min e - let e' ← mkAppM ``bool_to_prop_fmly #[βExpr] - imp (bool_to_prop_fmly β) e' p - | _ => throwError m!"Expected a `Forall` proposition, got {← ppExpr e}" + let (some βExpr, _) ← forallProp? e | throwError m!"Expected a `Forall` proposition, got {← ppExpr e}" + logInfo s!"prop variable testable: outer prop: {e}, inner prop: {βExpr}" + let p ← MetaTestable.runProp (NamedBinder var <| ∀ b : Bool, β b) cfg min e + let e' ← mkAppM ``bool_to_prop_fmly #[βExpr] + imp (bool_to_prop_fmly β) e' p instance (priority := high) unusedVarTestable {β : Prop} [Nonempty α] [MetaTestable β] : MetaTestable (NamedBinder var (α → β)) @@ -596,19 +583,17 @@ where if cfg.traceDiscarded || cfg.traceSuccesses then slimTrace s!"{var} is unused" logInfo s!"unused variable testable: outer prop: {e}" - match ← impProp? e with - | (some aExp, some e') => - let r ← MetaTestable.runProp β cfg min e' - let hExp ← mkAppOptM ``id #[e'] - let finalR ← addInfo s!"{var} is irrelevant (unused)" id hExp r - let h := (· <| Classical.ofNonempty) - let nInst ← synthInstance <| ← mkAppM ``Nonempty #[aExp] - logInfo "building h expression" - let hExp ← withLocalDeclD `h e fun h => do - mkLambdaFVars #[h] <| mkApp h nInst - logInfo m!"Calling imp; hExpr of type {← inferType hExp}" - imp h hExp finalR (PSum.inr <| fun x _ => x) - | _ => throwError m!"Expected an `Imp` proposition, got {← ppExpr e}" + let (some aExp, some e') ← impProp? e | throwError m!"Expected an `Imp` proposition, got {← ppExpr e}" + let r ← MetaTestable.runProp β cfg min e' + let hExp ← mkAppOptM ``id #[e'] + let finalR ← addInfo s!"{var} is irrelevant (unused)" id hExp r + let h := (· <| Classical.ofNonempty) + let nInst ← synthInstance <| ← mkAppM ``Nonempty #[aExp] + logInfo "building h expression" + let hExp ← withLocalDeclD `h e fun h => do + mkLambdaFVars #[h] <| mkApp h nInst + logInfo m!"Calling imp; hExpr of type {← inferType hExp}" + imp h hExp finalR (PSum.inr <| fun x _ => x) instance (priority := 2000) subtypeVarTestable {p : α → Prop} {β : α → Prop} [∀ x, PrintableProp (p x)] @@ -710,11 +695,30 @@ def MetaTestable.check (p : Prop) (cfg : Configuration := {}) | MetaTestResult.failure _ pf xs n => let msg := "Found a counter-example!" if cfg.quiet then - Lean.logWarning msg + Lean.logInfo msg else - Lean.logWarning <| Testable.formatFailure msg xs n + Lean.logInfo <| Testable.formatFailure msg xs n return some pf +def disproveM? (cfg : Configuration) (tgt: Expr) : MetaM <| Option Expr := do + let tgt' ← Decorations.addDecorations tgt + let inst ← try + synthInstance (← mkAppM ``MetaTestable #[tgt']) + catch _ => + throwError "Failed to create a `testable` instance for `{tgt}`." + let e ← mkAppOptM ``MetaTestable.check #[tgt, toExpr cfg, tgt', inst] + let expectedType := Lean.Expr.forallE `a + (Lean.Expr.const `Lean.Expr []) + (Lean.Expr.app + (Lean.Expr.const `Lean.Meta.MetaM []) + (Lean.Expr.app + (Lean.Expr.const ``Option [Level.zero]) + (Lean.Expr.const ``Lean.Expr []))) + (Lean.BinderInfo.default) + let code ← unsafe evalExpr (Expr → MetaM (Option Expr)) expectedType e + code tgt + + -- #eval MetaTestable.check (∀ (x y z a : Nat) (h1 : 3 < x) (h2 : 3 < y), x - y = y - x) -- Configuration.verbose -- #eval MetaTestable.check (∀ x : Nat, ∀ y : Nat, x + y = y + x) Configuration.verbose diff --git a/Test/MetaTestable.lean b/Test/MetaTestable.lean index 5b82b65..81f26ce 100644 --- a/Test/MetaTestable.lean +++ b/Test/MetaTestable.lean @@ -140,7 +140,7 @@ def de : Decorations.DecorationsOf (∀ a b : Nat, a ≤ b) := by mk_decorations #synth MetaTestable (NamedBinder "a" (∀ (a : Nat), NamedBinder "b" (∀ (b : Nat), a ≤ b))) - +set_option pp.universes true in #eval MetaTestable.check (∀ (a : Nat), False) (propExpr := Lean.Expr.forallE `a (Lean.Expr.const `Nat []) (Lean.Expr.const `False []) (Lean.BinderInfo.default)) /- @@ -169,7 +169,7 @@ set_option pp.universes true in (Lean.BinderInfo.default)) (Lean.BinderInfo.default))) - +set_option linter.unusedVariables false in #eval MetaTestable.check (∀ (a b : Nat), a < 0) (propExpr := (Lean.Expr.forallE `a (Lean.Expr.const `Nat []) @@ -192,3 +192,15 @@ set_option pp.universes true in #expr ∀ (a b : Nat), a < 0 #check Lean.Expr.lit (Lean.Literal.natVal 0) + +#expr Expr → MetaM (Option Expr) + + +elab "disprove%" t:term : term => do + let tgt ← elabType t + let cfg : Configuration := {} + let (some code') ← disproveM? cfg tgt | throwError "disprove% failed" + logInfo s!"disproof: {← ppExpr code'}; \ntype: {← ppExpr <| (← inferType code')}" + return tgt + +#check disprove% ∀ (a b : Nat), a < b From 4920676a0cfa3ae50bf12a498fe9a3195d180d87 Mon Sep 17 00:00:00 2001 From: Siddhartha Gadgil Date: Fri, 22 Nov 2024 15:34:16 +0530 Subject: [PATCH 14/30] fixed `and` an `or`, commented out noise --- Plausible/MetaTestable.lean | 44 +++++++++++++++++++------------------ Test/MetaTestable.lean | 4 ++++ 2 files changed, 27 insertions(+), 21 deletions(-) diff --git a/Plausible/MetaTestable.lean b/Plausible/MetaTestable.lean index 3694b44..ef10a1a 100644 --- a/Plausible/MetaTestable.lean +++ b/Plausible/MetaTestable.lean @@ -327,11 +327,11 @@ can be informative. -/ def addInfo (x : String) (h : q → p) (hExp: Expr) (r : MetaTestResult p) (p : Unit ⊕' (p → q) := PSum.inl ()) : (MetaM <| MetaTestResult q) := do if let failure h2 pf xs n := r then - logInfo s!"Adding info: {x} := {repr x}; type {← inferType hExp}" - logInfo s!"proof type: {← ppExpr <| ← inferType pf}" - logInfo s!"hExp type: {← ppExpr <| ← inferType hExp}" + -- logInfo s!"Adding info: {x} := {repr x}; type {← inferType hExp}" + -- logInfo s!"proof type: {← ppExpr <| ← inferType pf}" + -- logInfo s!"hExp type: {← ppExpr <| ← inferType hExp}" let pf' ← mkAppM ``mt #[hExp, pf] - logInfo s!"Proof: {pf'}" + -- logInfo s!"Proof: {pf'}" return failure (mt h h2) pf' (x :: xs) n else imp h hExp r p @@ -339,7 +339,7 @@ def addInfo (x : String) (h : q → p) (hExp: Expr) (r : MetaTestResult p) /-- Add some formatting to the information recorded by `addInfo`. -/ def addVarInfo {γ : Type _} [Repr γ] (var : String) (x : γ) (h : q → p) (hExp: Expr) (r : MetaTestResult p) (p : Unit ⊕' (p → q) := PSum.inl ()) : MetaM (MetaTestResult q) := do - logInfo "Adding var info" + -- logInfo "Adding var info" addInfo s!"{var} := {repr x}" h (hExp: Expr) r p def isFailure : MetaTestResult p → Bool @@ -357,7 +357,7 @@ abbrev ProxyExpr α [SampleableExt α] := ToExpr (SampleableExt.proxy α) def runProp (p : Prop) [MetaTestable p] : Configuration → Bool → Expr → MGen (MetaTestResult p) := fun cfg b e => do - logInfo s!"Testing {e}" + -- logInfo s!"Testing {e}" MetaTestable.run cfg b e /-- A `dbgTrace` with special formatting -/ @@ -366,20 +366,22 @@ def slimTrace {m : Type → Type _} [Pure m] (s : String) : m PUnit := instance andTestable [MetaTestable p] [MetaTestable q] : MetaTestable (p ∧ q) where run := fun cfg min e => do - let xp ← runProp p cfg min e - let xq ← runProp q cfg min e + let (some e₁, some e₂) ← andProp? e | throwError m!"Expected an `And` proposition, got {← ppExpr e}" + let xp ← runProp p cfg min e₁ + let xq ← runProp q cfg min e₂ and xp xq e instance orTestable [MetaTestable p] [MetaTestable q] : MetaTestable (p ∨ q) where run := fun cfg min e => do - let xp ← runProp p cfg min e + let (some e₁, some e₂) ← orProp? e | throwError m!"Expected an `Or` proposition, got {← ppExpr e}" + let xp ← runProp p cfg min e₁ -- As a little performance optimization we can just not run the second -- test if the first succeeds match xp with | success (PSum.inl h) => return success (PSum.inl h) | success (PSum.inr h) => return success (PSum.inr <| Or.inl h) | _ => - let xq ← runProp q cfg min e + let xq ← runProp q cfg min e₂ or xp xq e theorem iff_resolve (p q : Prop) : (p ↔ q) ↔ p ∧ q ∨ ¬p ∧ ¬q := by @@ -409,7 +411,7 @@ instance decGuardTestable [PrintableProp p] [Decidable p] {β : p → Prop} [∀ let decInstType ← mkAppM ``Decidable #[pExp] let inst ← synthInstance decInstType let falseRefl ← mkAppM ``Eq.refl #[mkConst ``false] - logInfo "Decidable used; building proof" + -- logInfo "Decidable used; building proof" let pf ← mkAppOptM ``of_decide_eq_true #[pExp, inst, falseRefl] let cod := mkApp βExp pf let hExp ← withLocalDeclD `x (← mkArrow pExp cod) fun x => do @@ -472,7 +474,7 @@ partial def minimizeAux [SampleableExt α] [ProxyExpr α] {β : α → Prop} [ let samp ← synthInstance instType let xInterp ← mkAppOptM ``SampleableExt.interp #[αExp, samp, xExpr] let e' ← mkAppM' βExpr #[xInterp] - logInfo "Minimize" + -- logInfo "Minimize" let res ← OptionT.lift <| MetaTestable.runProp (β (SampleableExt.interp candidate)) cfg true e' if res.isFailure then if cfg.traceShrink then @@ -515,7 +517,7 @@ instance varTestable [SampleableExt α] [ProxyExpr α] {β : α → Prop} [∀ x let (e', _) ← dsimp e' {} if cfg.traceSuccesses || cfg.traceDiscarded then slimTrace s!"{var} := {repr x}" - logInfo s!"variable testable: outer prop: {e}, inner prop: {e'}" + -- logInfo s!"variable testable: outer prop: {e}, inner prop: {e'}" let r ← MetaTestable.runProp (β <| SampleableExt.interp x) cfg false e' let ⟨finalX, finalR⟩ ← if isFailure r then @@ -534,7 +536,7 @@ instance varTestable [SampleableExt α] [ProxyExpr α] {β : α → Prop} [∀ x let samp ← synthInstance instType let xInterp ← mkAppOptM ``SampleableExt.interp #[αExp, samp, xExpr] let h := (· <| SampleableExt.interp finalX) - logInfo s!"Adding var info from varTestable, goal : {← ppExpr e}" + -- logInfo s!"Adding var info from varTestable, goal : {← ppExpr e}" let hExpr ← withLocalDeclD `x e fun x => do mkLambdaFVars #[x] (mkApp x xInterp) addVarInfo var finalX h hExpr finalR @@ -571,7 +573,7 @@ instance propVarTestable {β : Prop → Prop} [h: ∀ b : Bool, MetaTestable (β where run := fun cfg min e => do let (some βExpr, _) ← forallProp? e | throwError m!"Expected a `Forall` proposition, got {← ppExpr e}" - logInfo s!"prop variable testable: outer prop: {e}, inner prop: {βExpr}" + -- logInfo s!"prop variable testable: outer prop: {e}, inner prop: {βExpr}" let p ← MetaTestable.runProp (NamedBinder var <| ∀ b : Bool, β b) cfg min e let e' ← mkAppM ``bool_to_prop_fmly #[βExpr] imp (bool_to_prop_fmly β) e' p @@ -582,17 +584,17 @@ where run := fun cfg min e => do if cfg.traceDiscarded || cfg.traceSuccesses then slimTrace s!"{var} is unused" - logInfo s!"unused variable testable: outer prop: {e}" + -- logInfo s!"unused variable testable: outer prop: {e}" let (some aExp, some e') ← impProp? e | throwError m!"Expected an `Imp` proposition, got {← ppExpr e}" let r ← MetaTestable.runProp β cfg min e' let hExp ← mkAppOptM ``id #[e'] let finalR ← addInfo s!"{var} is irrelevant (unused)" id hExp r let h := (· <| Classical.ofNonempty) let nInst ← synthInstance <| ← mkAppM ``Nonempty #[aExp] - logInfo "building h expression" + -- logInfo "building h expression" let hExp ← withLocalDeclD `h e fun h => do mkLambdaFVars #[h] <| mkApp h nInst - logInfo m!"Calling imp; hExpr of type {← inferType hExp}" + -- logInfo m!"Calling imp; hExpr of type {← inferType hExp}" imp h hExp finalR (PSum.inr <| fun x _ => x) instance (priority := 2000) subtypeVarTestable {p : α → Prop} {β : α → Prop} @@ -619,11 +621,11 @@ instance (priority := low) decidableTestable {p : Prop} [PrintableProp p] [Decid let s := printProp p let inst ← synthInstance <| ← mkAppM ``Decidable #[e] let falseRefl ← mkAppM ``Eq.refl #[mkConst ``false] - logInfo "Decidable used; building proof" + -- logInfo "Decidable used; building proof" let pf' ← mkAppOptM ``of_decide_eq_false #[e, inst, falseRefl] - logInfo "Decidable proof built" + -- logInfo "Decidable proof built" checkDisproof pf' e - logInfo "Decidable proof checked" + -- logInfo "Decidable proof checked" return failure h pf' [s!"issue: {s} does not hold"] 0 end MetaTestable diff --git a/Test/MetaTestable.lean b/Test/MetaTestable.lean index 81f26ce..f379a71 100644 --- a/Test/MetaTestable.lean +++ b/Test/MetaTestable.lean @@ -204,3 +204,7 @@ elab "disprove%" t:term : term => do return tgt #check disprove% ∀ (a b : Nat), a < b + +#check disprove% ∀ (a b : Nat), a < b ∨ b < a + +#check disprove% False ∧ False From dbba754459474cbdd8990034308823cbc5637258 Mon Sep 17 00:00:00 2001 From: Siddhartha Gadgil Date: Fri, 22 Nov 2024 16:02:49 +0530 Subject: [PATCH 15/30] purged `-- logInfo` --- Plausible/MetaTestable.lean | 18 ------------------ 1 file changed, 18 deletions(-) diff --git a/Plausible/MetaTestable.lean b/Plausible/MetaTestable.lean index ef10a1a..67f7056 100644 --- a/Plausible/MetaTestable.lean +++ b/Plausible/MetaTestable.lean @@ -327,11 +327,7 @@ can be informative. -/ def addInfo (x : String) (h : q → p) (hExp: Expr) (r : MetaTestResult p) (p : Unit ⊕' (p → q) := PSum.inl ()) : (MetaM <| MetaTestResult q) := do if let failure h2 pf xs n := r then - -- logInfo s!"Adding info: {x} := {repr x}; type {← inferType hExp}" - -- logInfo s!"proof type: {← ppExpr <| ← inferType pf}" - -- logInfo s!"hExp type: {← ppExpr <| ← inferType hExp}" let pf' ← mkAppM ``mt #[hExp, pf] - -- logInfo s!"Proof: {pf'}" return failure (mt h h2) pf' (x :: xs) n else imp h hExp r p @@ -339,7 +335,6 @@ def addInfo (x : String) (h : q → p) (hExp: Expr) (r : MetaTestResult p) /-- Add some formatting to the information recorded by `addInfo`. -/ def addVarInfo {γ : Type _} [Repr γ] (var : String) (x : γ) (h : q → p) (hExp: Expr) (r : MetaTestResult p) (p : Unit ⊕' (p → q) := PSum.inl ()) : MetaM (MetaTestResult q) := do - -- logInfo "Adding var info" addInfo s!"{var} := {repr x}" h (hExp: Expr) r p def isFailure : MetaTestResult p → Bool @@ -357,7 +352,6 @@ abbrev ProxyExpr α [SampleableExt α] := ToExpr (SampleableExt.proxy α) def runProp (p : Prop) [MetaTestable p] : Configuration → Bool → Expr → MGen (MetaTestResult p) := fun cfg b e => do - -- logInfo s!"Testing {e}" MetaTestable.run cfg b e /-- A `dbgTrace` with special formatting -/ @@ -411,7 +405,6 @@ instance decGuardTestable [PrintableProp p] [Decidable p] {β : p → Prop} [∀ let decInstType ← mkAppM ``Decidable #[pExp] let inst ← synthInstance decInstType let falseRefl ← mkAppM ``Eq.refl #[mkConst ``false] - -- logInfo "Decidable used; building proof" let pf ← mkAppOptM ``of_decide_eq_true #[pExp, inst, falseRefl] let cod := mkApp βExp pf let hExp ← withLocalDeclD `x (← mkArrow pExp cod) fun x => do @@ -474,7 +467,6 @@ partial def minimizeAux [SampleableExt α] [ProxyExpr α] {β : α → Prop} [ let samp ← synthInstance instType let xInterp ← mkAppOptM ``SampleableExt.interp #[αExp, samp, xExpr] let e' ← mkAppM' βExpr #[xInterp] - -- logInfo "Minimize" let res ← OptionT.lift <| MetaTestable.runProp (β (SampleableExt.interp candidate)) cfg true e' if res.isFailure then if cfg.traceShrink then @@ -511,13 +503,11 @@ instance varTestable [SampleableExt α] [ProxyExpr α] {β : α → Prop} [∀ x let v ← mkFreshLevelMVar let instType := mkApp (mkConst ``SampleableExt [u, v]) αExp let samp ← synthInstance instType - -- logInfo s!"Instance: {samp}" let xInterp ← mkAppOptM ``SampleableExt.interp #[αExp, samp, xExpr] let e' ← mkAppM' βExp #[xInterp] let (e', _) ← dsimp e' {} if cfg.traceSuccesses || cfg.traceDiscarded then slimTrace s!"{var} := {repr x}" - -- logInfo s!"variable testable: outer prop: {e}, inner prop: {e'}" let r ← MetaTestable.runProp (β <| SampleableExt.interp x) cfg false e' let ⟨finalX, finalR⟩ ← if isFailure r then @@ -536,7 +526,6 @@ instance varTestable [SampleableExt α] [ProxyExpr α] {β : α → Prop} [∀ x let samp ← synthInstance instType let xInterp ← mkAppOptM ``SampleableExt.interp #[αExp, samp, xExpr] let h := (· <| SampleableExt.interp finalX) - -- logInfo s!"Adding var info from varTestable, goal : {← ppExpr e}" let hExpr ← withLocalDeclD `x e fun x => do mkLambdaFVars #[x] (mkApp x xInterp) addVarInfo var finalX h hExpr finalR @@ -573,7 +562,6 @@ instance propVarTestable {β : Prop → Prop} [h: ∀ b : Bool, MetaTestable (β where run := fun cfg min e => do let (some βExpr, _) ← forallProp? e | throwError m!"Expected a `Forall` proposition, got {← ppExpr e}" - -- logInfo s!"prop variable testable: outer prop: {e}, inner prop: {βExpr}" let p ← MetaTestable.runProp (NamedBinder var <| ∀ b : Bool, β b) cfg min e let e' ← mkAppM ``bool_to_prop_fmly #[βExpr] imp (bool_to_prop_fmly β) e' p @@ -584,17 +572,14 @@ where run := fun cfg min e => do if cfg.traceDiscarded || cfg.traceSuccesses then slimTrace s!"{var} is unused" - -- logInfo s!"unused variable testable: outer prop: {e}" let (some aExp, some e') ← impProp? e | throwError m!"Expected an `Imp` proposition, got {← ppExpr e}" let r ← MetaTestable.runProp β cfg min e' let hExp ← mkAppOptM ``id #[e'] let finalR ← addInfo s!"{var} is irrelevant (unused)" id hExp r let h := (· <| Classical.ofNonempty) let nInst ← synthInstance <| ← mkAppM ``Nonempty #[aExp] - -- logInfo "building h expression" let hExp ← withLocalDeclD `h e fun h => do mkLambdaFVars #[h] <| mkApp h nInst - -- logInfo m!"Calling imp; hExpr of type {← inferType hExp}" imp h hExp finalR (PSum.inr <| fun x _ => x) instance (priority := 2000) subtypeVarTestable {p : α → Prop} {β : α → Prop} @@ -621,11 +606,8 @@ instance (priority := low) decidableTestable {p : Prop} [PrintableProp p] [Decid let s := printProp p let inst ← synthInstance <| ← mkAppM ``Decidable #[e] let falseRefl ← mkAppM ``Eq.refl #[mkConst ``false] - -- logInfo "Decidable used; building proof" let pf' ← mkAppOptM ``of_decide_eq_false #[e, inst, falseRefl] - -- logInfo "Decidable proof built" checkDisproof pf' e - -- logInfo "Decidable proof checked" return failure h pf' [s!"issue: {s} does not hold"] 0 end MetaTestable From df56c20cda9eaeda8784332114ad7372163861d5 Mon Sep 17 00:00:00 2001 From: Siddhartha Gadgil Date: Fri, 22 Nov 2024 18:53:18 +0530 Subject: [PATCH 16/30] fixed subtype --- Plausible/MetaTestable.lean | 44 +++++++++++++++++++++++++++++++------ 1 file changed, 37 insertions(+), 7 deletions(-) diff --git a/Plausible/MetaTestable.lean b/Plausible/MetaTestable.lean index 67f7056..eb271dc 100644 --- a/Plausible/MetaTestable.lean +++ b/Plausible/MetaTestable.lean @@ -126,6 +126,24 @@ def forallProp? (e: Expr) : MetaM (Option Expr × Option Expr) := do else return (none, none) +def forallPropProp? (e: Expr) : MetaM (Option Expr × Option Expr × Option Expr) := do + let u ← mkFreshLevelMVar + let α ← mkFreshExprMVar (mkSort u) + let prop := mkSort levelZero + let fmlyType ← mkArrow α prop + let fmly ← mkFreshExprMVar (some fmlyType) + let p ← mkFreshExprMVar (some fmlyType) + let e' ← withLocalDeclD `x α fun x => do + let y ← mkAppM' fmly #[x] + let y' ← mkAppM' p #[x] + let cod ← mkArrow y' y + mkForallFVars #[x] cod + if ← isDefEq e' e then + return (← Lean.getExprMVarAssignment? p.mvarId!, ← Lean.getExprMVarAssignment? fmly.mvarId!, ← Lean.getExprMVarAssignment? α.mvarId!) + else + return (none, none, none) + + def impProp? (e: Expr) : MetaM (Option Expr × Option Expr) := do let u ← mkFreshLevelMVar let α ← mkFreshExprMVar (mkSort u) @@ -582,20 +600,31 @@ where mkLambdaFVars #[h] <| mkApp h nInst imp h hExp finalR (PSum.inr <| fun x _ => x) +theorem prop_iff_subtype (p : α → Prop) (β : α → Prop) : NamedBinder var (∀ (x : α), NamedBinder var' (p x → β x)) ↔ ∀ (x : Subtype p), β x.val := by simp [Subtype.forall, NamedBinder] + instance (priority := 2000) subtypeVarTestable {p : α → Prop} {β : α → Prop} [∀ x, PrintableProp (p x)] - [∀ x, MetaTestable (β x)] + [∀ x, MetaTestable (β x)][ToExpr α] [SampleableExt (Subtype p)] [ProxyExpr (Subtype p)] {var'} : MetaTestable (NamedBinder var <| (x : α) → NamedBinder var' <| p x → β x) where - run cfg min e := + run cfg min e := do + let (some pExp, some βExp, some _) ← forallPropProp? e | throwError m!"Expected a `Forall` proposition with arrow, got {← ppExpr e}" + let subType ← mkAppM ``Subtype #[pExp] letI (x : Subtype p) : MetaTestable (β x) := { run := fun cfg min e => do - let r ← MetaTestable.runProp (β x.val) cfg min e - addInfo s!"guard: {printProp (p x)} (by construction)" id e r (PSum.inr id) } + let xExp := toExpr x.val + let y ← mkAppM' βExp #[xExp] + let r ← MetaTestable.runProp (β x.val) cfg min y + let idExp ← mkAppOptM ``id #[e] + addInfo s!"guard: {printProp (p x)} (by construction)" id idExp r (PSum.inr id) } do - let r ← @MetaTestable.run (∀ x : Subtype p, β x.val) (@varTestable var _ _ _ _ _) cfg min e - have := by simp [Subtype.forall, NamedBinder] - iff this e r + let e' ← withLocalDeclD `x subType fun x => do + let x' ← mkAppM ``Subtype.val #[x] + let y ← mkAppM' βExp #[x'] + mkForallFVars #[x] y + let r ← @MetaTestable.run (∀ x : Subtype p, β x.val) (@varTestable var _ _ _ _ _) cfg min e' + let hExp ← mkAppM ``prop_iff_subtype #[pExp, βExp] + iff (prop_iff_subtype p β) hExp r instance (priority := low) decidableTestable {p : Prop} [PrintableProp p] [Decidable p] : MetaTestable p where @@ -608,6 +637,7 @@ instance (priority := low) decidableTestable {p : Prop} [PrintableProp p] [Decid let falseRefl ← mkAppM ``Eq.refl #[mkConst ``false] let pf' ← mkAppOptM ``of_decide_eq_false #[e, inst, falseRefl] checkDisproof pf' e + -- logInfo "Decidable proof checked" return failure h pf' [s!"issue: {s} does not hold"] 0 end MetaTestable From 8e01cc2467f3c9d8c71de6f914343aecd548a61b Mon Sep 17 00:00:00 2001 From: Siddhartha Gadgil Date: Sat, 23 Nov 2024 19:23:40 +0530 Subject: [PATCH 17/30] Vacuous tactic --- Plausible/Vacuous.lean | 69 ++++++++++++++++++++++++++++++++++++++++++ Test/MetaTestable.lean | 17 +++++++++-- 2 files changed, 84 insertions(+), 2 deletions(-) create mode 100644 Plausible/Vacuous.lean diff --git a/Plausible/Vacuous.lean b/Plausible/Vacuous.lean new file mode 100644 index 0000000..b109b3f --- /dev/null +++ b/Plausible/Vacuous.lean @@ -0,0 +1,69 @@ +import Lean +import Plausible +import Plausible.MetaTestable + +open Lean Meta Elab Tactic Plausible + +partial def proveVacuous? (p: Expr) (cfg : Configuration := {}) : MetaM <| Option Expr := do + match p with + | .forallE n d b bi => + withLocalDecl n bi d fun x => do + let b := b.instantiate1 x + let contra? ← try + disproveM? cfg d -- proof of ¬d + catch _ => + pure none + match contra? with + | some contra => + logInfo m!"Vacuous Implication. Hypothesis {← ppExpr d} is never satisfied" + let pfFalse ← mkAppM' contra #[x] + let pfBody ← mkAppOptM ``False.elim #[b, pfFalse] + mkLambdaFVars #[x] pfBody + | none => + let inner ← proveVacuous? b cfg + inner.mapM fun pf => mkLambdaFVars #[x] pf + | _ => + return none + +open Lean.Parser.Tactic + +syntax (name := vacuousSyntax) "vacuous" (config)? : tactic +elab_rules : tactic | `(tactic| vacuous $[$cfg]?) => withMainContext do + let cfg ← elabConfig (mkOptionalNode cfg) let (_, g) ← (← getMainGoal).revert ((← getLocalHyps).map (Expr.fvarId!)) + let cfg := { cfg with + traceDiscarded := cfg.traceDiscarded || (← isTracingEnabledFor `plausible.discarded), + traceSuccesses := cfg.traceSuccesses || (← isTracingEnabledFor `plausible.success), + traceShrink := cfg.traceShrink || (← isTracingEnabledFor `plausible.shrink.steps), + traceShrinkCandidates := cfg.traceShrinkCandidates + || (← isTracingEnabledFor `plausible.shrink.candidates) } + g.withContext do + let tgt ← g.getType + match ← proveVacuous? tgt cfg with + | some pf => + g.assign pf + | none => + throwError "no vacuous proof found" + + +elab "prove_vacuous" type:term : term => do + let p ← Term.elabType type + let vacuous ← proveVacuous? p + match vacuous with + | some pf => + return pf + | none => + logWarning m!"No vacuous proof found" + return mkConst ``False + + +#check prove_vacuous ((2 : Nat) < 1) → 1 ≤ 3 + +#check prove_vacuous (∀ n: Nat, n < (4: Nat)) → 1 ≤ 3 + +example : (∀ n: Nat, n < (4: Nat)) → 1 ≤ 3 := by vacuous + +example : (2 : Nat) < 1 → 1 ≤ 3 := by vacuous + +example (h: ∀ n: Nat, n < (4: Nat)) : 4 ≤ 3 := by vacuous + +example (m: Nat) (h: ∀ n: Nat, n < (4: Nat)) : m ≤ 3 := by vacuous diff --git a/Test/MetaTestable.lean b/Test/MetaTestable.lean index f379a71..41c0fa3 100644 --- a/Test/MetaTestable.lean +++ b/Test/MetaTestable.lean @@ -141,7 +141,7 @@ def de : Decorations.DecorationsOf (∀ a b : Nat, a ≤ b) := by mk_decorations #synth MetaTestable (NamedBinder "a" (∀ (a : Nat), NamedBinder "b" (∀ (b : Nat), a ≤ b))) set_option pp.universes true in -#eval MetaTestable.check (∀ (a : Nat), False) (propExpr := Lean.Expr.forallE `a (Lean.Expr.const `Nat []) (Lean.Expr.const `False []) (Lean.BinderInfo.default)) +#eval MetaTestable.check (∀ (_a : Nat), False) (propExpr := Lean.Expr.forallE `a (Lean.Expr.const `Nat []) (Lean.Expr.const `False []) (Lean.BinderInfo.default)) /- AppBuilder for 'mkAppM', result contains metavariables @@ -189,7 +189,7 @@ set_option linter.unusedVariables false in (Lean.BinderInfo.default))) #expr ∀ (a b : Nat), a < b -#expr ∀ (a b : Nat), a < 0 +#expr ∀ (a _b : Nat), a < 0 #check Lean.Expr.lit (Lean.Literal.natVal 0) @@ -208,3 +208,16 @@ elab "disprove%" t:term : term => do #check disprove% ∀ (a b : Nat), a < b ∨ b < a #check disprove% False ∧ False + +def eg {p: Prop}(h : ¬¬p) : p := + by + by_cases h':p + · exact h' + · exfalso + contradiction + +/- +def eg : ∀ {p : Prop}, ¬¬p → p := +fun {p} h => if h' : p then h' else False.elim (absurd h' h) +-/ +#print eg From 7ea09449ba8769d484582ae0dfe73ee95674e8b3 Mon Sep 17 00:00:00 2001 From: Siddhartha Gadgil Date: Sat, 23 Nov 2024 19:44:27 +0530 Subject: [PATCH 18/30] separated tests --- Plausible/Vacuous.lean | 24 ----------- Test/MetaTestable.lean | 51 ----------------------- Test/Vacuous.lean | 91 ++++++++++++++++++++++++++++++++++++++++++ 3 files changed, 91 insertions(+), 75 deletions(-) create mode 100644 Test/Vacuous.lean diff --git a/Plausible/Vacuous.lean b/Plausible/Vacuous.lean index b109b3f..a0660cb 100644 --- a/Plausible/Vacuous.lean +++ b/Plausible/Vacuous.lean @@ -43,27 +43,3 @@ elab_rules : tactic | `(tactic| vacuous $[$cfg]?) => withMainContext do g.assign pf | none => throwError "no vacuous proof found" - - -elab "prove_vacuous" type:term : term => do - let p ← Term.elabType type - let vacuous ← proveVacuous? p - match vacuous with - | some pf => - return pf - | none => - logWarning m!"No vacuous proof found" - return mkConst ``False - - -#check prove_vacuous ((2 : Nat) < 1) → 1 ≤ 3 - -#check prove_vacuous (∀ n: Nat, n < (4: Nat)) → 1 ≤ 3 - -example : (∀ n: Nat, n < (4: Nat)) → 1 ≤ 3 := by vacuous - -example : (2 : Nat) < 1 → 1 ≤ 3 := by vacuous - -example (h: ∀ n: Nat, n < (4: Nat)) : 4 ≤ 3 := by vacuous - -example (m: Nat) (h: ∀ n: Nat, n < (4: Nat)) : m ≤ 3 := by vacuous diff --git a/Test/MetaTestable.lean b/Test/MetaTestable.lean index 41c0fa3..68185da 100644 --- a/Test/MetaTestable.lean +++ b/Test/MetaTestable.lean @@ -63,9 +63,6 @@ info: Implies: 1 = 0; 2 ≠ 0 #guard_msgs in #decompose_prop 1 = 0 ↔ 2 ≠ 0 -#check MetaTestResult.failure - - def eg_fail : MetaTestResult False := @MetaTestResult.failure False (fun (x: False) ↦ x) (Lean.Expr.lam `x (Lean.Expr.const `False []) (Lean.Expr.bvar 0) (Lean.BinderInfo.default)) [] 0 @@ -95,47 +92,12 @@ elab "expr%" e:term : term => do logInfo s!"{repr e}" logInfo s!"{← reduce e}" return e --- - -structure MyType where - x : Nat - y : Nat - h : x ≤ y - deriving Repr - -#check fun (x y : Nat) (h : x ≤ y) => expr% (MyType.mk x y h) - -instance : Shrinkable MyType where - shrink := fun ⟨x, y, _⟩ => - let proxy := Shrinkable.shrink (x, y - x) - proxy.map (fun (fst, snd) => ⟨fst, fst + snd, by omega⟩) - -instance : SampleableExt MyType := - SampleableExt.mkSelfContained do - let x ← SampleableExt.interpSample Nat - let xyDiff ← SampleableExt.interpSample Nat - return ⟨x, x + xyDiff, by omega⟩ - -set_option diagnostics true in --- #eval MetaTestable.check <| ∀ a b : MyType, a.y ≤ b.x → a.x ≤ b.y open Decorations Lean Elab Tactic -elab "mk_decorations" : tactic => do - let goal ← getMainGoal - let goalType ← goal.getType - if let .app (.const ``Decorations.DecorationsOf _) body := goalType then - closeMainGoal `mk_decorations (← Decorations.addDecorations body) - - -def de : Decorations.DecorationsOf (∀ a b : Nat, a ≤ b) := by mk_decorations - -#print de #synth MetaTestable <| (1: Nat) = 0 -#synth Testable (NamedBinder "a" (∀ (a : Nat), a ≤ 1)) - #synth MetaTestable (NamedBinder "a" (∀ (a : Nat), a ≤ 1)) #synth MetaTestable (NamedBinder "a" (∀ (a : Nat), NamedBinder "b" (∀ (b : Nat), a ≤ b))) @@ -208,16 +170,3 @@ elab "disprove%" t:term : term => do #check disprove% ∀ (a b : Nat), a < b ∨ b < a #check disprove% False ∧ False - -def eg {p: Prop}(h : ¬¬p) : p := - by - by_cases h':p - · exact h' - · exfalso - contradiction - -/- -def eg : ∀ {p : Prop}, ¬¬p → p := -fun {p} h => if h' : p then h' else False.elim (absurd h' h) --/ -#print eg diff --git a/Test/Vacuous.lean b/Test/Vacuous.lean new file mode 100644 index 0000000..c928f02 --- /dev/null +++ b/Test/Vacuous.lean @@ -0,0 +1,91 @@ +import Plausible.Vacuous +open Lean Meta Elab + +elab "prove_vacuous" type:term : term => do + let p ← Term.elabType type + let vacuous ← proveVacuous? p + match vacuous with + | some pf => + return pf + | none => + logWarning m!"No vacuous proof found" + return mkConst ``False + +/-- +info: +=================== +Found a counter-example! +issue: 2 < 1 does not hold +(0 shrinks) +------------------- +--- +info: Vacuous Implication. Hypothesis 2 < 1 is never satisfied +--- +info: fun a => False.elim (of_decide_eq_false (Eq.refl false) a) : 2 < 1 → 1 ≤ 3 +-/ +#guard_msgs in +#check prove_vacuous ((2 : Nat) < 1) → 1 ≤ 3 + +/-- +info: +=================== +Found a counter-example! +n := 6 +issue: 6 < 4 does not hold +(0 shrinks) +------------------- +--- +info: Vacuous Implication. Hypothesis ∀ (n : Nat), n < 4 is never satisfied +--- +info: fun a => + False.elim + (mt (fun x => x (Plausible.SampleableExt.interp 6)) (of_decide_eq_false (Eq.refl false)) + a) : (∀ (n : Nat), n < 4) → 1 ≤ 3 +-/ +#guard_msgs in +#check prove_vacuous (∀ n: Nat, n < (4: Nat)) → 1 ≤ 3 + +/-- +info: +=================== +Found a counter-example! +n := 6 +issue: 6 < 4 does not hold +(0 shrinks) +------------------- +--- +info: Vacuous Implication. Hypothesis ∀ (n : Nat), n < 4 is never satisfied +-/ +#guard_msgs in +-- Can negate hypothesis in goal statement +example : (∀ n: Nat, n < (4: Nat)) → 4 ≤ 3 := by vacuous + +/-- +info: +=================== +Found a counter-example! +n := 6 +issue: 6 < 4 does not hold +(0 shrinks) +------------------- +--- +info: Vacuous Implication. Hypothesis ∀ (n : Nat), n < 4 is never satisfied +-/ +#guard_msgs in +-- Can negate hypothesis in local context +example (h: ∀ n: Nat, n < (4: Nat)) : 4 ≤ 3 := by vacuous + +/-- +info: +=================== +Found a counter-example! +n := 6 +issue: 6 < 4 does not hold +(0 shrinks) +------------------- +--- +info: Vacuous Implication. Hypothesis ∀ (n : Nat), n < 4 is never satisfied +-/ +#guard_msgs in +-- Can negate hypothesis other than the first one +example (m: Nat) (h: ∀ n: Nat, n < (4: Nat)) : m ≤ 3 := by vacuous From 6897196fc56f5226a14d244cd83bc76ce6aab661 Mon Sep 17 00:00:00 2001 From: Siddhartha Gadgil Date: Sat, 23 Nov 2024 21:22:42 +0530 Subject: [PATCH 19/30] clean up of tests --- Plausible/MetaTestable.lean | 36 -- Test/MetaTestable.lean | 737 +++++++++++++++++++++++++++++++++--- Test/Vacuous.lean | 18 +- 3 files changed, 698 insertions(+), 93 deletions(-) diff --git a/Plausible/MetaTestable.lean b/Plausible/MetaTestable.lean index eb271dc..8262fd2 100644 --- a/Plausible/MetaTestable.lean +++ b/Plausible/MetaTestable.lean @@ -29,47 +29,12 @@ counter-example it will then use a `Shrinkable` instance to reduce the example. This allows the user to create new instances and apply `Plausible` to new situations. -### What do I do if I'm testing a property about my newly defined type? - -Let us consider a type made for a new formalization: - -```lean -structure MyType where - x : Nat - y : Nat - h : x ≤ y - deriving Repr -``` - -How do we test a property about `MyType`? For instance, let us consider -`MetaTestable.check <| ∀ a b : MyType, a.y ≤ b.x → a.x ≤ b.y`. Writing this -property as is will give us an error because we do not have an instance -of `Shrinkable MyType` and `SampleableExt MyType`. We can define one as follows: - -```lean -instance : Shrinkable MyType where - shrink := fun ⟨x, y, _⟩ => - let proxy := Shrinkable.shrink (x, y - x) - proxy.map (fun (fst, snd) => ⟨fst, fst + snd, by omega⟩) - -instance : SampleableExt MyType := - SampleableExt.mkSelfContained do - let x ← SampleableExt.interpSample Nat - let xyDiff ← SampleableExt.interpSample Nat - return ⟨x, x + xyDiff, by omega⟩ -``` - -Again, we take advantage of the fact that other types have useful -`Shrinkable` implementations, in this case `Prod`. ## Main definitions * `MetaTestable` class * `MetaTestable.check`: a way to test a proposition using random examples -## References - -* https://hackage.haskell.org/package/QuickCheck -/ @@ -328,7 +293,6 @@ def imp (h : q → p) (hExp: Expr) (r : MetaTestResult p) (p : Unit ⊕' (p → q) := PSum.inl ()) : MetaM (MetaTestResult q) := match r with | failure h2 pf xs n => do - logInfo s!"Implication: {hExp}; type {← inferType hExp}" let pf' ← mkAppM ``mt #[hExp, pf] return failure (mt h h2) pf' xs n | success h2 => return success <| combine p h2 diff --git a/Test/MetaTestable.lean b/Test/MetaTestable.lean index 68185da..cfb5008 100644 --- a/Test/MetaTestable.lean +++ b/Test/MetaTestable.lean @@ -4,6 +4,7 @@ import Plausible.MetaTestable open Plausible Plausible.MetaTestable open Lean Meta Elab Term Tactic +-- Testing the pattern matching functions open Lean Elab Term in elab "#decompose_prop" t:term : command => Command.liftTermElabM do @@ -63,23 +64,70 @@ info: Implies: 1 = 0; 2 ≠ 0 #guard_msgs in #decompose_prop 1 = 0 ↔ 2 ≠ 0 -def eg_fail : MetaTestResult False := - @MetaTestResult.failure False (fun (x: False) ↦ x) - (Lean.Expr.lam `x (Lean.Expr.const `False []) (Lean.Expr.bvar 0) (Lean.BinderInfo.default)) [] 0 +-- Elaborater to help with testing +elab "disprove%" t:term : term => do + let tgt ← elabType t + let cfg : Configuration := {} + let (some code') ← disproveM? cfg tgt | throwError "disprove% failed" + logInfo s!"disproof: {← ppExpr code'}; \ntype: {← ppExpr <| (← inferType code')}" + return tgt -def disproofExpr {p: Prop} : MetaTestResult p → MetaM Lean.Expr - | MetaTestResult.failure _ pfExpr _ _ => do - return pfExpr - | _ => - throwError "disproofExpr: expected failure" +-- Testing the disproveM? function +/-- +info: +=================== +Found a counter-example! +a := 0 +b := 0 +issue: 0 < 0 does not hold +(0 shrinks) +------------------- +--- +info: disproof: mt (fun x => x (SampleableExt.interp 0)) (mt (fun x => x (SampleableExt.interp 0)) (of_decide_eq_false (Eq.refl false))); ⏎ +type: ¬∀ (a b : Nat), a < b +--- +info: ∀ (a b : Nat), a < b : Prop +-/ +#guard_msgs in +#check disprove% ∀ (a b : Nat), a < b -elab "disproof_expr_eg%" : term => do - disproofExpr eg_fail +/-- +info: +=================== +Found a counter-example! +a := 0 +b := 0 +issue: 0 < 0 does not hold +issue: 0 < 0 does not hold +(0 shrinks) +------------------- +--- +info: disproof: mt (fun x => x (SampleableExt.interp 0)) + (mt (fun x => x (SampleableExt.interp 0)) + (Or.rec (of_decide_eq_false (Eq.refl false)) (of_decide_eq_false (Eq.refl false)))); ⏎ +type: ¬∀ (a b : Nat), a < b ∨ b < a +--- +info: ∀ (a b : Nat), a < b ∨ b < a : Prop +-/ +#guard_msgs in +#check disprove% ∀ (a b : Nat), a < b ∨ b < a -/- -fun x => x : False → False +/-- +info: +=================== +Found a counter-example! +issue: False does not hold +(0 shrinks) +------------------- +--- +info: disproof: fun h => of_decide_eq_false (Eq.refl false) h.left; ⏎ +type: False ∧ False → False +--- +info: False ∧ False : Prop -/ -#check disproof_expr_eg% +#guard_msgs in +#check disprove% False ∧ False + elab "#expr" e:term : command => Command.liftTermElabM do @@ -93,26 +141,340 @@ elab "expr%" e:term : term => do logInfo s!"{← reduce e}" return e -open Decorations Lean Elab Tactic - - -#synth MetaTestable <| (1: Nat) = 0 +-- Testing the MetaTestable class can be inferred +example : MetaTestable <| (1: Nat) = 0 := inferInstance -#synth MetaTestable (NamedBinder "a" (∀ (a : Nat), a ≤ 1)) +example : MetaTestable (NamedBinder "a" (∀ (a : Nat), a ≤ 1)) := inferInstance -#synth MetaTestable (NamedBinder "a" (∀ (a : Nat), NamedBinder "b" (∀ (b : Nat), a ≤ b))) +example : MetaTestable (NamedBinder "a" (∀ (a : Nat), NamedBinder "b" (∀ (b : Nat), a ≤ b))) := inferInstance -set_option pp.universes true in +-- Main tests: finding counter-examples +/-- +info: +=================== +Found a counter-example! +_a is irrelevant (unused) +issue: False does not hold +(0 shrinks) +------------------- +--- +info: some (Lean.Expr.app + (Lean.Expr.app + (Lean.Expr.app + (Lean.Expr.app + (Lean.Expr.const `mt []) + (Lean.Expr.forallE `a (Lean.Expr.const `Nat []) (Lean.Expr.const `False []) (Lean.BinderInfo.default))) + (Lean.Expr.const `False [])) + (Lean.Expr.lam + `h + (Lean.Expr.forallE `a (Lean.Expr.const `Nat []) (Lean.Expr.const `False []) (Lean.BinderInfo.default)) + (Lean.Expr.app + (Lean.Expr.bvar 0) + (Lean.Expr.app + (Lean.Expr.app + (Lean.Expr.const `instNonemptyOfInhabited [Lean.Level.succ (Lean.Level.zero)]) + (Lean.Expr.const `Nat [])) + (Lean.Expr.const `instInhabitedNat []))) + (Lean.BinderInfo.default))) + (Lean.Expr.app + (Lean.Expr.app + (Lean.Expr.app (Lean.Expr.app (Lean.Expr.const `mt []) (Lean.Expr.const `False [])) (Lean.Expr.const `False [])) + (Lean.Expr.app (Lean.Expr.const `id [Lean.Level.zero]) (Lean.Expr.const `False []))) + (Lean.Expr.app + (Lean.Expr.app + (Lean.Expr.app (Lean.Expr.const `of_decide_eq_false []) (Lean.Expr.const `False [])) + (Lean.Expr.const `instDecidableFalse [])) + (Lean.Expr.app + (Lean.Expr.app (Lean.Expr.const `Eq.refl [Lean.Level.succ (Lean.Level.zero)]) (Lean.Expr.const `Bool [])) + (Lean.Expr.const `Bool.false []))))) +-/ +#guard_msgs in #eval MetaTestable.check (∀ (_a : Nat), False) (propExpr := Lean.Expr.forallE `a (Lean.Expr.const `Nat []) (Lean.Expr.const `False []) (Lean.BinderInfo.default)) -/- -AppBuilder for 'mkAppM', result contains metavariables - SampleableExt Nat + +/-- +info: +=================== +Found a counter-example! +a := 1 +issue: 1 < 1 does not hold +(0 shrinks) +------------------- +--- +info: some (Lean.Expr.app + (Lean.Expr.app + (Lean.Expr.app + (Lean.Expr.app + (Lean.Expr.const `mt []) + (Lean.Expr.forallE `a (Lean.Expr.const `Nat []) (Lean.Expr.const `False []) (Lean.BinderInfo.default))) + (Lean.Expr.const `False [])) + (Lean.Expr.lam + `x + (Lean.Expr.forallE `a (Lean.Expr.const `Nat []) (Lean.Expr.const `False []) (Lean.BinderInfo.default)) + (Lean.Expr.app + (Lean.Expr.bvar 0) + (Lean.Expr.app + (Lean.Expr.app + (Lean.Expr.app + (Lean.Expr.const `Plausible.SampleableExt.interp [Lean.Level.succ (Lean.Level.zero), Lean.Level.zero]) + (Lean.Expr.const `Nat [])) + (Lean.Expr.const `Plausible.Nat.sampleableExt [])) + (Lean.Expr.app + (Lean.Expr.app + (Lean.Expr.app (Lean.Expr.const `OfNat.ofNat [Lean.Level.zero]) (Lean.Expr.const `Nat [])) + (Lean.Expr.lit (Lean.Literal.natVal 1))) + (Lean.Expr.app (Lean.Expr.const `instOfNatNat []) (Lean.Expr.lit (Lean.Literal.natVal 1)))))) + (Lean.BinderInfo.default))) + (Lean.Expr.app + (Lean.Expr.app + (Lean.Expr.app (Lean.Expr.const `of_decide_eq_false []) (Lean.Expr.const `False [])) + (Lean.Expr.const `instDecidableFalse [])) + (Lean.Expr.app + (Lean.Expr.app (Lean.Expr.const `Eq.refl [Lean.Level.succ (Lean.Level.zero)]) (Lean.Expr.const `Bool [])) + (Lean.Expr.const `Bool.false [])))) -/ -set_option pp.universes true in +#guard_msgs in #eval MetaTestable.check (∀ (a : Nat), a < 1) (propExpr := Lean.Expr.forallE `a (Lean.Expr.const `Nat []) (Lean.Expr.const `False []) (Lean.BinderInfo.default)) - +/-- +info: +=================== +Found a counter-example! +a := 0 +b := 0 +issue: 0 < 0 does not hold +(0 shrinks) +------------------- +--- +info: some (Lean.Expr.app + (Lean.Expr.app + (Lean.Expr.app + (Lean.Expr.app + (Lean.Expr.const `mt []) + (Lean.Expr.forallE + `a + (Lean.Expr.const `Nat []) + (Lean.Expr.forallE + `b + (Lean.Expr.const `Nat []) + (Lean.Expr.app + (Lean.Expr.app + (Lean.Expr.app + (Lean.Expr.app (Lean.Expr.const `LT.lt [Lean.Level.succ (Lean.Level.zero)]) (Lean.Expr.const `Nat [])) + (Lean.Expr.const `instLTNat [])) + (Lean.Expr.bvar 1)) + (Lean.Expr.bvar 0)) + (Lean.BinderInfo.default)) + (Lean.BinderInfo.default))) + (Lean.Expr.forallE + `b + (Lean.Expr.const `Nat []) + (Lean.Expr.app + (Lean.Expr.app + (Lean.Expr.app + (Lean.Expr.app (Lean.Expr.const `LT.lt [Lean.Level.succ (Lean.Level.zero)]) (Lean.Expr.const `Nat [])) + (Lean.Expr.const `instLTNat [])) + (Lean.Expr.app + (Lean.Expr.app + (Lean.Expr.app + (Lean.Expr.const `Plausible.SampleableExt.interp [Lean.Level.succ (Lean.Level.zero), Lean.Level.zero]) + (Lean.Expr.const `Nat [])) + (Lean.Expr.const `Plausible.Nat.sampleableExt [])) + (Lean.Expr.app + (Lean.Expr.app + (Lean.Expr.app (Lean.Expr.const `OfNat.ofNat [Lean.Level.zero]) (Lean.Expr.const `Nat [])) + (Lean.Expr.lit (Lean.Literal.natVal 0))) + (Lean.Expr.app (Lean.Expr.const `instOfNatNat []) (Lean.Expr.lit (Lean.Literal.natVal 0)))))) + (Lean.Expr.bvar 0)) + (Lean.BinderInfo.default))) + (Lean.Expr.lam + `x + (Lean.Expr.forallE + `a + (Lean.Expr.const `Nat []) + (Lean.Expr.forallE + `b + (Lean.Expr.const `Nat []) + (Lean.Expr.app + (Lean.Expr.app + (Lean.Expr.app + (Lean.Expr.app (Lean.Expr.const `LT.lt [Lean.Level.succ (Lean.Level.zero)]) (Lean.Expr.const `Nat [])) + (Lean.Expr.const `instLTNat [])) + (Lean.Expr.bvar 1)) + (Lean.Expr.bvar 0)) + (Lean.BinderInfo.default)) + (Lean.BinderInfo.default)) + (Lean.Expr.app + (Lean.Expr.bvar 0) + (Lean.Expr.app + (Lean.Expr.app + (Lean.Expr.app + (Lean.Expr.const `Plausible.SampleableExt.interp [Lean.Level.succ (Lean.Level.zero), Lean.Level.zero]) + (Lean.Expr.const `Nat [])) + (Lean.Expr.const `Plausible.Nat.sampleableExt [])) + (Lean.Expr.app + (Lean.Expr.app + (Lean.Expr.app (Lean.Expr.const `OfNat.ofNat [Lean.Level.zero]) (Lean.Expr.const `Nat [])) + (Lean.Expr.lit (Lean.Literal.natVal 0))) + (Lean.Expr.app (Lean.Expr.const `instOfNatNat []) (Lean.Expr.lit (Lean.Literal.natVal 0)))))) + (Lean.BinderInfo.default))) + (Lean.Expr.app + (Lean.Expr.app + (Lean.Expr.app + (Lean.Expr.app + (Lean.Expr.const `mt []) + (Lean.Expr.forallE + `b + (Lean.Expr.const `Nat []) + (Lean.Expr.app + (Lean.Expr.app + (Lean.Expr.app + (Lean.Expr.app (Lean.Expr.const `LT.lt [Lean.Level.succ (Lean.Level.zero)]) (Lean.Expr.const `Nat [])) + (Lean.Expr.const `instLTNat [])) + (Lean.Expr.app + (Lean.Expr.app + (Lean.Expr.app + (Lean.Expr.const + `Plausible.SampleableExt.interp + [Lean.Level.succ (Lean.Level.zero), Lean.Level.zero]) + (Lean.Expr.const `Nat [])) + (Lean.Expr.const `Plausible.Nat.sampleableExt [])) + (Lean.Expr.app + (Lean.Expr.app + (Lean.Expr.app (Lean.Expr.const `OfNat.ofNat [Lean.Level.zero]) (Lean.Expr.const `Nat [])) + (Lean.Expr.lit (Lean.Literal.natVal 0))) + (Lean.Expr.app (Lean.Expr.const `instOfNatNat []) (Lean.Expr.lit (Lean.Literal.natVal 0)))))) + (Lean.Expr.bvar 0)) + (Lean.BinderInfo.default))) + (Lean.Expr.app + (Lean.Expr.app + (Lean.Expr.app + (Lean.Expr.app (Lean.Expr.const `LT.lt [Lean.Level.succ (Lean.Level.zero)]) (Lean.Expr.const `Nat [])) + (Lean.Expr.const `instLTNat [])) + (Lean.Expr.app + (Lean.Expr.app + (Lean.Expr.app + (Lean.Expr.const `Plausible.SampleableExt.interp [Lean.Level.succ (Lean.Level.zero), Lean.Level.zero]) + (Lean.Expr.const `Nat [])) + (Lean.Expr.const `Plausible.Nat.sampleableExt [])) + (Lean.Expr.app + (Lean.Expr.app + (Lean.Expr.app (Lean.Expr.const `OfNat.ofNat [Lean.Level.zero]) (Lean.Expr.const `Nat [])) + (Lean.Expr.lit (Lean.Literal.natVal 0))) + (Lean.Expr.app (Lean.Expr.const `instOfNatNat []) (Lean.Expr.lit (Lean.Literal.natVal 0)))))) + (Lean.Expr.app + (Lean.Expr.app + (Lean.Expr.app + (Lean.Expr.const `Plausible.SampleableExt.interp [Lean.Level.succ (Lean.Level.zero), Lean.Level.zero]) + (Lean.Expr.const `Nat [])) + (Lean.Expr.const `Plausible.Nat.sampleableExt [])) + (Lean.Expr.app + (Lean.Expr.app + (Lean.Expr.app (Lean.Expr.const `OfNat.ofNat [Lean.Level.zero]) (Lean.Expr.const `Nat [])) + (Lean.Expr.lit (Lean.Literal.natVal 0))) + (Lean.Expr.app (Lean.Expr.const `instOfNatNat []) (Lean.Expr.lit (Lean.Literal.natVal 0))))))) + (Lean.Expr.lam + `x + (Lean.Expr.forallE + `b + (Lean.Expr.const `Nat []) + (Lean.Expr.app + (Lean.Expr.app + (Lean.Expr.app + (Lean.Expr.app (Lean.Expr.const `LT.lt [Lean.Level.succ (Lean.Level.zero)]) (Lean.Expr.const `Nat [])) + (Lean.Expr.const `instLTNat [])) + (Lean.Expr.app + (Lean.Expr.app + (Lean.Expr.app + (Lean.Expr.const + `Plausible.SampleableExt.interp + [Lean.Level.succ (Lean.Level.zero), Lean.Level.zero]) + (Lean.Expr.const `Nat [])) + (Lean.Expr.const `Plausible.Nat.sampleableExt [])) + (Lean.Expr.app + (Lean.Expr.app + (Lean.Expr.app (Lean.Expr.const `OfNat.ofNat [Lean.Level.zero]) (Lean.Expr.const `Nat [])) + (Lean.Expr.lit (Lean.Literal.natVal 0))) + (Lean.Expr.app (Lean.Expr.const `instOfNatNat []) (Lean.Expr.lit (Lean.Literal.natVal 0)))))) + (Lean.Expr.bvar 0)) + (Lean.BinderInfo.default)) + (Lean.Expr.app + (Lean.Expr.bvar 0) + (Lean.Expr.app + (Lean.Expr.app + (Lean.Expr.app + (Lean.Expr.const `Plausible.SampleableExt.interp [Lean.Level.succ (Lean.Level.zero), Lean.Level.zero]) + (Lean.Expr.const `Nat [])) + (Lean.Expr.const `Plausible.Nat.sampleableExt [])) + (Lean.Expr.app + (Lean.Expr.app + (Lean.Expr.app (Lean.Expr.const `OfNat.ofNat [Lean.Level.zero]) (Lean.Expr.const `Nat [])) + (Lean.Expr.lit (Lean.Literal.natVal 0))) + (Lean.Expr.app (Lean.Expr.const `instOfNatNat []) (Lean.Expr.lit (Lean.Literal.natVal 0)))))) + (Lean.BinderInfo.default))) + (Lean.Expr.app + (Lean.Expr.app + (Lean.Expr.app + (Lean.Expr.const `of_decide_eq_false []) + (Lean.Expr.app + (Lean.Expr.app + (Lean.Expr.app + (Lean.Expr.app (Lean.Expr.const `LT.lt [Lean.Level.succ (Lean.Level.zero)]) (Lean.Expr.const `Nat [])) + (Lean.Expr.const `instLTNat [])) + (Lean.Expr.app + (Lean.Expr.app + (Lean.Expr.app + (Lean.Expr.const + `Plausible.SampleableExt.interp + [Lean.Level.succ (Lean.Level.zero), Lean.Level.zero]) + (Lean.Expr.const `Nat [])) + (Lean.Expr.const `Plausible.Nat.sampleableExt [])) + (Lean.Expr.app + (Lean.Expr.app + (Lean.Expr.app (Lean.Expr.const `OfNat.ofNat [Lean.Level.zero]) (Lean.Expr.const `Nat [])) + (Lean.Expr.lit (Lean.Literal.natVal 0))) + (Lean.Expr.app (Lean.Expr.const `instOfNatNat []) (Lean.Expr.lit (Lean.Literal.natVal 0)))))) + (Lean.Expr.app + (Lean.Expr.app + (Lean.Expr.app + (Lean.Expr.const `Plausible.SampleableExt.interp [Lean.Level.succ (Lean.Level.zero), Lean.Level.zero]) + (Lean.Expr.const `Nat [])) + (Lean.Expr.const `Plausible.Nat.sampleableExt [])) + (Lean.Expr.app + (Lean.Expr.app + (Lean.Expr.app (Lean.Expr.const `OfNat.ofNat [Lean.Level.zero]) (Lean.Expr.const `Nat [])) + (Lean.Expr.lit (Lean.Literal.natVal 0))) + (Lean.Expr.app (Lean.Expr.const `instOfNatNat []) (Lean.Expr.lit (Lean.Literal.natVal 0))))))) + (Lean.Expr.app + (Lean.Expr.app + (Lean.Expr.const `Nat.decLt []) + (Lean.Expr.app + (Lean.Expr.app + (Lean.Expr.app + (Lean.Expr.const `Plausible.SampleableExt.interp [Lean.Level.succ (Lean.Level.zero), Lean.Level.zero]) + (Lean.Expr.const `Nat [])) + (Lean.Expr.const `Plausible.Nat.sampleableExt [])) + (Lean.Expr.app + (Lean.Expr.app + (Lean.Expr.app (Lean.Expr.const `OfNat.ofNat [Lean.Level.zero]) (Lean.Expr.const `Nat [])) + (Lean.Expr.lit (Lean.Literal.natVal 0))) + (Lean.Expr.app (Lean.Expr.const `instOfNatNat []) (Lean.Expr.lit (Lean.Literal.natVal 0)))))) + (Lean.Expr.app + (Lean.Expr.app + (Lean.Expr.app + (Lean.Expr.const `Plausible.SampleableExt.interp [Lean.Level.succ (Lean.Level.zero), Lean.Level.zero]) + (Lean.Expr.const `Nat [])) + (Lean.Expr.const `Plausible.Nat.sampleableExt [])) + (Lean.Expr.app + (Lean.Expr.app + (Lean.Expr.app (Lean.Expr.const `OfNat.ofNat [Lean.Level.zero]) (Lean.Expr.const `Nat [])) + (Lean.Expr.lit (Lean.Literal.natVal 0))) + (Lean.Expr.app (Lean.Expr.const `instOfNatNat []) (Lean.Expr.lit (Lean.Literal.natVal 0))))))) + (Lean.Expr.app + (Lean.Expr.app (Lean.Expr.const `Eq.refl [Lean.Level.succ (Lean.Level.zero)]) (Lean.Expr.const `Bool [])) + (Lean.Expr.const `Bool.false []))))) +-/ +#guard_msgs in #eval MetaTestable.check (∀ (a b : Nat), a < b) (propExpr := (Lean.Expr.forallE `a (Lean.Expr.const `Nat []) @@ -131,8 +493,308 @@ set_option pp.universes true in (Lean.BinderInfo.default)) (Lean.BinderInfo.default))) -set_option linter.unusedVariables false in -#eval MetaTestable.check (∀ (a b : Nat), a < 0) (propExpr := (Lean.Expr.forallE +/-- +info: +=================== +Found a counter-example! +a := 0 +_b is irrelevant (unused) +issue: 0 < 0 does not hold +(0 shrinks) +------------------- +--- +info: some (Lean.Expr.app + (Lean.Expr.app + (Lean.Expr.app + (Lean.Expr.app + (Lean.Expr.const `mt []) + (Lean.Expr.forallE + `a + (Lean.Expr.const `Nat []) + (Lean.Expr.forallE + `b + (Lean.Expr.const `Nat []) + (Lean.Expr.app + (Lean.Expr.app + (Lean.Expr.app + (Lean.Expr.app (Lean.Expr.const `LT.lt [Lean.Level.succ (Lean.Level.zero)]) (Lean.Expr.const `Nat [])) + (Lean.Expr.const `instLTNat [])) + (Lean.Expr.bvar 1)) + (Lean.Expr.lit (Lean.Literal.natVal 0))) + (Lean.BinderInfo.default)) + (Lean.BinderInfo.default))) + (Lean.Expr.forallE + `b + (Lean.Expr.const `Nat []) + (Lean.Expr.app + (Lean.Expr.app + (Lean.Expr.app + (Lean.Expr.app (Lean.Expr.const `LT.lt [Lean.Level.succ (Lean.Level.zero)]) (Lean.Expr.const `Nat [])) + (Lean.Expr.const `instLTNat [])) + (Lean.Expr.app + (Lean.Expr.app + (Lean.Expr.app + (Lean.Expr.const `Plausible.SampleableExt.interp [Lean.Level.succ (Lean.Level.zero), Lean.Level.zero]) + (Lean.Expr.const `Nat [])) + (Lean.Expr.const `Plausible.Nat.sampleableExt [])) + (Lean.Expr.app + (Lean.Expr.app + (Lean.Expr.app (Lean.Expr.const `OfNat.ofNat [Lean.Level.zero]) (Lean.Expr.const `Nat [])) + (Lean.Expr.lit (Lean.Literal.natVal 0))) + (Lean.Expr.app (Lean.Expr.const `instOfNatNat []) (Lean.Expr.lit (Lean.Literal.natVal 0)))))) + (Lean.Expr.lit (Lean.Literal.natVal 0))) + (Lean.BinderInfo.default))) + (Lean.Expr.lam + `x + (Lean.Expr.forallE + `a + (Lean.Expr.const `Nat []) + (Lean.Expr.forallE + `b + (Lean.Expr.const `Nat []) + (Lean.Expr.app + (Lean.Expr.app + (Lean.Expr.app + (Lean.Expr.app (Lean.Expr.const `LT.lt [Lean.Level.succ (Lean.Level.zero)]) (Lean.Expr.const `Nat [])) + (Lean.Expr.const `instLTNat [])) + (Lean.Expr.bvar 1)) + (Lean.Expr.lit (Lean.Literal.natVal 0))) + (Lean.BinderInfo.default)) + (Lean.BinderInfo.default)) + (Lean.Expr.app + (Lean.Expr.bvar 0) + (Lean.Expr.app + (Lean.Expr.app + (Lean.Expr.app + (Lean.Expr.const `Plausible.SampleableExt.interp [Lean.Level.succ (Lean.Level.zero), Lean.Level.zero]) + (Lean.Expr.const `Nat [])) + (Lean.Expr.const `Plausible.Nat.sampleableExt [])) + (Lean.Expr.app + (Lean.Expr.app + (Lean.Expr.app (Lean.Expr.const `OfNat.ofNat [Lean.Level.zero]) (Lean.Expr.const `Nat [])) + (Lean.Expr.lit (Lean.Literal.natVal 0))) + (Lean.Expr.app (Lean.Expr.const `instOfNatNat []) (Lean.Expr.lit (Lean.Literal.natVal 0)))))) + (Lean.BinderInfo.default))) + (Lean.Expr.app + (Lean.Expr.app + (Lean.Expr.app + (Lean.Expr.app + (Lean.Expr.const `mt []) + (Lean.Expr.forallE + `b + (Lean.Expr.const `Nat []) + (Lean.Expr.app + (Lean.Expr.app + (Lean.Expr.app + (Lean.Expr.app (Lean.Expr.const `LT.lt [Lean.Level.succ (Lean.Level.zero)]) (Lean.Expr.const `Nat [])) + (Lean.Expr.const `instLTNat [])) + (Lean.Expr.app + (Lean.Expr.app + (Lean.Expr.app + (Lean.Expr.const + `Plausible.SampleableExt.interp + [Lean.Level.succ (Lean.Level.zero), Lean.Level.zero]) + (Lean.Expr.const `Nat [])) + (Lean.Expr.const `Plausible.Nat.sampleableExt [])) + (Lean.Expr.app + (Lean.Expr.app + (Lean.Expr.app (Lean.Expr.const `OfNat.ofNat [Lean.Level.zero]) (Lean.Expr.const `Nat [])) + (Lean.Expr.lit (Lean.Literal.natVal 0))) + (Lean.Expr.app (Lean.Expr.const `instOfNatNat []) (Lean.Expr.lit (Lean.Literal.natVal 0)))))) + (Lean.Expr.app + (Lean.Expr.app + (Lean.Expr.app (Lean.Expr.const `OfNat.ofNat [Lean.Level.zero]) (Lean.Expr.const `Nat [])) + (Lean.Expr.lit (Lean.Literal.natVal 0))) + (Lean.Expr.app (Lean.Expr.const `instOfNatNat []) (Lean.Expr.lit (Lean.Literal.natVal 0))))) + (Lean.BinderInfo.default))) + (Lean.Expr.app + (Lean.Expr.app + (Lean.Expr.app + (Lean.Expr.app (Lean.Expr.const `LT.lt [Lean.Level.succ (Lean.Level.zero)]) (Lean.Expr.const `Nat [])) + (Lean.Expr.const `instLTNat [])) + (Lean.Expr.app + (Lean.Expr.app + (Lean.Expr.app + (Lean.Expr.const `Plausible.SampleableExt.interp [Lean.Level.succ (Lean.Level.zero), Lean.Level.zero]) + (Lean.Expr.const `Nat [])) + (Lean.Expr.const `Plausible.Nat.sampleableExt [])) + (Lean.Expr.app + (Lean.Expr.app + (Lean.Expr.app (Lean.Expr.const `OfNat.ofNat [Lean.Level.zero]) (Lean.Expr.const `Nat [])) + (Lean.Expr.lit (Lean.Literal.natVal 0))) + (Lean.Expr.app (Lean.Expr.const `instOfNatNat []) (Lean.Expr.lit (Lean.Literal.natVal 0)))))) + (Lean.Expr.app + (Lean.Expr.app + (Lean.Expr.app (Lean.Expr.const `OfNat.ofNat [Lean.Level.zero]) (Lean.Expr.const `Nat [])) + (Lean.Expr.lit (Lean.Literal.natVal 0))) + (Lean.Expr.app (Lean.Expr.const `instOfNatNat []) (Lean.Expr.lit (Lean.Literal.natVal 0)))))) + (Lean.Expr.lam + `h + (Lean.Expr.forallE + `b + (Lean.Expr.const `Nat []) + (Lean.Expr.app + (Lean.Expr.app + (Lean.Expr.app + (Lean.Expr.app (Lean.Expr.const `LT.lt [Lean.Level.succ (Lean.Level.zero)]) (Lean.Expr.const `Nat [])) + (Lean.Expr.const `instLTNat [])) + (Lean.Expr.app + (Lean.Expr.app + (Lean.Expr.app + (Lean.Expr.const + `Plausible.SampleableExt.interp + [Lean.Level.succ (Lean.Level.zero), Lean.Level.zero]) + (Lean.Expr.const `Nat [])) + (Lean.Expr.const `Plausible.Nat.sampleableExt [])) + (Lean.Expr.app + (Lean.Expr.app + (Lean.Expr.app (Lean.Expr.const `OfNat.ofNat [Lean.Level.zero]) (Lean.Expr.const `Nat [])) + (Lean.Expr.lit (Lean.Literal.natVal 0))) + (Lean.Expr.app (Lean.Expr.const `instOfNatNat []) (Lean.Expr.lit (Lean.Literal.natVal 0)))))) + (Lean.Expr.app + (Lean.Expr.app + (Lean.Expr.app (Lean.Expr.const `OfNat.ofNat [Lean.Level.zero]) (Lean.Expr.const `Nat [])) + (Lean.Expr.lit (Lean.Literal.natVal 0))) + (Lean.Expr.app (Lean.Expr.const `instOfNatNat []) (Lean.Expr.lit (Lean.Literal.natVal 0))))) + (Lean.BinderInfo.default)) + (Lean.Expr.app + (Lean.Expr.bvar 0) + (Lean.Expr.app + (Lean.Expr.app + (Lean.Expr.const `instNonemptyOfInhabited [Lean.Level.succ (Lean.Level.zero)]) + (Lean.Expr.const `Nat [])) + (Lean.Expr.const `instInhabitedNat []))) + (Lean.BinderInfo.default))) + (Lean.Expr.app + (Lean.Expr.app + (Lean.Expr.app + (Lean.Expr.app + (Lean.Expr.const `mt []) + (Lean.Expr.app + (Lean.Expr.app + (Lean.Expr.app + (Lean.Expr.app (Lean.Expr.const `LT.lt [Lean.Level.succ (Lean.Level.zero)]) (Lean.Expr.const `Nat [])) + (Lean.Expr.const `instLTNat [])) + (Lean.Expr.app + (Lean.Expr.app + (Lean.Expr.app + (Lean.Expr.const + `Plausible.SampleableExt.interp + [Lean.Level.succ (Lean.Level.zero), Lean.Level.zero]) + (Lean.Expr.const `Nat [])) + (Lean.Expr.const `Plausible.Nat.sampleableExt [])) + (Lean.Expr.app + (Lean.Expr.app + (Lean.Expr.app (Lean.Expr.const `OfNat.ofNat [Lean.Level.zero]) (Lean.Expr.const `Nat [])) + (Lean.Expr.lit (Lean.Literal.natVal 0))) + (Lean.Expr.app (Lean.Expr.const `instOfNatNat []) (Lean.Expr.lit (Lean.Literal.natVal 0)))))) + (Lean.Expr.app + (Lean.Expr.app + (Lean.Expr.app (Lean.Expr.const `OfNat.ofNat [Lean.Level.zero]) (Lean.Expr.const `Nat [])) + (Lean.Expr.lit (Lean.Literal.natVal 0))) + (Lean.Expr.app (Lean.Expr.const `instOfNatNat []) (Lean.Expr.lit (Lean.Literal.natVal 0)))))) + (Lean.Expr.app + (Lean.Expr.app + (Lean.Expr.app + (Lean.Expr.app (Lean.Expr.const `LT.lt [Lean.Level.succ (Lean.Level.zero)]) (Lean.Expr.const `Nat [])) + (Lean.Expr.const `instLTNat [])) + (Lean.Expr.app + (Lean.Expr.app + (Lean.Expr.app + (Lean.Expr.const + `Plausible.SampleableExt.interp + [Lean.Level.succ (Lean.Level.zero), Lean.Level.zero]) + (Lean.Expr.const `Nat [])) + (Lean.Expr.const `Plausible.Nat.sampleableExt [])) + (Lean.Expr.app + (Lean.Expr.app + (Lean.Expr.app (Lean.Expr.const `OfNat.ofNat [Lean.Level.zero]) (Lean.Expr.const `Nat [])) + (Lean.Expr.lit (Lean.Literal.natVal 0))) + (Lean.Expr.app (Lean.Expr.const `instOfNatNat []) (Lean.Expr.lit (Lean.Literal.natVal 0)))))) + (Lean.Expr.app + (Lean.Expr.app + (Lean.Expr.app (Lean.Expr.const `OfNat.ofNat [Lean.Level.zero]) (Lean.Expr.const `Nat [])) + (Lean.Expr.lit (Lean.Literal.natVal 0))) + (Lean.Expr.app (Lean.Expr.const `instOfNatNat []) (Lean.Expr.lit (Lean.Literal.natVal 0)))))) + (Lean.Expr.app + (Lean.Expr.const `id [Lean.Level.zero]) + (Lean.Expr.app + (Lean.Expr.app + (Lean.Expr.app + (Lean.Expr.app (Lean.Expr.const `LT.lt [Lean.Level.succ (Lean.Level.zero)]) (Lean.Expr.const `Nat [])) + (Lean.Expr.const `instLTNat [])) + (Lean.Expr.app + (Lean.Expr.app + (Lean.Expr.app + (Lean.Expr.const + `Plausible.SampleableExt.interp + [Lean.Level.succ (Lean.Level.zero), Lean.Level.zero]) + (Lean.Expr.const `Nat [])) + (Lean.Expr.const `Plausible.Nat.sampleableExt [])) + (Lean.Expr.app + (Lean.Expr.app + (Lean.Expr.app (Lean.Expr.const `OfNat.ofNat [Lean.Level.zero]) (Lean.Expr.const `Nat [])) + (Lean.Expr.lit (Lean.Literal.natVal 0))) + (Lean.Expr.app (Lean.Expr.const `instOfNatNat []) (Lean.Expr.lit (Lean.Literal.natVal 0)))))) + (Lean.Expr.app + (Lean.Expr.app + (Lean.Expr.app (Lean.Expr.const `OfNat.ofNat [Lean.Level.zero]) (Lean.Expr.const `Nat [])) + (Lean.Expr.lit (Lean.Literal.natVal 0))) + (Lean.Expr.app (Lean.Expr.const `instOfNatNat []) (Lean.Expr.lit (Lean.Literal.natVal 0))))))) + (Lean.Expr.app + (Lean.Expr.app + (Lean.Expr.app + (Lean.Expr.const `of_decide_eq_false []) + (Lean.Expr.app + (Lean.Expr.app + (Lean.Expr.app + (Lean.Expr.app (Lean.Expr.const `LT.lt [Lean.Level.succ (Lean.Level.zero)]) (Lean.Expr.const `Nat [])) + (Lean.Expr.const `instLTNat [])) + (Lean.Expr.app + (Lean.Expr.app + (Lean.Expr.app + (Lean.Expr.const + `Plausible.SampleableExt.interp + [Lean.Level.succ (Lean.Level.zero), Lean.Level.zero]) + (Lean.Expr.const `Nat [])) + (Lean.Expr.const `Plausible.Nat.sampleableExt [])) + (Lean.Expr.app + (Lean.Expr.app + (Lean.Expr.app (Lean.Expr.const `OfNat.ofNat [Lean.Level.zero]) (Lean.Expr.const `Nat [])) + (Lean.Expr.lit (Lean.Literal.natVal 0))) + (Lean.Expr.app (Lean.Expr.const `instOfNatNat []) (Lean.Expr.lit (Lean.Literal.natVal 0)))))) + (Lean.Expr.app + (Lean.Expr.app + (Lean.Expr.app (Lean.Expr.const `OfNat.ofNat [Lean.Level.zero]) (Lean.Expr.const `Nat [])) + (Lean.Expr.lit (Lean.Literal.natVal 0))) + (Lean.Expr.app (Lean.Expr.const `instOfNatNat []) (Lean.Expr.lit (Lean.Literal.natVal 0)))))) + (Lean.Expr.app + (Lean.Expr.app + (Lean.Expr.const `Nat.decLt []) + (Lean.Expr.app + (Lean.Expr.app + (Lean.Expr.app + (Lean.Expr.const + `Plausible.SampleableExt.interp + [Lean.Level.succ (Lean.Level.zero), Lean.Level.zero]) + (Lean.Expr.const `Nat [])) + (Lean.Expr.const `Plausible.Nat.sampleableExt [])) + (Lean.Expr.app + (Lean.Expr.app + (Lean.Expr.app (Lean.Expr.const `OfNat.ofNat [Lean.Level.zero]) (Lean.Expr.const `Nat [])) + (Lean.Expr.lit (Lean.Literal.natVal 0))) + (Lean.Expr.app (Lean.Expr.const `instOfNatNat []) (Lean.Expr.lit (Lean.Literal.natVal 0)))))) + (Lean.Expr.app + (Lean.Expr.app + (Lean.Expr.app (Lean.Expr.const `OfNat.ofNat [Lean.Level.zero]) (Lean.Expr.const `Nat [])) + (Lean.Expr.lit (Lean.Literal.natVal 0))) + (Lean.Expr.app (Lean.Expr.const `instOfNatNat []) (Lean.Expr.lit (Lean.Literal.natVal 0)))))) + (Lean.Expr.app + (Lean.Expr.app (Lean.Expr.const `Eq.refl [Lean.Level.succ (Lean.Level.zero)]) (Lean.Expr.const `Bool [])) + (Lean.Expr.const `Bool.false [])))))) +-/ +#guard_msgs in +#eval MetaTestable.check (∀ (a _b : Nat), a < 0) (propExpr := (Lean.Expr.forallE `a (Lean.Expr.const `Nat []) (Lean.Expr.forallE @@ -149,24 +811,3 @@ set_option linter.unusedVariables false in (Lean.Expr.lit (Lean.Literal.natVal 0))) (Lean.BinderInfo.default)) (Lean.BinderInfo.default))) - -#expr ∀ (a b : Nat), a < b -#expr ∀ (a _b : Nat), a < 0 - -#check Lean.Expr.lit (Lean.Literal.natVal 0) - -#expr Expr → MetaM (Option Expr) - - -elab "disprove%" t:term : term => do - let tgt ← elabType t - let cfg : Configuration := {} - let (some code') ← disproveM? cfg tgt | throwError "disprove% failed" - logInfo s!"disproof: {← ppExpr code'}; \ntype: {← ppExpr <| (← inferType code')}" - return tgt - -#check disprove% ∀ (a b : Nat), a < b - -#check disprove% ∀ (a b : Nat), a < b ∨ b < a - -#check disprove% False ∧ False diff --git a/Test/Vacuous.lean b/Test/Vacuous.lean index c928f02..91a7005 100644 --- a/Test/Vacuous.lean +++ b/Test/Vacuous.lean @@ -30,8 +30,8 @@ info: fun a => False.elim (of_decide_eq_false (Eq.refl false) a) : 2 < 1 → 1 info: =================== Found a counter-example! -n := 6 -issue: 6 < 4 does not hold +n := 5 +issue: 5 < 4 does not hold (0 shrinks) ------------------- --- @@ -39,7 +39,7 @@ info: Vacuous Implication. Hypothesis ∀ (n : Nat), n < 4 is never satisfied --- info: fun a => False.elim - (mt (fun x => x (Plausible.SampleableExt.interp 6)) (of_decide_eq_false (Eq.refl false)) + (mt (fun x => x (Plausible.SampleableExt.interp 5)) (of_decide_eq_false (Eq.refl false)) a) : (∀ (n : Nat), n < 4) → 1 ≤ 3 -/ #guard_msgs in @@ -49,8 +49,8 @@ info: fun a => info: =================== Found a counter-example! -n := 6 -issue: 6 < 4 does not hold +n := 5 +issue: 5 < 4 does not hold (0 shrinks) ------------------- --- @@ -64,8 +64,8 @@ example : (∀ n: Nat, n < (4: Nat)) → 4 ≤ 3 := by vacuous info: =================== Found a counter-example! -n := 6 -issue: 6 < 4 does not hold +n := 5 +issue: 5 < 4 does not hold (0 shrinks) ------------------- --- @@ -79,8 +79,8 @@ example (h: ∀ n: Nat, n < (4: Nat)) : 4 ≤ 3 := by vacuous info: =================== Found a counter-example! -n := 6 -issue: 6 < 4 does not hold +n := 5 +issue: 5 < 4 does not hold (0 shrinks) ------------------- --- From 466ad9164ed960edfd4a2159e6bc590f0584a09e Mon Sep 17 00:00:00 2001 From: Siddhartha Gadgil Date: Mon, 25 Nov 2024 07:41:05 +0530 Subject: [PATCH 20/30] random search implemented --- Plausible/MetaTestable.lean | 21 +++--- Plausible/RandomSearch.lean | 147 ++++++++++++++++++++++++++++++++++++ 2 files changed, 157 insertions(+), 11 deletions(-) create mode 100644 Plausible/RandomSearch.lean diff --git a/Plausible/MetaTestable.lean b/Plausible/MetaTestable.lean index 8262fd2..ea2642a 100644 --- a/Plausible/MetaTestable.lean +++ b/Plausible/MetaTestable.lean @@ -77,6 +77,15 @@ def andProp? (e: Expr) : MetaM (Option Expr × Option Expr) := do else return (none, none) +def notProp? (e: Expr) : MetaM (Option Expr) := do + let prop := mkSort levelZero + let α ← mkFreshExprMVar prop + let e' ← mkAppM ``Not #[α] + if ← isDefEq e' e then + return (← Lean.getExprMVarAssignment? α.mvarId!) + else + return none + def forallProp? (e: Expr) : MetaM (Option Expr × Option Expr) := do let u ← mkFreshLevelMVar let α ← mkFreshExprMVar (mkSort u) @@ -696,17 +705,7 @@ def disproveM? (cfg : Configuration) (tgt: Expr) : MetaM <| Option Expr := do let code ← unsafe evalExpr (Expr → MetaM (Option Expr)) expectedType e code tgt +-- Negate --- #eval MetaTestable.check (∀ (x y z a : Nat) (h1 : 3 < x) (h2 : 3 < y), x - y = y - x) --- Configuration.verbose --- #eval MetaTestable.check (∀ x : Nat, ∀ y : Nat, x + y = y + x) Configuration.verbose --- #eval MetaTestable.check (∀ (x : (Nat × Nat)), x.fst - x.snd - 10 = x.snd - x.fst - 10) --- Configuration.verbose --- #eval MetaTestable.check (∀ (x : Nat) (h : 10 < x), 5 < x) Configuration.verbose - -macro tk:"#test " e:term : command => `(command| #eval%$tk MetaTestable.check $e) - --- #test ∀ (x : Nat) (h : 5 < x), 10 < x --- #test ∀ (x : Nat) (h : 10 < x), 5 < x end Plausible diff --git a/Plausible/RandomSearch.lean b/Plausible/RandomSearch.lean new file mode 100644 index 0000000..092c265 --- /dev/null +++ b/Plausible/RandomSearch.lean @@ -0,0 +1,147 @@ +import Plausible.MetaTestable +open Lean Meta Elab Term + +namespace Plausible + +#check Decidable.not_not +#check Classical.choice + +open Classical in +theorem not_not (p: Prop) : ¬ ¬ p → p := + fun h' => if h : p then h else absurd h h' + +open Classical in +theorem not_and {p q : Prop} : ¬ (p ∧ q) → ¬ p ∨ ¬ q := + fun h => if hp : p then Or.inr fun hq => h ⟨hp, hq⟩ else Or.inl hp + +theorem induced_or {p₁ q₁ p₂ q₂ : Prop} (h₁ : p₁ → q₁) (h₂ : p₂ → q₂) : p₁ ∨ p₂ → q₁ ∨ q₂ + | Or.inl h => Or.inl (h₁ h) + | Or.inr h => Or.inr (h₂ h) + +theorem induced_and {p₁ q₁ p₂ q₂ : Prop} (h₁ : p₁ → q₁) (h₂ : p₂ → q₂) : p₁ ∧ p₂ → q₁ ∧ q₂ + | ⟨h1, h2⟩ => ⟨h₁ h1, h₂ h2⟩ + +theorem induced_exists {p q : α → Prop}(f : ∀x : α, p x → q x) : (∃ x, p x) → ∃ x, q x + | ⟨x, h⟩ => ⟨x, f x h⟩ + +theorem not_or_mp {p q : Prop} (h : ¬ (p ∨ q)) : ¬ p ∧ ¬ q := + ⟨fun h' => h (Or.inl h'), fun h' => h (Or.inr h')⟩ + +theorem not_forall {p : α → Prop} : ¬ (∀ x, p x) → ∃ x, ¬ p x := by + intro h + apply not_not + intro contra + have l : ∀ (x : α), p x := by + intro x + by_cases h' : p x + · exact h' + · exfalso + exact contra ⟨x, h'⟩ + contradiction + + +#print not_not +#check not_and_of_not_or_not +#check not_or +#check not_exists_of_forall_not +#check not_forall_of_exists_not +#check Iff.mp + +partial def negate (e: Expr) : MetaM <| Expr × Expr := do + match ← notProp? e with + | some e' => + let negProof ← mkAppOptM ``id #[e] + return (e', negProof) + | _ => + match ← andProp? e with + | (some e₁, some e₂) => + let (negProp₁, negProof₁) ← negate e₁ + let (negProp₂, negProof₂) ← negate e₂ + let negProp ← mkAppM ``Or #[negProp₁, negProp₂] + let negPf ← withLocalDeclD `h (← mkAppM ``Not #[negProp]) fun h => do + let h' ← mkAppM ``not_or_mp #[h] + let negProof ← mkAppM ``induced_and #[negProof₁, negProof₂, h'] + mkLambdaFVars #[h] negProof + return (negProp, negPf) + | _ => + match ← orProp? e with + | (some e₁, some e₂) => + let (negProp₁, negProof₁) ← negate e₁ + let (negProp₂, negProof₂) ← negate e₂ + let negProp ← mkAppM ``And #[negProp₁, negProp₂] + let negPf ← withLocalDeclD `h (← mkAppM ``Not #[negProp]) fun h => do + let h' ← mkAppM ``not_and #[h] + let negProof ← mkAppM ``induced_or #[negProof₁, negProof₂, h'] + mkLambdaFVars #[h] negProof + return (negProp, negPf) + | _ => + match ← existsProp? e with + | (some p, some α) => + logInfo m!"Negating existential quantifier {← ppExpr e}, family: {← ppExpr p}, domain: {← ppExpr α}" + let (negProp, negProofFamily) ← withLocalDeclD `x α fun x => do + let y ← mkAppM' p #[x] + let (y, _) ← dsimp y {} + let (negProp, negProof) ← negate y + pure (← mkForallFVars #[x] negProp, ← mkLambdaFVars #[x] negProof) + logInfo m!"Negation of {← ppExpr e} is {← ppExpr negProp}" + let negPf ← withLocalDeclD `h (← mkAppM ``Not #[negProp]) fun h => do + let h' ← mkAppM ``not_forall #[h] + let negProof ← mkAppM ``induced_exists #[negProofFamily, h'] + mkLambdaFVars #[h] negProof + return (negProp, negPf) + | _ => + let negProp ← mkAppM ``Not #[e] + let negProof ← mkAppM ``not_not #[e] + return (negProp, negProof) + +open Elab Term +elab "neg_neg" t:term : term => do + let e ← elabTerm t none + let (negProp, negProof) ← negate e + let pfType ← inferType negProof + -- disproof of negation implies `e` + let goal ← mkArrow (← mkAppM ``Not #[negProp]) e + let check ← isDefEq pfType goal + logInfo m!"Negation of {← ppExpr e} is {← ppExpr negProp}" + logInfo m!"Proof of negation is {← ppExpr negProof} with type {← ppExpr <| ← inferType negProof}, should be {← ppExpr goal}" + logInfo m!"check: {check}" + return negProp + +#check neg_neg (1 = 4) + +#check neg_neg ¬(1 = 4) + +#check neg_neg ¬(1 = 4) ∧ (2 < 4) + +#check neg_neg (1 = 4) ∨ (¬ (2 < 4)) + +#check neg_neg ∃ x : Nat, x > 0 + +open Lean.Parser.Tactic Tactic +syntax (name := searchSyntax) "random_search" (config)? : tactic +elab_rules : tactic | `(tactic| random_search $[$cfg]?) => withMainContext do + let cfg ← elabConfig (mkOptionalNode cfg) let (_, g) ← (← getMainGoal).revert ((← getLocalHyps).map (Expr.fvarId!)) + let cfg := { cfg with + traceDiscarded := cfg.traceDiscarded || (← isTracingEnabledFor `plausible.discarded), + traceSuccesses := cfg.traceSuccesses || (← isTracingEnabledFor `plausible.success), + traceShrink := cfg.traceShrink || (← isTracingEnabledFor `plausible.shrink.steps), + traceShrinkCandidates := cfg.traceShrinkCandidates + || (← isTracingEnabledFor `plausible.shrink.candidates) } + g.withContext do + let tgt ← g.getType + let (negProp, negProof) ← negate tgt + match ← disproveM? cfg negProp with + | some pf => + let pf' ← mkAppM' negProof #[pf] + logInfo m!"Found a proof by negated counter-example!" + g.assign pf' + | none => + throwError "could not find a proof by negated counter-example" + +example : ¬ (1 = 4) := by + random_search + +example : ∃ n : Nat, n > 0 ∧ n < 4 := by + random_search + +end Plausible From 241a26fce4d98a439bd0cbc643e4f081f44a0c32 Mon Sep 17 00:00:00 2001 From: Siddhartha Gadgil Date: Mon, 25 Nov 2024 08:25:00 +0530 Subject: [PATCH 21/30] clean up comments and tests --- Plausible/MetaTestable.lean | 26 ++++------------- Plausible/RandomSearch.lean | 32 ++++---------------- Plausible/Vacuous.lean | 5 ++++ Test/RandomSeach.lean | 42 +++++++++++++++++++++++++++ Test/Vacuous.lean | 58 ++----------------------------------- 5 files changed, 60 insertions(+), 103 deletions(-) create mode 100644 Test/RandomSeach.lean diff --git a/Plausible/MetaTestable.lean b/Plausible/MetaTestable.lean index ea2642a..d8f75f9 100644 --- a/Plausible/MetaTestable.lean +++ b/Plausible/MetaTestable.lean @@ -12,30 +12,14 @@ open Lean /-! # `MetaTestable` Class -MetaTestable propositions have a procedure that can generate counter-examples -together with a proof that they invalidate the proposition. - -This is a port of the Haskell QuickCheck library. - -## Creating Customized Instances - -The type classes `MetaTestable`, `SampleableExt` and `Shrinkable` are the -means by which `Plausible` creates samples and tests them. For instance, -the proposition `∀ i j : Nat, i ≤ j` has a `MetaTestable` instance because `Nat` -is sampleable and `i ≤ j` is decidable. Once `Plausible` finds the `MetaTestable` -instance, it can start using the instance to repeatedly creating samples -and checking whether they satisfy the property. Once it has found a -counter-example it will then use a `Shrinkable` instance to reduce the -example. This allows the user to create new instances and apply -`Plausible` to new situations. +MetaTestable propositions have a procedure that can generate counter-examples together with a proof that they invalidate the proposition. This is a refinement of the `Testable` class. +Instances of `Testable` are built using `Sampleable` instances. For `MetaTestable` instances, we also need an expression for the proxy in a sampleable type, which is an additional requirement. ## Main definitions * `MetaTestable` class -* `MetaTestable.check`: a way to test a proposition using random examples - - +* `MetaTestable.check`: a way to test a proposition using random examples, givng a proof of the counter-example -/ namespace Plausible @@ -117,7 +101,6 @@ def forallPropProp? (e: Expr) : MetaM (Option Expr × Option Expr × Option Expr else return (none, none, none) - def impProp? (e: Expr) : MetaM (Option Expr × Option Expr) := do let u ← mkFreshLevelMVar let α ← mkFreshExprMVar (mkSort u) @@ -522,7 +505,7 @@ instance varTestable [SampleableExt α] [ProxyExpr α] {β : α → Prop} [∀ x addVarInfo var finalX h hExpr finalR - +-- Typeclass inference does not seem to work across `mkSelfContained`, so we need to provide instances for the basic types instance : ProxyExpr Bool := (inferInstance : ToExpr Bool) instance : ProxyExpr String := (inferInstance : ToExpr String) instance : ProxyExpr Nat := (inferInstance : ToExpr Nat) @@ -575,6 +558,7 @@ where theorem prop_iff_subtype (p : α → Prop) (β : α → Prop) : NamedBinder var (∀ (x : α), NamedBinder var' (p x → β x)) ↔ ∀ (x : Subtype p), β x.val := by simp [Subtype.forall, NamedBinder] +-- This is unlikely to ever be used successfully because of the instance of `SampleableExt (Subtype p)` needed. Should probably have classes for implications instead. instance (priority := 2000) subtypeVarTestable {p : α → Prop} {β : α → Prop} [∀ x, PrintableProp (p x)] [∀ x, MetaTestable (β x)][ToExpr α] diff --git a/Plausible/RandomSearch.lean b/Plausible/RandomSearch.lean index 092c265..18b7049 100644 --- a/Plausible/RandomSearch.lean +++ b/Plausible/RandomSearch.lean @@ -1,11 +1,14 @@ import Plausible.MetaTestable +/-! +## Random Search + +The `random_search` tactic tries to prove a result by searching for a counter-example to its negation using `MetaTestable`. If a counter-example is found, we can prove the result by showing that the negation is false. +-/ + open Lean Meta Elab Term namespace Plausible -#check Decidable.not_not -#check Classical.choice - open Classical in theorem not_not (p: Prop) : ¬ ¬ p → p := fun h' => if h : p then h else absurd h h' @@ -40,13 +43,6 @@ theorem not_forall {p : α → Prop} : ¬ (∀ x, p x) → ∃ x, ¬ p x := by contradiction -#print not_not -#check not_and_of_not_or_not -#check not_or -#check not_exists_of_forall_not -#check not_forall_of_exists_not -#check Iff.mp - partial def negate (e: Expr) : MetaM <| Expr × Expr := do match ← notProp? e with | some e' => @@ -107,16 +103,6 @@ elab "neg_neg" t:term : term => do logInfo m!"check: {check}" return negProp -#check neg_neg (1 = 4) - -#check neg_neg ¬(1 = 4) - -#check neg_neg ¬(1 = 4) ∧ (2 < 4) - -#check neg_neg (1 = 4) ∨ (¬ (2 < 4)) - -#check neg_neg ∃ x : Nat, x > 0 - open Lean.Parser.Tactic Tactic syntax (name := searchSyntax) "random_search" (config)? : tactic elab_rules : tactic | `(tactic| random_search $[$cfg]?) => withMainContext do @@ -138,10 +124,4 @@ elab_rules : tactic | `(tactic| random_search $[$cfg]?) => withMainContext do | none => throwError "could not find a proof by negated counter-example" -example : ¬ (1 = 4) := by - random_search - -example : ∃ n : Nat, n > 0 ∧ n < 4 := by - random_search - end Plausible diff --git a/Plausible/Vacuous.lean b/Plausible/Vacuous.lean index a0660cb..ec24703 100644 --- a/Plausible/Vacuous.lean +++ b/Plausible/Vacuous.lean @@ -1,6 +1,11 @@ import Lean import Plausible import Plausible.MetaTestable +/-! +## Vacuous Implication + +The `vacuous` tactic is used to prove vacuous implications. We use the plausible search for counterexamples (actually at the `MetaTestable` level) to find a counterexample to the hypothesis. If a counterexample is found, we can prove the vacuous implication by showing that the hypothesis is never satisfied. +-/ open Lean Meta Elab Tactic Plausible diff --git a/Test/RandomSeach.lean b/Test/RandomSeach.lean new file mode 100644 index 0000000..a28a36a --- /dev/null +++ b/Test/RandomSeach.lean @@ -0,0 +1,42 @@ +import Plausible.RandomSearch + +open Plausible + +/- +¬1 = 4 : Prop +-/ +#check neg_neg (1 = 4) + +/- +Negation of ¬1 = 4 is 1 = 4 +-/ +#check neg_neg ¬(1 = 4) + +/- +Proof of negation is fun h => + induced_and id (not_not (2 < 4)) + (not_or_mp h) with type ¬(1 = 4 ∨ ¬2 < 4) → ¬1 = 4 ∧ 2 < 4, should be ¬(1 = 4 ∨ ¬2 < 4) → ¬1 = 4 ∧ 2 < 4 +-/ +#check neg_neg ¬(1 = 4) ∧ (2 < 4) + +/- +check: true +-/ +#check neg_neg (1 = 4) ∨ (¬ (2 < 4)) + +/- +check: true +-/ +#check neg_neg ∃ x : Nat, x > 0 + +example : ¬ (1 = 4) := by + /- + Found a proof by negated counter-example! + -/ + random_search + +example : ∃ n : Nat, n > 0 ∧ n < 4 := by + /- + Found a proof by negated counter-example! + -/ + random_search diff --git a/Test/Vacuous.lean b/Test/Vacuous.lean index 91a7005..9e34ff8 100644 --- a/Test/Vacuous.lean +++ b/Test/Vacuous.lean @@ -26,66 +26,12 @@ info: fun a => False.elim (of_decide_eq_false (Eq.refl false) a) : 2 < 1 → 1 #guard_msgs in #check prove_vacuous ((2 : Nat) < 1) → 1 ≤ 3 -/-- -info: -=================== -Found a counter-example! -n := 5 -issue: 5 < 4 does not hold -(0 shrinks) -------------------- ---- -info: Vacuous Implication. Hypothesis ∀ (n : Nat), n < 4 is never satisfied ---- -info: fun a => - False.elim - (mt (fun x => x (Plausible.SampleableExt.interp 5)) (of_decide_eq_false (Eq.refl false)) - a) : (∀ (n : Nat), n < 4) → 1 ≤ 3 --/ -#guard_msgs in +-- We cannot use `guard_msgs` here because of random search. Need to add seeds as a parameter. + #check prove_vacuous (∀ n: Nat, n < (4: Nat)) → 1 ≤ 3 -/-- -info: -=================== -Found a counter-example! -n := 5 -issue: 5 < 4 does not hold -(0 shrinks) -------------------- ---- -info: Vacuous Implication. Hypothesis ∀ (n : Nat), n < 4 is never satisfied --/ -#guard_msgs in --- Can negate hypothesis in goal statement example : (∀ n: Nat, n < (4: Nat)) → 4 ≤ 3 := by vacuous -/-- -info: -=================== -Found a counter-example! -n := 5 -issue: 5 < 4 does not hold -(0 shrinks) -------------------- ---- -info: Vacuous Implication. Hypothesis ∀ (n : Nat), n < 4 is never satisfied --/ -#guard_msgs in --- Can negate hypothesis in local context example (h: ∀ n: Nat, n < (4: Nat)) : 4 ≤ 3 := by vacuous -/-- -info: -=================== -Found a counter-example! -n := 5 -issue: 5 < 4 does not hold -(0 shrinks) -------------------- ---- -info: Vacuous Implication. Hypothesis ∀ (n : Nat), n < 4 is never satisfied --/ -#guard_msgs in --- Can negate hypothesis other than the first one example (m: Nat) (h: ∀ n: Nat, n < (4: Nat)) : m ≤ 3 := by vacuous From 0187e13d45b9bb394774b4d7e2bc4b3e141b1e1d Mon Sep 17 00:00:00 2001 From: Siddhartha Gadgil Date: Tue, 26 Nov 2024 09:18:49 +0530 Subject: [PATCH 22/30] optional expression in failure, rename run --- Plausible/MetaTestable.lean | 90 +++--- Test/MetaTestable.lean | 624 +----------------------------------- 2 files changed, 46 insertions(+), 668 deletions(-) diff --git a/Plausible/MetaTestable.lean b/Plausible/MetaTestable.lean index d8f75f9..18ea244 100644 --- a/Plausible/MetaTestable.lean +++ b/Plausible/MetaTestable.lean @@ -204,13 +204,13 @@ inductive MetaTestResult (p : Prop) where guarantee that there will be no false positive. The last component, `n`, is the number of times that the counter-example was shrunk. -/ - | failure : ¬ p → Expr → List String → Nat → MetaTestResult p + | failure : ¬ p → Option Expr → List String → Nat → MetaTestResult p deriving Inhabited /-- `MetaTestable p` uses random examples to try to disprove `p`. -/ class MetaTestable (p : Prop) where - run (cfg : Configuration) (minimize : Bool) (propExpr : Expr) : MGen (MetaTestResult p) + runExpr (cfg : Configuration) (minimize : Bool) (propExpr : Expr) : MGen (MetaTestResult p) namespace MetaTestResult @@ -236,22 +236,22 @@ def checkDisproof (pf prop: Expr) : MetaM Unit := do /-- Combine the test result for properties `p` and `q` to create a test for their conjunction. -/ def and : MetaTestResult p → MetaTestResult q → Expr → MetaM (MetaTestResult (p ∧ q)) - | failure h pf xs n, _, e => do + | failure h pf? xs n, _, e => do let (some e₁, some e₂) ← andProp? e | throwError m!"Expected an `And` proposition, got {← ppExpr e}" - let pf' ← withLocalDeclD `h e fun h => do + let pf'? ← + pf?.mapM fun pf => withLocalDeclD `h e fun h => do let x ← mkAppOptM ``And.left #[e₁, e₂, h] - let e' ← mkAppM' pf #[x] + let e' ← mkAppM' pf #[x] mkLambdaFVars #[h] e' - checkDisproof pf' e - return failure (fun h2 => h h2.left) pf' xs n - | _, failure h pf xs n, e => do + -- checkDisproof pf' e + return failure (fun h2 => h h2.left) pf'? xs n + | _, failure h pf? xs n, e => do let (some e₁, some e₂) ← andProp? e | throwError m!"Expected an `And` proposition, got {← ppExpr e}" - - let pf' ← withLocalDeclD `h e fun h => do + let pf' ← pf?.mapM fun pf => withLocalDeclD `h e fun h => do let x ← mkAppOptM ``And.right #[e₁, e₂, h] let e' ← mkAppM' pf #[x] mkLambdaFVars #[h] e' - checkDisproof pf' e + -- checkDisproof pf' e return failure (fun h2 => h h2.right) pf' xs n | success h1, success h2, _ => return success <| combine (combine (PSum.inr And.intro) h1) h2 @@ -284,8 +284,8 @@ to find counter-examples to `q`. -/ def imp (h : q → p) (hExp: Expr) (r : MetaTestResult p) (p : Unit ⊕' (p → q) := PSum.inl ()) : MetaM (MetaTestResult q) := match r with - | failure h2 pf xs n => do - let pf' ← mkAppM ``mt #[hExp, pf] + | failure h2 pf? xs n => do + let pf' ← pf?.mapM fun pf => mkAppM ``mt #[hExp, pf] return failure (mt h h2) pf' xs n | success h2 => return success <| combine p h2 | gaveUp n => return gaveUp n @@ -300,8 +300,8 @@ we record that value using this function so that our counter-examples can be informative. -/ def addInfo (x : String) (h : q → p) (hExp: Expr) (r : MetaTestResult p) (p : Unit ⊕' (p → q) := PSum.inl ()) : (MetaM <| MetaTestResult q) := do - if let failure h2 pf xs n := r then - let pf' ← mkAppM ``mt #[hExp, pf] + if let failure h2 pf? xs n := r then + let pf' ← pf?.mapM fun pf => mkAppM ``mt #[hExp, pf] return failure (mt h h2) pf' (x :: xs) n else imp h hExp r p @@ -325,31 +325,31 @@ open MetaTestResult abbrev ProxyExpr α [SampleableExt α] := ToExpr (SampleableExt.proxy α) -def runProp (p : Prop) [MetaTestable p] : Configuration → Bool → Expr → MGen (MetaTestResult p) := fun cfg b e => do - MetaTestable.run cfg b e +def runPropExpr (p : Prop) [MetaTestable p] : Configuration → Bool → Expr → MGen (MetaTestResult p) := fun cfg b e => do + runExpr cfg b e /-- A `dbgTrace` with special formatting -/ def slimTrace {m : Type → Type _} [Pure m] (s : String) : m PUnit := dbgTrace s!"[Plausible: {s}]" (fun _ => pure ()) instance andTestable [MetaTestable p] [MetaTestable q] : MetaTestable (p ∧ q) where - run := fun cfg min e => do + runExpr := fun cfg min e => do let (some e₁, some e₂) ← andProp? e | throwError m!"Expected an `And` proposition, got {← ppExpr e}" - let xp ← runProp p cfg min e₁ - let xq ← runProp q cfg min e₂ + let xp ← runPropExpr p cfg min e₁ + let xq ← runPropExpr q cfg min e₂ and xp xq e instance orTestable [MetaTestable p] [MetaTestable q] : MetaTestable (p ∨ q) where - run := fun cfg min e => do + runExpr := fun cfg min e => do let (some e₁, some e₂) ← orProp? e | throwError m!"Expected an `Or` proposition, got {← ppExpr e}" - let xp ← runProp p cfg min e₁ + let xp ← runPropExpr p cfg min e₁ -- As a little performance optimization we can just not run the second -- test if the first succeeds match xp with | success (PSum.inl h) => return success (PSum.inl h) | success (PSum.inr h) => return success (PSum.inr <| Or.inl h) | _ => - let xq ← runProp q cfg min e₂ + let xq ← runPropExpr q cfg min e₂ or xp xq e theorem iff_resolve (p q : Prop) : (p ↔ q) ↔ p ∧ q ∨ ¬p ∧ ¬q := by @@ -360,9 +360,9 @@ theorem iff_resolve (p q : Prop) : (p ↔ q) ↔ p ∧ q ∨ ¬p ∧ ¬q := by rcases h with ⟨hleft, hright⟩ | ⟨hleft, hright⟩ <;> simp [hleft, hright] instance iffTestable [MetaTestable ((p ∧ q) ∨ (¬ p ∧ ¬ q))] : MetaTestable (p ↔ q) where - run := fun cfg min e => do + runExpr := fun cfg min e => do let some (α, β) ← iffProp? e | throwError m!"Expected an `Iff` proposition, got {← ppExpr e}" - let h ← runProp ((p ∧ q) ∨ (¬ p ∧ ¬ q)) cfg min e + let h ← runPropExpr ((p ∧ q) ∨ (¬ p ∧ ¬ q)) cfg min e let hExp ← mkAppM ``iff_resolve #[α, β] iff (iff_resolve p q) hExp h @@ -370,11 +370,11 @@ variable {var : String} instance decGuardTestable [PrintableProp p] [Decidable p] {β : p → Prop} [∀ h, MetaTestable (β h)] : MetaTestable (NamedBinder var <| ∀ h, β h) where - run := fun cfg min e => do + runExpr := fun cfg min e => do if h : p then let (some βExp, some pExp) ← forallProp? e | throwError m!"Expected a `Forall` proposition, got {← ppExpr e}" let yExp ← mkAppM' βExp #[pExp] - let res := runProp (β h) cfg min yExp + let res := runPropExpr (β h) cfg min yExp let h' := (· <| h) let decInstType ← mkAppM ``Decidable #[pExp] let inst ← synthInstance decInstType @@ -395,16 +395,16 @@ instance decGuardTestable [PrintableProp p] [Decidable p] {β : p → Prop} [∀ instance forallTypesTestable {f : Type → Prop} [MetaTestable (f Int)] : MetaTestable (NamedBinder var <| ∀ x, f x) where - run := fun cfg min e => do - let r ← runProp (f Int) cfg min e + runExpr := fun cfg min e => do + let r ← runPropExpr (f Int) cfg min e addVarInfo var "Int" (· <| Int) e r -- TODO: only in mathlib: @[pp_with_univ] instance (priority := 100) forallTypesULiftTestable.{u} {f : Type u → Prop} [MetaTestable (f (ULift.{u} Int))] : MetaTestable (NamedBinder var <| ∀ x, f x) where - run cfg min e := do - let r ← runProp (f (ULift Int)) cfg min e + runExpr cfg min e := do + let r ← runPropExpr (f (ULift Int)) cfg min e addVarInfo var "ULift Int" (· <| ULift Int) e r /-- @@ -441,7 +441,7 @@ partial def minimizeAux [SampleableExt α] [ProxyExpr α] {β : α → Prop} [ let samp ← synthInstance instType let xInterp ← mkAppOptM ``SampleableExt.interp #[αExp, samp, xExpr] let e' ← mkAppM' βExpr #[xInterp] - let res ← OptionT.lift <| MetaTestable.runProp (β (SampleableExt.interp candidate)) cfg true e' + let res ← OptionT.lift <| runPropExpr (β (SampleableExt.interp candidate)) cfg true e' if res.isFailure then if cfg.traceShrink then slimTrace s!"{var} shrunk to {repr candidate} from {repr x}" @@ -468,7 +468,7 @@ open Lean Meta Elab Term Tactic in bound variable with it. -/ instance varTestable [SampleableExt α] [ProxyExpr α] {β : α → Prop} [∀ x, MetaTestable (β x)] : MetaTestable (NamedBinder var <| ∀ x : α, β x) where - run := fun cfg min e => do + runExpr := fun cfg min e => do let (some βExp, some αExp) ← forallProp? e | throwError m!"Expected a `Forall` proposition, got {← ppExpr e}" let x ← SampleableExt.sample let xExpr := toExpr x @@ -482,7 +482,7 @@ instance varTestable [SampleableExt α] [ProxyExpr α] {β : α → Prop} [∀ x let (e', _) ← dsimp e' {} if cfg.traceSuccesses || cfg.traceDiscarded then slimTrace s!"{var} := {repr x}" - let r ← MetaTestable.runProp (β <| SampleableExt.interp x) cfg false e' + let r ← runPropExpr (β <| SampleableExt.interp x) cfg false e' let ⟨finalX, finalR⟩ ← if isFailure r then if cfg.traceSuccesses then @@ -534,20 +534,20 @@ theorem bool_to_prop_fmly (β : Prop → Prop): NamedBinder var (∀ (p : Prop), instance propVarTestable {β : Prop → Prop} [h: ∀ b : Bool, MetaTestable (β b)] : MetaTestable (NamedBinder var <| ∀ p : Prop, β p) where - run := fun cfg min e => do + runExpr := fun cfg min e => do let (some βExpr, _) ← forallProp? e | throwError m!"Expected a `Forall` proposition, got {← ppExpr e}" - let p ← MetaTestable.runProp (NamedBinder var <| ∀ b : Bool, β b) cfg min e + let p ← runPropExpr (NamedBinder var <| ∀ b : Bool, β b) cfg min e let e' ← mkAppM ``bool_to_prop_fmly #[βExpr] imp (bool_to_prop_fmly β) e' p instance (priority := high) unusedVarTestable {β : Prop} [Nonempty α] [MetaTestable β] : MetaTestable (NamedBinder var (α → β)) where - run := fun cfg min e => do + runExpr := fun cfg min e => do if cfg.traceDiscarded || cfg.traceSuccesses then slimTrace s!"{var} is unused" let (some aExp, some e') ← impProp? e | throwError m!"Expected an `Imp` proposition, got {← ppExpr e}" - let r ← MetaTestable.runProp β cfg min e' + let r ← runPropExpr β cfg min e' let hExp ← mkAppOptM ``id #[e'] let finalR ← addInfo s!"{var} is irrelevant (unused)" id hExp r let h := (· <| Classical.ofNonempty) @@ -564,14 +564,14 @@ instance (priority := 2000) subtypeVarTestable {p : α → Prop} {β : α → Pr [∀ x, MetaTestable (β x)][ToExpr α] [SampleableExt (Subtype p)] [ProxyExpr (Subtype p)] {var'} : MetaTestable (NamedBinder var <| (x : α) → NamedBinder var' <| p x → β x) where - run cfg min e := do + runExpr cfg min e := do let (some pExp, some βExp, some _) ← forallPropProp? e | throwError m!"Expected a `Forall` proposition with arrow, got {← ppExpr e}" let subType ← mkAppM ``Subtype #[pExp] letI (x : Subtype p) : MetaTestable (β x) := - { run := fun cfg min e => do + { runExpr := fun cfg min e => do let xExp := toExpr x.val let y ← mkAppM' βExp #[xExp] - let r ← MetaTestable.runProp (β x.val) cfg min y + let r ← runPropExpr (β x.val) cfg min y let idExp ← mkAppOptM ``id #[e] addInfo s!"guard: {printProp (p x)} (by construction)" id idExp r (PSum.inr id) } do @@ -579,13 +579,13 @@ instance (priority := 2000) subtypeVarTestable {p : α → Prop} {β : α → Pr let x' ← mkAppM ``Subtype.val #[x] let y ← mkAppM' βExp #[x'] mkForallFVars #[x] y - let r ← @MetaTestable.run (∀ x : Subtype p, β x.val) (@varTestable var _ _ _ _ _) cfg min e' + let r ← @runExpr (∀ x : Subtype p, β x.val) (@varTestable var _ _ _ _ _) cfg min e' let hExp ← mkAppM ``prop_iff_subtype #[pExp, βExp] iff (prop_iff_subtype p β) hExp r instance (priority := low) decidableTestable {p : Prop} [PrintableProp p] [Decidable p] : MetaTestable p where - run := fun _ _ e _ => do + runExpr := fun _ _ e _ => do if h : p then return success (PSum.inr h) else @@ -632,7 +632,7 @@ def MetaTestable.runSuiteAux (p : Prop) [MetaTestable p] (propExpr: Expr) (cfg : if cfg.traceSuccesses then slimTrace s!"New sample" slimTrace s!"Retrying up to {cfg.numRetries} times until guards hold" - let x ← retry (ReaderT.run (MetaTestable.runProp p cfg true propExpr) ⟨size⟩) cfg.numRetries + let x ← retry (ReaderT.run (runPropExpr p cfg true propExpr) ⟨size⟩) cfg.numRetries match x with | success (PSum.inl ()) => runSuiteAux p propExpr cfg r n | gaveUp g => runSuiteAux p propExpr cfg (giveUp g r) n @@ -669,7 +669,7 @@ def MetaTestable.check (p : Prop) (cfg : Configuration := {}) Lean.logInfo msg else Lean.logInfo <| Testable.formatFailure msg xs n - return some pf + return pf def disproveM? (cfg : Configuration) (tgt: Expr) : MetaM <| Option Expr := do let tgt' ← Decorations.addDecorations tgt diff --git a/Test/MetaTestable.lean b/Test/MetaTestable.lean index cfb5008..802ea4b 100644 --- a/Test/MetaTestable.lean +++ b/Test/MetaTestable.lean @@ -148,333 +148,12 @@ example : MetaTestable (NamedBinder "a" (∀ (a : Nat), a ≤ 1)) := inferInstan example : MetaTestable (NamedBinder "a" (∀ (a : Nat), NamedBinder "b" (∀ (b : Nat), a ≤ b))) := inferInstance --- Main tests: finding counter-examples -/-- -info: -=================== -Found a counter-example! -_a is irrelevant (unused) -issue: False does not hold -(0 shrinks) -------------------- ---- -info: some (Lean.Expr.app - (Lean.Expr.app - (Lean.Expr.app - (Lean.Expr.app - (Lean.Expr.const `mt []) - (Lean.Expr.forallE `a (Lean.Expr.const `Nat []) (Lean.Expr.const `False []) (Lean.BinderInfo.default))) - (Lean.Expr.const `False [])) - (Lean.Expr.lam - `h - (Lean.Expr.forallE `a (Lean.Expr.const `Nat []) (Lean.Expr.const `False []) (Lean.BinderInfo.default)) - (Lean.Expr.app - (Lean.Expr.bvar 0) - (Lean.Expr.app - (Lean.Expr.app - (Lean.Expr.const `instNonemptyOfInhabited [Lean.Level.succ (Lean.Level.zero)]) - (Lean.Expr.const `Nat [])) - (Lean.Expr.const `instInhabitedNat []))) - (Lean.BinderInfo.default))) - (Lean.Expr.app - (Lean.Expr.app - (Lean.Expr.app (Lean.Expr.app (Lean.Expr.const `mt []) (Lean.Expr.const `False [])) (Lean.Expr.const `False [])) - (Lean.Expr.app (Lean.Expr.const `id [Lean.Level.zero]) (Lean.Expr.const `False []))) - (Lean.Expr.app - (Lean.Expr.app - (Lean.Expr.app (Lean.Expr.const `of_decide_eq_false []) (Lean.Expr.const `False [])) - (Lean.Expr.const `instDecidableFalse [])) - (Lean.Expr.app - (Lean.Expr.app (Lean.Expr.const `Eq.refl [Lean.Level.succ (Lean.Level.zero)]) (Lean.Expr.const `Bool [])) - (Lean.Expr.const `Bool.false []))))) --/ -#guard_msgs in +-- Main tests: finding counterexamples #eval MetaTestable.check (∀ (_a : Nat), False) (propExpr := Lean.Expr.forallE `a (Lean.Expr.const `Nat []) (Lean.Expr.const `False []) (Lean.BinderInfo.default)) -/-- -info: -=================== -Found a counter-example! -a := 1 -issue: 1 < 1 does not hold -(0 shrinks) -------------------- ---- -info: some (Lean.Expr.app - (Lean.Expr.app - (Lean.Expr.app - (Lean.Expr.app - (Lean.Expr.const `mt []) - (Lean.Expr.forallE `a (Lean.Expr.const `Nat []) (Lean.Expr.const `False []) (Lean.BinderInfo.default))) - (Lean.Expr.const `False [])) - (Lean.Expr.lam - `x - (Lean.Expr.forallE `a (Lean.Expr.const `Nat []) (Lean.Expr.const `False []) (Lean.BinderInfo.default)) - (Lean.Expr.app - (Lean.Expr.bvar 0) - (Lean.Expr.app - (Lean.Expr.app - (Lean.Expr.app - (Lean.Expr.const `Plausible.SampleableExt.interp [Lean.Level.succ (Lean.Level.zero), Lean.Level.zero]) - (Lean.Expr.const `Nat [])) - (Lean.Expr.const `Plausible.Nat.sampleableExt [])) - (Lean.Expr.app - (Lean.Expr.app - (Lean.Expr.app (Lean.Expr.const `OfNat.ofNat [Lean.Level.zero]) (Lean.Expr.const `Nat [])) - (Lean.Expr.lit (Lean.Literal.natVal 1))) - (Lean.Expr.app (Lean.Expr.const `instOfNatNat []) (Lean.Expr.lit (Lean.Literal.natVal 1)))))) - (Lean.BinderInfo.default))) - (Lean.Expr.app - (Lean.Expr.app - (Lean.Expr.app (Lean.Expr.const `of_decide_eq_false []) (Lean.Expr.const `False [])) - (Lean.Expr.const `instDecidableFalse [])) - (Lean.Expr.app - (Lean.Expr.app (Lean.Expr.const `Eq.refl [Lean.Level.succ (Lean.Level.zero)]) (Lean.Expr.const `Bool [])) - (Lean.Expr.const `Bool.false [])))) --/ -#guard_msgs in #eval MetaTestable.check (∀ (a : Nat), a < 1) (propExpr := Lean.Expr.forallE `a (Lean.Expr.const `Nat []) (Lean.Expr.const `False []) (Lean.BinderInfo.default)) -/-- -info: -=================== -Found a counter-example! -a := 0 -b := 0 -issue: 0 < 0 does not hold -(0 shrinks) -------------------- ---- -info: some (Lean.Expr.app - (Lean.Expr.app - (Lean.Expr.app - (Lean.Expr.app - (Lean.Expr.const `mt []) - (Lean.Expr.forallE - `a - (Lean.Expr.const `Nat []) - (Lean.Expr.forallE - `b - (Lean.Expr.const `Nat []) - (Lean.Expr.app - (Lean.Expr.app - (Lean.Expr.app - (Lean.Expr.app (Lean.Expr.const `LT.lt [Lean.Level.succ (Lean.Level.zero)]) (Lean.Expr.const `Nat [])) - (Lean.Expr.const `instLTNat [])) - (Lean.Expr.bvar 1)) - (Lean.Expr.bvar 0)) - (Lean.BinderInfo.default)) - (Lean.BinderInfo.default))) - (Lean.Expr.forallE - `b - (Lean.Expr.const `Nat []) - (Lean.Expr.app - (Lean.Expr.app - (Lean.Expr.app - (Lean.Expr.app (Lean.Expr.const `LT.lt [Lean.Level.succ (Lean.Level.zero)]) (Lean.Expr.const `Nat [])) - (Lean.Expr.const `instLTNat [])) - (Lean.Expr.app - (Lean.Expr.app - (Lean.Expr.app - (Lean.Expr.const `Plausible.SampleableExt.interp [Lean.Level.succ (Lean.Level.zero), Lean.Level.zero]) - (Lean.Expr.const `Nat [])) - (Lean.Expr.const `Plausible.Nat.sampleableExt [])) - (Lean.Expr.app - (Lean.Expr.app - (Lean.Expr.app (Lean.Expr.const `OfNat.ofNat [Lean.Level.zero]) (Lean.Expr.const `Nat [])) - (Lean.Expr.lit (Lean.Literal.natVal 0))) - (Lean.Expr.app (Lean.Expr.const `instOfNatNat []) (Lean.Expr.lit (Lean.Literal.natVal 0)))))) - (Lean.Expr.bvar 0)) - (Lean.BinderInfo.default))) - (Lean.Expr.lam - `x - (Lean.Expr.forallE - `a - (Lean.Expr.const `Nat []) - (Lean.Expr.forallE - `b - (Lean.Expr.const `Nat []) - (Lean.Expr.app - (Lean.Expr.app - (Lean.Expr.app - (Lean.Expr.app (Lean.Expr.const `LT.lt [Lean.Level.succ (Lean.Level.zero)]) (Lean.Expr.const `Nat [])) - (Lean.Expr.const `instLTNat [])) - (Lean.Expr.bvar 1)) - (Lean.Expr.bvar 0)) - (Lean.BinderInfo.default)) - (Lean.BinderInfo.default)) - (Lean.Expr.app - (Lean.Expr.bvar 0) - (Lean.Expr.app - (Lean.Expr.app - (Lean.Expr.app - (Lean.Expr.const `Plausible.SampleableExt.interp [Lean.Level.succ (Lean.Level.zero), Lean.Level.zero]) - (Lean.Expr.const `Nat [])) - (Lean.Expr.const `Plausible.Nat.sampleableExt [])) - (Lean.Expr.app - (Lean.Expr.app - (Lean.Expr.app (Lean.Expr.const `OfNat.ofNat [Lean.Level.zero]) (Lean.Expr.const `Nat [])) - (Lean.Expr.lit (Lean.Literal.natVal 0))) - (Lean.Expr.app (Lean.Expr.const `instOfNatNat []) (Lean.Expr.lit (Lean.Literal.natVal 0)))))) - (Lean.BinderInfo.default))) - (Lean.Expr.app - (Lean.Expr.app - (Lean.Expr.app - (Lean.Expr.app - (Lean.Expr.const `mt []) - (Lean.Expr.forallE - `b - (Lean.Expr.const `Nat []) - (Lean.Expr.app - (Lean.Expr.app - (Lean.Expr.app - (Lean.Expr.app (Lean.Expr.const `LT.lt [Lean.Level.succ (Lean.Level.zero)]) (Lean.Expr.const `Nat [])) - (Lean.Expr.const `instLTNat [])) - (Lean.Expr.app - (Lean.Expr.app - (Lean.Expr.app - (Lean.Expr.const - `Plausible.SampleableExt.interp - [Lean.Level.succ (Lean.Level.zero), Lean.Level.zero]) - (Lean.Expr.const `Nat [])) - (Lean.Expr.const `Plausible.Nat.sampleableExt [])) - (Lean.Expr.app - (Lean.Expr.app - (Lean.Expr.app (Lean.Expr.const `OfNat.ofNat [Lean.Level.zero]) (Lean.Expr.const `Nat [])) - (Lean.Expr.lit (Lean.Literal.natVal 0))) - (Lean.Expr.app (Lean.Expr.const `instOfNatNat []) (Lean.Expr.lit (Lean.Literal.natVal 0)))))) - (Lean.Expr.bvar 0)) - (Lean.BinderInfo.default))) - (Lean.Expr.app - (Lean.Expr.app - (Lean.Expr.app - (Lean.Expr.app (Lean.Expr.const `LT.lt [Lean.Level.succ (Lean.Level.zero)]) (Lean.Expr.const `Nat [])) - (Lean.Expr.const `instLTNat [])) - (Lean.Expr.app - (Lean.Expr.app - (Lean.Expr.app - (Lean.Expr.const `Plausible.SampleableExt.interp [Lean.Level.succ (Lean.Level.zero), Lean.Level.zero]) - (Lean.Expr.const `Nat [])) - (Lean.Expr.const `Plausible.Nat.sampleableExt [])) - (Lean.Expr.app - (Lean.Expr.app - (Lean.Expr.app (Lean.Expr.const `OfNat.ofNat [Lean.Level.zero]) (Lean.Expr.const `Nat [])) - (Lean.Expr.lit (Lean.Literal.natVal 0))) - (Lean.Expr.app (Lean.Expr.const `instOfNatNat []) (Lean.Expr.lit (Lean.Literal.natVal 0)))))) - (Lean.Expr.app - (Lean.Expr.app - (Lean.Expr.app - (Lean.Expr.const `Plausible.SampleableExt.interp [Lean.Level.succ (Lean.Level.zero), Lean.Level.zero]) - (Lean.Expr.const `Nat [])) - (Lean.Expr.const `Plausible.Nat.sampleableExt [])) - (Lean.Expr.app - (Lean.Expr.app - (Lean.Expr.app (Lean.Expr.const `OfNat.ofNat [Lean.Level.zero]) (Lean.Expr.const `Nat [])) - (Lean.Expr.lit (Lean.Literal.natVal 0))) - (Lean.Expr.app (Lean.Expr.const `instOfNatNat []) (Lean.Expr.lit (Lean.Literal.natVal 0))))))) - (Lean.Expr.lam - `x - (Lean.Expr.forallE - `b - (Lean.Expr.const `Nat []) - (Lean.Expr.app - (Lean.Expr.app - (Lean.Expr.app - (Lean.Expr.app (Lean.Expr.const `LT.lt [Lean.Level.succ (Lean.Level.zero)]) (Lean.Expr.const `Nat [])) - (Lean.Expr.const `instLTNat [])) - (Lean.Expr.app - (Lean.Expr.app - (Lean.Expr.app - (Lean.Expr.const - `Plausible.SampleableExt.interp - [Lean.Level.succ (Lean.Level.zero), Lean.Level.zero]) - (Lean.Expr.const `Nat [])) - (Lean.Expr.const `Plausible.Nat.sampleableExt [])) - (Lean.Expr.app - (Lean.Expr.app - (Lean.Expr.app (Lean.Expr.const `OfNat.ofNat [Lean.Level.zero]) (Lean.Expr.const `Nat [])) - (Lean.Expr.lit (Lean.Literal.natVal 0))) - (Lean.Expr.app (Lean.Expr.const `instOfNatNat []) (Lean.Expr.lit (Lean.Literal.natVal 0)))))) - (Lean.Expr.bvar 0)) - (Lean.BinderInfo.default)) - (Lean.Expr.app - (Lean.Expr.bvar 0) - (Lean.Expr.app - (Lean.Expr.app - (Lean.Expr.app - (Lean.Expr.const `Plausible.SampleableExt.interp [Lean.Level.succ (Lean.Level.zero), Lean.Level.zero]) - (Lean.Expr.const `Nat [])) - (Lean.Expr.const `Plausible.Nat.sampleableExt [])) - (Lean.Expr.app - (Lean.Expr.app - (Lean.Expr.app (Lean.Expr.const `OfNat.ofNat [Lean.Level.zero]) (Lean.Expr.const `Nat [])) - (Lean.Expr.lit (Lean.Literal.natVal 0))) - (Lean.Expr.app (Lean.Expr.const `instOfNatNat []) (Lean.Expr.lit (Lean.Literal.natVal 0)))))) - (Lean.BinderInfo.default))) - (Lean.Expr.app - (Lean.Expr.app - (Lean.Expr.app - (Lean.Expr.const `of_decide_eq_false []) - (Lean.Expr.app - (Lean.Expr.app - (Lean.Expr.app - (Lean.Expr.app (Lean.Expr.const `LT.lt [Lean.Level.succ (Lean.Level.zero)]) (Lean.Expr.const `Nat [])) - (Lean.Expr.const `instLTNat [])) - (Lean.Expr.app - (Lean.Expr.app - (Lean.Expr.app - (Lean.Expr.const - `Plausible.SampleableExt.interp - [Lean.Level.succ (Lean.Level.zero), Lean.Level.zero]) - (Lean.Expr.const `Nat [])) - (Lean.Expr.const `Plausible.Nat.sampleableExt [])) - (Lean.Expr.app - (Lean.Expr.app - (Lean.Expr.app (Lean.Expr.const `OfNat.ofNat [Lean.Level.zero]) (Lean.Expr.const `Nat [])) - (Lean.Expr.lit (Lean.Literal.natVal 0))) - (Lean.Expr.app (Lean.Expr.const `instOfNatNat []) (Lean.Expr.lit (Lean.Literal.natVal 0)))))) - (Lean.Expr.app - (Lean.Expr.app - (Lean.Expr.app - (Lean.Expr.const `Plausible.SampleableExt.interp [Lean.Level.succ (Lean.Level.zero), Lean.Level.zero]) - (Lean.Expr.const `Nat [])) - (Lean.Expr.const `Plausible.Nat.sampleableExt [])) - (Lean.Expr.app - (Lean.Expr.app - (Lean.Expr.app (Lean.Expr.const `OfNat.ofNat [Lean.Level.zero]) (Lean.Expr.const `Nat [])) - (Lean.Expr.lit (Lean.Literal.natVal 0))) - (Lean.Expr.app (Lean.Expr.const `instOfNatNat []) (Lean.Expr.lit (Lean.Literal.natVal 0))))))) - (Lean.Expr.app - (Lean.Expr.app - (Lean.Expr.const `Nat.decLt []) - (Lean.Expr.app - (Lean.Expr.app - (Lean.Expr.app - (Lean.Expr.const `Plausible.SampleableExt.interp [Lean.Level.succ (Lean.Level.zero), Lean.Level.zero]) - (Lean.Expr.const `Nat [])) - (Lean.Expr.const `Plausible.Nat.sampleableExt [])) - (Lean.Expr.app - (Lean.Expr.app - (Lean.Expr.app (Lean.Expr.const `OfNat.ofNat [Lean.Level.zero]) (Lean.Expr.const `Nat [])) - (Lean.Expr.lit (Lean.Literal.natVal 0))) - (Lean.Expr.app (Lean.Expr.const `instOfNatNat []) (Lean.Expr.lit (Lean.Literal.natVal 0)))))) - (Lean.Expr.app - (Lean.Expr.app - (Lean.Expr.app - (Lean.Expr.const `Plausible.SampleableExt.interp [Lean.Level.succ (Lean.Level.zero), Lean.Level.zero]) - (Lean.Expr.const `Nat [])) - (Lean.Expr.const `Plausible.Nat.sampleableExt [])) - (Lean.Expr.app - (Lean.Expr.app - (Lean.Expr.app (Lean.Expr.const `OfNat.ofNat [Lean.Level.zero]) (Lean.Expr.const `Nat [])) - (Lean.Expr.lit (Lean.Literal.natVal 0))) - (Lean.Expr.app (Lean.Expr.const `instOfNatNat []) (Lean.Expr.lit (Lean.Literal.natVal 0))))))) - (Lean.Expr.app - (Lean.Expr.app (Lean.Expr.const `Eq.refl [Lean.Level.succ (Lean.Level.zero)]) (Lean.Expr.const `Bool [])) - (Lean.Expr.const `Bool.false []))))) --/ -#guard_msgs in #eval MetaTestable.check (∀ (a b : Nat), a < b) (propExpr := (Lean.Expr.forallE `a (Lean.Expr.const `Nat []) @@ -493,307 +172,6 @@ info: some (Lean.Expr.app (Lean.BinderInfo.default)) (Lean.BinderInfo.default))) -/-- -info: -=================== -Found a counter-example! -a := 0 -_b is irrelevant (unused) -issue: 0 < 0 does not hold -(0 shrinks) -------------------- ---- -info: some (Lean.Expr.app - (Lean.Expr.app - (Lean.Expr.app - (Lean.Expr.app - (Lean.Expr.const `mt []) - (Lean.Expr.forallE - `a - (Lean.Expr.const `Nat []) - (Lean.Expr.forallE - `b - (Lean.Expr.const `Nat []) - (Lean.Expr.app - (Lean.Expr.app - (Lean.Expr.app - (Lean.Expr.app (Lean.Expr.const `LT.lt [Lean.Level.succ (Lean.Level.zero)]) (Lean.Expr.const `Nat [])) - (Lean.Expr.const `instLTNat [])) - (Lean.Expr.bvar 1)) - (Lean.Expr.lit (Lean.Literal.natVal 0))) - (Lean.BinderInfo.default)) - (Lean.BinderInfo.default))) - (Lean.Expr.forallE - `b - (Lean.Expr.const `Nat []) - (Lean.Expr.app - (Lean.Expr.app - (Lean.Expr.app - (Lean.Expr.app (Lean.Expr.const `LT.lt [Lean.Level.succ (Lean.Level.zero)]) (Lean.Expr.const `Nat [])) - (Lean.Expr.const `instLTNat [])) - (Lean.Expr.app - (Lean.Expr.app - (Lean.Expr.app - (Lean.Expr.const `Plausible.SampleableExt.interp [Lean.Level.succ (Lean.Level.zero), Lean.Level.zero]) - (Lean.Expr.const `Nat [])) - (Lean.Expr.const `Plausible.Nat.sampleableExt [])) - (Lean.Expr.app - (Lean.Expr.app - (Lean.Expr.app (Lean.Expr.const `OfNat.ofNat [Lean.Level.zero]) (Lean.Expr.const `Nat [])) - (Lean.Expr.lit (Lean.Literal.natVal 0))) - (Lean.Expr.app (Lean.Expr.const `instOfNatNat []) (Lean.Expr.lit (Lean.Literal.natVal 0)))))) - (Lean.Expr.lit (Lean.Literal.natVal 0))) - (Lean.BinderInfo.default))) - (Lean.Expr.lam - `x - (Lean.Expr.forallE - `a - (Lean.Expr.const `Nat []) - (Lean.Expr.forallE - `b - (Lean.Expr.const `Nat []) - (Lean.Expr.app - (Lean.Expr.app - (Lean.Expr.app - (Lean.Expr.app (Lean.Expr.const `LT.lt [Lean.Level.succ (Lean.Level.zero)]) (Lean.Expr.const `Nat [])) - (Lean.Expr.const `instLTNat [])) - (Lean.Expr.bvar 1)) - (Lean.Expr.lit (Lean.Literal.natVal 0))) - (Lean.BinderInfo.default)) - (Lean.BinderInfo.default)) - (Lean.Expr.app - (Lean.Expr.bvar 0) - (Lean.Expr.app - (Lean.Expr.app - (Lean.Expr.app - (Lean.Expr.const `Plausible.SampleableExt.interp [Lean.Level.succ (Lean.Level.zero), Lean.Level.zero]) - (Lean.Expr.const `Nat [])) - (Lean.Expr.const `Plausible.Nat.sampleableExt [])) - (Lean.Expr.app - (Lean.Expr.app - (Lean.Expr.app (Lean.Expr.const `OfNat.ofNat [Lean.Level.zero]) (Lean.Expr.const `Nat [])) - (Lean.Expr.lit (Lean.Literal.natVal 0))) - (Lean.Expr.app (Lean.Expr.const `instOfNatNat []) (Lean.Expr.lit (Lean.Literal.natVal 0)))))) - (Lean.BinderInfo.default))) - (Lean.Expr.app - (Lean.Expr.app - (Lean.Expr.app - (Lean.Expr.app - (Lean.Expr.const `mt []) - (Lean.Expr.forallE - `b - (Lean.Expr.const `Nat []) - (Lean.Expr.app - (Lean.Expr.app - (Lean.Expr.app - (Lean.Expr.app (Lean.Expr.const `LT.lt [Lean.Level.succ (Lean.Level.zero)]) (Lean.Expr.const `Nat [])) - (Lean.Expr.const `instLTNat [])) - (Lean.Expr.app - (Lean.Expr.app - (Lean.Expr.app - (Lean.Expr.const - `Plausible.SampleableExt.interp - [Lean.Level.succ (Lean.Level.zero), Lean.Level.zero]) - (Lean.Expr.const `Nat [])) - (Lean.Expr.const `Plausible.Nat.sampleableExt [])) - (Lean.Expr.app - (Lean.Expr.app - (Lean.Expr.app (Lean.Expr.const `OfNat.ofNat [Lean.Level.zero]) (Lean.Expr.const `Nat [])) - (Lean.Expr.lit (Lean.Literal.natVal 0))) - (Lean.Expr.app (Lean.Expr.const `instOfNatNat []) (Lean.Expr.lit (Lean.Literal.natVal 0)))))) - (Lean.Expr.app - (Lean.Expr.app - (Lean.Expr.app (Lean.Expr.const `OfNat.ofNat [Lean.Level.zero]) (Lean.Expr.const `Nat [])) - (Lean.Expr.lit (Lean.Literal.natVal 0))) - (Lean.Expr.app (Lean.Expr.const `instOfNatNat []) (Lean.Expr.lit (Lean.Literal.natVal 0))))) - (Lean.BinderInfo.default))) - (Lean.Expr.app - (Lean.Expr.app - (Lean.Expr.app - (Lean.Expr.app (Lean.Expr.const `LT.lt [Lean.Level.succ (Lean.Level.zero)]) (Lean.Expr.const `Nat [])) - (Lean.Expr.const `instLTNat [])) - (Lean.Expr.app - (Lean.Expr.app - (Lean.Expr.app - (Lean.Expr.const `Plausible.SampleableExt.interp [Lean.Level.succ (Lean.Level.zero), Lean.Level.zero]) - (Lean.Expr.const `Nat [])) - (Lean.Expr.const `Plausible.Nat.sampleableExt [])) - (Lean.Expr.app - (Lean.Expr.app - (Lean.Expr.app (Lean.Expr.const `OfNat.ofNat [Lean.Level.zero]) (Lean.Expr.const `Nat [])) - (Lean.Expr.lit (Lean.Literal.natVal 0))) - (Lean.Expr.app (Lean.Expr.const `instOfNatNat []) (Lean.Expr.lit (Lean.Literal.natVal 0)))))) - (Lean.Expr.app - (Lean.Expr.app - (Lean.Expr.app (Lean.Expr.const `OfNat.ofNat [Lean.Level.zero]) (Lean.Expr.const `Nat [])) - (Lean.Expr.lit (Lean.Literal.natVal 0))) - (Lean.Expr.app (Lean.Expr.const `instOfNatNat []) (Lean.Expr.lit (Lean.Literal.natVal 0)))))) - (Lean.Expr.lam - `h - (Lean.Expr.forallE - `b - (Lean.Expr.const `Nat []) - (Lean.Expr.app - (Lean.Expr.app - (Lean.Expr.app - (Lean.Expr.app (Lean.Expr.const `LT.lt [Lean.Level.succ (Lean.Level.zero)]) (Lean.Expr.const `Nat [])) - (Lean.Expr.const `instLTNat [])) - (Lean.Expr.app - (Lean.Expr.app - (Lean.Expr.app - (Lean.Expr.const - `Plausible.SampleableExt.interp - [Lean.Level.succ (Lean.Level.zero), Lean.Level.zero]) - (Lean.Expr.const `Nat [])) - (Lean.Expr.const `Plausible.Nat.sampleableExt [])) - (Lean.Expr.app - (Lean.Expr.app - (Lean.Expr.app (Lean.Expr.const `OfNat.ofNat [Lean.Level.zero]) (Lean.Expr.const `Nat [])) - (Lean.Expr.lit (Lean.Literal.natVal 0))) - (Lean.Expr.app (Lean.Expr.const `instOfNatNat []) (Lean.Expr.lit (Lean.Literal.natVal 0)))))) - (Lean.Expr.app - (Lean.Expr.app - (Lean.Expr.app (Lean.Expr.const `OfNat.ofNat [Lean.Level.zero]) (Lean.Expr.const `Nat [])) - (Lean.Expr.lit (Lean.Literal.natVal 0))) - (Lean.Expr.app (Lean.Expr.const `instOfNatNat []) (Lean.Expr.lit (Lean.Literal.natVal 0))))) - (Lean.BinderInfo.default)) - (Lean.Expr.app - (Lean.Expr.bvar 0) - (Lean.Expr.app - (Lean.Expr.app - (Lean.Expr.const `instNonemptyOfInhabited [Lean.Level.succ (Lean.Level.zero)]) - (Lean.Expr.const `Nat [])) - (Lean.Expr.const `instInhabitedNat []))) - (Lean.BinderInfo.default))) - (Lean.Expr.app - (Lean.Expr.app - (Lean.Expr.app - (Lean.Expr.app - (Lean.Expr.const `mt []) - (Lean.Expr.app - (Lean.Expr.app - (Lean.Expr.app - (Lean.Expr.app (Lean.Expr.const `LT.lt [Lean.Level.succ (Lean.Level.zero)]) (Lean.Expr.const `Nat [])) - (Lean.Expr.const `instLTNat [])) - (Lean.Expr.app - (Lean.Expr.app - (Lean.Expr.app - (Lean.Expr.const - `Plausible.SampleableExt.interp - [Lean.Level.succ (Lean.Level.zero), Lean.Level.zero]) - (Lean.Expr.const `Nat [])) - (Lean.Expr.const `Plausible.Nat.sampleableExt [])) - (Lean.Expr.app - (Lean.Expr.app - (Lean.Expr.app (Lean.Expr.const `OfNat.ofNat [Lean.Level.zero]) (Lean.Expr.const `Nat [])) - (Lean.Expr.lit (Lean.Literal.natVal 0))) - (Lean.Expr.app (Lean.Expr.const `instOfNatNat []) (Lean.Expr.lit (Lean.Literal.natVal 0)))))) - (Lean.Expr.app - (Lean.Expr.app - (Lean.Expr.app (Lean.Expr.const `OfNat.ofNat [Lean.Level.zero]) (Lean.Expr.const `Nat [])) - (Lean.Expr.lit (Lean.Literal.natVal 0))) - (Lean.Expr.app (Lean.Expr.const `instOfNatNat []) (Lean.Expr.lit (Lean.Literal.natVal 0)))))) - (Lean.Expr.app - (Lean.Expr.app - (Lean.Expr.app - (Lean.Expr.app (Lean.Expr.const `LT.lt [Lean.Level.succ (Lean.Level.zero)]) (Lean.Expr.const `Nat [])) - (Lean.Expr.const `instLTNat [])) - (Lean.Expr.app - (Lean.Expr.app - (Lean.Expr.app - (Lean.Expr.const - `Plausible.SampleableExt.interp - [Lean.Level.succ (Lean.Level.zero), Lean.Level.zero]) - (Lean.Expr.const `Nat [])) - (Lean.Expr.const `Plausible.Nat.sampleableExt [])) - (Lean.Expr.app - (Lean.Expr.app - (Lean.Expr.app (Lean.Expr.const `OfNat.ofNat [Lean.Level.zero]) (Lean.Expr.const `Nat [])) - (Lean.Expr.lit (Lean.Literal.natVal 0))) - (Lean.Expr.app (Lean.Expr.const `instOfNatNat []) (Lean.Expr.lit (Lean.Literal.natVal 0)))))) - (Lean.Expr.app - (Lean.Expr.app - (Lean.Expr.app (Lean.Expr.const `OfNat.ofNat [Lean.Level.zero]) (Lean.Expr.const `Nat [])) - (Lean.Expr.lit (Lean.Literal.natVal 0))) - (Lean.Expr.app (Lean.Expr.const `instOfNatNat []) (Lean.Expr.lit (Lean.Literal.natVal 0)))))) - (Lean.Expr.app - (Lean.Expr.const `id [Lean.Level.zero]) - (Lean.Expr.app - (Lean.Expr.app - (Lean.Expr.app - (Lean.Expr.app (Lean.Expr.const `LT.lt [Lean.Level.succ (Lean.Level.zero)]) (Lean.Expr.const `Nat [])) - (Lean.Expr.const `instLTNat [])) - (Lean.Expr.app - (Lean.Expr.app - (Lean.Expr.app - (Lean.Expr.const - `Plausible.SampleableExt.interp - [Lean.Level.succ (Lean.Level.zero), Lean.Level.zero]) - (Lean.Expr.const `Nat [])) - (Lean.Expr.const `Plausible.Nat.sampleableExt [])) - (Lean.Expr.app - (Lean.Expr.app - (Lean.Expr.app (Lean.Expr.const `OfNat.ofNat [Lean.Level.zero]) (Lean.Expr.const `Nat [])) - (Lean.Expr.lit (Lean.Literal.natVal 0))) - (Lean.Expr.app (Lean.Expr.const `instOfNatNat []) (Lean.Expr.lit (Lean.Literal.natVal 0)))))) - (Lean.Expr.app - (Lean.Expr.app - (Lean.Expr.app (Lean.Expr.const `OfNat.ofNat [Lean.Level.zero]) (Lean.Expr.const `Nat [])) - (Lean.Expr.lit (Lean.Literal.natVal 0))) - (Lean.Expr.app (Lean.Expr.const `instOfNatNat []) (Lean.Expr.lit (Lean.Literal.natVal 0))))))) - (Lean.Expr.app - (Lean.Expr.app - (Lean.Expr.app - (Lean.Expr.const `of_decide_eq_false []) - (Lean.Expr.app - (Lean.Expr.app - (Lean.Expr.app - (Lean.Expr.app (Lean.Expr.const `LT.lt [Lean.Level.succ (Lean.Level.zero)]) (Lean.Expr.const `Nat [])) - (Lean.Expr.const `instLTNat [])) - (Lean.Expr.app - (Lean.Expr.app - (Lean.Expr.app - (Lean.Expr.const - `Plausible.SampleableExt.interp - [Lean.Level.succ (Lean.Level.zero), Lean.Level.zero]) - (Lean.Expr.const `Nat [])) - (Lean.Expr.const `Plausible.Nat.sampleableExt [])) - (Lean.Expr.app - (Lean.Expr.app - (Lean.Expr.app (Lean.Expr.const `OfNat.ofNat [Lean.Level.zero]) (Lean.Expr.const `Nat [])) - (Lean.Expr.lit (Lean.Literal.natVal 0))) - (Lean.Expr.app (Lean.Expr.const `instOfNatNat []) (Lean.Expr.lit (Lean.Literal.natVal 0)))))) - (Lean.Expr.app - (Lean.Expr.app - (Lean.Expr.app (Lean.Expr.const `OfNat.ofNat [Lean.Level.zero]) (Lean.Expr.const `Nat [])) - (Lean.Expr.lit (Lean.Literal.natVal 0))) - (Lean.Expr.app (Lean.Expr.const `instOfNatNat []) (Lean.Expr.lit (Lean.Literal.natVal 0)))))) - (Lean.Expr.app - (Lean.Expr.app - (Lean.Expr.const `Nat.decLt []) - (Lean.Expr.app - (Lean.Expr.app - (Lean.Expr.app - (Lean.Expr.const - `Plausible.SampleableExt.interp - [Lean.Level.succ (Lean.Level.zero), Lean.Level.zero]) - (Lean.Expr.const `Nat [])) - (Lean.Expr.const `Plausible.Nat.sampleableExt [])) - (Lean.Expr.app - (Lean.Expr.app - (Lean.Expr.app (Lean.Expr.const `OfNat.ofNat [Lean.Level.zero]) (Lean.Expr.const `Nat [])) - (Lean.Expr.lit (Lean.Literal.natVal 0))) - (Lean.Expr.app (Lean.Expr.const `instOfNatNat []) (Lean.Expr.lit (Lean.Literal.natVal 0)))))) - (Lean.Expr.app - (Lean.Expr.app - (Lean.Expr.app (Lean.Expr.const `OfNat.ofNat [Lean.Level.zero]) (Lean.Expr.const `Nat [])) - (Lean.Expr.lit (Lean.Literal.natVal 0))) - (Lean.Expr.app (Lean.Expr.const `instOfNatNat []) (Lean.Expr.lit (Lean.Literal.natVal 0)))))) - (Lean.Expr.app - (Lean.Expr.app (Lean.Expr.const `Eq.refl [Lean.Level.succ (Lean.Level.zero)]) (Lean.Expr.const `Bool [])) - (Lean.Expr.const `Bool.false [])))))) --/ -#guard_msgs in #eval MetaTestable.check (∀ (a _b : Nat), a < 0) (propExpr := (Lean.Expr.forallE `a (Lean.Expr.const `Nat []) From b8b8453a7fbc5dbe4433c971a67307ae47bfebe1 Mon Sep 17 00:00:00 2001 From: Siddhartha Gadgil Date: Tue, 26 Nov 2024 10:40:04 +0530 Subject: [PATCH 23/30] propExpr optional --- Plausible/MetaTestable.lean | 267 ++++++++++++++++++++++++------------ 1 file changed, 180 insertions(+), 87 deletions(-) diff --git a/Plausible/MetaTestable.lean b/Plausible/MetaTestable.lean index 18ea244..248116c 100644 --- a/Plausible/MetaTestable.lean +++ b/Plausible/MetaTestable.lean @@ -210,7 +210,7 @@ inductive MetaTestResult (p : Prop) where /-- `MetaTestable p` uses random examples to try to disprove `p`. -/ class MetaTestable (p : Prop) where - runExpr (cfg : Configuration) (minimize : Bool) (propExpr : Expr) : MGen (MetaTestResult p) + runExpr (cfg : Configuration) (minimize : Bool) (propExpr? : Option Expr) : MGen (MetaTestResult p) namespace MetaTestResult @@ -235,23 +235,30 @@ def checkDisproof (pf prop: Expr) : MetaM Unit := do throwError m!"Expected a proof of {negProp}, got proof of {← ppExpr pfType}" /-- Combine the test result for properties `p` and `q` to create a test for their conjunction. -/ -def and : MetaTestResult p → MetaTestResult q → Expr → MetaM (MetaTestResult (p ∧ q)) - | failure h pf? xs n, _, e => do - let (some e₁, some e₂) ← andProp? e | throwError m!"Expected an `And` proposition, got {← ppExpr e}" +def and : MetaTestResult p → MetaTestResult q → Option Expr → MetaM (MetaTestResult (p ∧ q)) + | failure h pf? xs n, _, e? => do let pf'? ← - pf?.mapM fun pf => withLocalDeclD `h e fun h => do - let x ← mkAppOptM ``And.left #[e₁, e₂, h] - let e' ← mkAppM' pf #[x] - mkLambdaFVars #[h] e' - -- checkDisproof pf' e + match e? with + | none => pure none + | some e => + let (some e₁, some e₂) ← andProp? e | throwError m!"Expected an `And` proposition, got {← ppExpr e}" + pf?.mapM fun pf => withLocalDeclD `h e fun h => do + let x ← mkAppOptM ``And.left #[e₁, e₂, h] + let e' ← mkAppM' pf #[x] + mkLambdaFVars #[h] e' + -- checkDisproof pf' e return failure (fun h2 => h h2.left) pf'? xs n - | _, failure h pf? xs n, e => do - let (some e₁, some e₂) ← andProp? e | throwError m!"Expected an `And` proposition, got {← ppExpr e}" - let pf' ← pf?.mapM fun pf => withLocalDeclD `h e fun h => do - let x ← mkAppOptM ``And.right #[e₁, e₂, h] - let e' ← mkAppM' pf #[x] - mkLambdaFVars #[h] e' - -- checkDisproof pf' e + | _, failure h pf? xs n, e? => do + let pf' ← + match e? with + | none => pure none + | some e => + let (some e₁, some e₂) ← andProp? e | throwError m!"Expected an `And` proposition, got {← ppExpr e}" + pf?.mapM fun pf => withLocalDeclD `h e fun h => do + let x ← mkAppOptM ``And.right #[e₁, e₂, h] + let e' ← mkAppM' pf #[x] + mkLambdaFVars #[h] e' + -- checkDisproof pf' e return failure (fun h2 => h h2.right) pf' xs n | success h1, success h2, _ => return success <| combine (combine (PSum.inr And.intro) h1) h2 @@ -261,17 +268,21 @@ def and : MetaTestResult p → MetaTestResult q → Expr → MetaM (MetaTestRes /-- Combine the test result for properties `p` and `q` to create a test for their disjunction. -/ -def or : MetaTestResult p → MetaTestResult q → Expr → MetaM (MetaTestResult (p ∨ q)) - | failure h1 pf1 xs n, failure h2 pf2 ys m, e => do - let (some α, some β) ← orProp? e | throwError m!"Expected an `Or` proposition, got {← ppExpr e}" - let motive ← withLocalDeclD `h e fun h => do - mkLambdaFVars #[h] <| mkConst ``False - let pf ← mkAppOptM ``Or.rec #[α, β, motive, pf1, pf2] +def or : MetaTestResult p → MetaTestResult q → Option Expr → MetaM (MetaTestResult (p ∨ q)) + | failure h1 pf1 xs n, failure h2 pf2 ys m, e? => do + let pf ← + match e? with + | none => pure none + | some e => + let (some α, some β) ← orProp? e | throwError m!"Expected an `Or` proposition, got {← ppExpr e}" + let motive ← withLocalDeclD `h e fun h => do + mkLambdaFVars #[h] <| mkConst ``False + mkAppOptM ``Or.rec #[α, β, motive, pf1, pf2] let h3 := fun h => match h with | Or.inl h3 => h1 h3 | Or.inr h3 => h2 h3 - checkDisproof pf e + -- checkDisproof pf e return failure h3 pf (xs ++ ys) (n + m) | success h, _, _ => return success <| combine (PSum.inr Or.inl) h | _, success h, _ => return success <| combine (PSum.inr Or.inr) h @@ -281,35 +292,36 @@ def or : MetaTestResult p → MetaTestResult q → Expr → MetaM (MetaTestResu /-- If `q → p`, then `¬ p → ¬ q` which means that testing `p` can allow us to find counter-examples to `q`. -/ -def imp (h : q → p) (hExp: Expr) (r : MetaTestResult p) +def imp (h : q → p) (hExp?: Option Expr) (r : MetaTestResult p) (p : Unit ⊕' (p → q) := PSum.inl ()) : MetaM (MetaTestResult q) := match r with | failure h2 pf? xs n => do - let pf' ← pf?.mapM fun pf => mkAppM ``mt #[hExp, pf] + let pf' ← + hExp?.bindM fun hExp => pf?.mapM fun pf => mkAppM ``mt #[hExp, pf] return failure (mt h h2) pf' xs n | success h2 => return success <| combine p h2 | gaveUp n => return gaveUp n /-- Test `q` by testing `p` and proving the equivalence between the two. -/ -def iff (h : q ↔ p) (hExp: Expr) (r : MetaTestResult p) : MetaM (MetaTestResult q) := do - let hExp' ← mkAppM ``Iff.mp #[hExp] +def iff (h : q ↔ p) (hExp? : Option Expr) (r : MetaTestResult p) : MetaM (MetaTestResult q) := do + let hExp' ← hExp?.mapM fun hExp => mkAppM ``Iff.mp #[hExp] imp h.mp hExp' r (PSum.inr h.mpr) /-- When we assign a value to a universally quantified variable, we record that value using this function so that our counter-examples can be informative. -/ -def addInfo (x : String) (h : q → p) (hExp: Expr) (r : MetaTestResult p) +def addInfo (x : String) (h : q → p) (hExp?: Option Expr) (r : MetaTestResult p) (p : Unit ⊕' (p → q) := PSum.inl ()) : (MetaM <| MetaTestResult q) := do if let failure h2 pf? xs n := r then - let pf' ← pf?.mapM fun pf => mkAppM ``mt #[hExp, pf] + let pf' ← hExp?.bindM fun hExp => pf?.mapM fun pf => mkAppM ``mt #[hExp, pf] return failure (mt h h2) pf' (x :: xs) n else - imp h hExp r p + imp h hExp? r p /-- Add some formatting to the information recorded by `addInfo`. -/ -def addVarInfo {γ : Type _} [Repr γ] (var : String) (x : γ) (h : q → p) (hExp: Expr) (r : MetaTestResult p) +def addVarInfo {γ : Type _} [Repr γ] (var : String) (x : γ) (h : q → p) (hExp: Option Expr) (r : MetaTestResult p) (p : Unit ⊕' (p → q) := PSum.inl ()) : MetaM (MetaTestResult q) := do - addInfo s!"{var} := {repr x}" h (hExp: Expr) r p + addInfo s!"{var} := {repr x}" h (hExp: Option Expr) r p def isFailure : MetaTestResult p → Bool | failure .. => true @@ -325,23 +337,32 @@ open MetaTestResult abbrev ProxyExpr α [SampleableExt α] := ToExpr (SampleableExt.proxy α) -def runPropExpr (p : Prop) [MetaTestable p] : Configuration → Bool → Expr → MGen (MetaTestResult p) := fun cfg b e => do +def runPropExpr (p : Prop) [MetaTestable p] : Configuration → Bool → Option Expr → MGen (MetaTestResult p) := fun cfg b e => do runExpr cfg b e /-- A `dbgTrace` with special formatting -/ def slimTrace {m : Type → Type _} [Pure m] (s : String) : m PUnit := dbgTrace s!"[Plausible: {s}]" (fun _ => pure ()) + instance andTestable [MetaTestable p] [MetaTestable q] : MetaTestable (p ∧ q) where - runExpr := fun cfg min e => do - let (some e₁, some e₂) ← andProp? e | throwError m!"Expected an `And` proposition, got {← ppExpr e}" + runExpr := fun cfg min e? => do + let (e₁, e₂) ← match e? with + | some e => do + let (some e₁, some e₂) ← andProp? e | throwError m!"Expected an `And` proposition, got {← ppExpr e}" + pure (some e₁, some e₂) + | none => pure (none, none) let xp ← runPropExpr p cfg min e₁ let xq ← runPropExpr q cfg min e₂ - and xp xq e + and xp xq e? instance orTestable [MetaTestable p] [MetaTestable q] : MetaTestable (p ∨ q) where - runExpr := fun cfg min e => do - let (some e₁, some e₂) ← orProp? e | throwError m!"Expected an `Or` proposition, got {← ppExpr e}" + runExpr := fun cfg min e? => do + let (e₁, e₂) ← match e? with + | some e => do + let (some e₁, some e₂) ← orProp? e | throwError m!"Expected an `Or` proposition, got {← ppExpr e}" + pure (some e₁, some e₂) + | none => pure (none, none) let xp ← runPropExpr p cfg min e₁ -- As a little performance optimization we can just not run the second -- test if the first succeeds @@ -350,7 +371,7 @@ instance orTestable [MetaTestable p] [MetaTestable q] : MetaTestable (p ∨ q) w | success (PSum.inr h) => return success (PSum.inr <| Or.inl h) | _ => let xq ← runPropExpr q cfg min e₂ - or xp xq e + or xp xq e? theorem iff_resolve (p q : Prop) : (p ↔ q) ↔ p ∧ q ∨ ¬p ∧ ¬q := by constructor @@ -360,28 +381,39 @@ theorem iff_resolve (p q : Prop) : (p ↔ q) ↔ p ∧ q ∨ ¬p ∧ ¬q := by rcases h with ⟨hleft, hright⟩ | ⟨hleft, hright⟩ <;> simp [hleft, hright] instance iffTestable [MetaTestable ((p ∧ q) ∨ (¬ p ∧ ¬ q))] : MetaTestable (p ↔ q) where - runExpr := fun cfg min e => do - let some (α, β) ← iffProp? e | throwError m!"Expected an `Iff` proposition, got {← ppExpr e}" - let h ← runPropExpr ((p ∧ q) ∨ (¬ p ∧ ¬ q)) cfg min e - let hExp ← mkAppM ``iff_resolve #[α, β] + runExpr := fun cfg min e? => do + let hExp ← + match e? with + | none => pure none + | some e => do + let some (α, β) ← iffProp? e | throwError m!"Expected an `Iff` proposition, got {← ppExpr e}" + mkAppM ``iff_resolve #[α, β] + let h ← runPropExpr ((p ∧ q) ∨ (¬ p ∧ ¬ q)) cfg min e? iff (iff_resolve p q) hExp h variable {var : String} instance decGuardTestable [PrintableProp p] [Decidable p] {β : p → Prop} [∀ h, MetaTestable (β h)] : MetaTestable (NamedBinder var <| ∀ h, β h) where - runExpr := fun cfg min e => do + runExpr := fun cfg min e? => do if h : p then + match e? with + | none => + let res := runPropExpr (β h) cfg min none + let s := printProp p + ← (fun r => addInfo s!"guard: {s}" (· <| h) none r (PSum.inr <| fun q _ => q)) <$> res | some e => let (some βExp, some pExp) ← forallProp? e | throwError m!"Expected a `Forall` proposition, got {← ppExpr e}" - let yExp ← mkAppM' βExp #[pExp] - let res := runPropExpr (β h) cfg min yExp let h' := (· <| h) + let yExp ← + mkAppM' βExp #[pExp] + let res := runPropExpr (β h) cfg min yExp let decInstType ← mkAppM ``Decidable #[pExp] let inst ← synthInstance decInstType let falseRefl ← mkAppM ``Eq.refl #[mkConst ``false] let pf ← mkAppOptM ``of_decide_eq_true #[pExp, inst, falseRefl] let cod := mkApp βExp pf - let hExp ← withLocalDeclD `x (← mkArrow pExp cod) fun x => do + let hExp ← + withLocalDeclD `x (← mkArrow pExp cod) fun x => do let y ← mkAppM' x #[cod] mkLambdaFVars #[x] y let s := printProp p @@ -425,36 +457,49 @@ candidate that falsifies a property and recursively shrinking that one. The process is guaranteed to terminate because `shrink x` produces a proof that all the values it produces are smaller (according to `SizeOf`) than `x`. -/ -partial def minimizeAux [SampleableExt α] [ProxyExpr α] {β : α → Prop} [∀ x, MetaTestable (β x)] (αExp βExpr: Expr) (cfg : Configuration) +partial def minimizeAux [SampleableExt α] [ProxyExpr α] {β : α → Prop} [∀ x, MetaTestable (β x)] (αExp? βExp?: Option Expr) (cfg : Configuration) (var : String) (x : SampleableExt.proxy α) (n : Nat) : OptionT MGen (Σ x, MetaTestResult (β (SampleableExt.interp x))) := do let candidates := SampleableExt.shrink.shrink x if cfg.traceShrinkCandidates then slimTrace s!"Candidates for {var} := {repr x}:\n {repr candidates}" - for candidate in candidates do - if cfg.traceShrinkCandidates then - slimTrace s!"Trying {var} := {repr candidate}" - let xExpr := toExpr candidate - let .sort u := ← inferType αExp | throwError m!"Expected a sort, got {αExp}" - let v ← mkFreshLevelMVar - let instType := mkApp (mkConst ``SampleableExt [u, v]) αExp - let samp ← synthInstance instType - let xInterp ← mkAppOptM ``SampleableExt.interp #[αExp, samp, xExpr] - let e' ← mkAppM' βExpr #[xInterp] - let res ← OptionT.lift <| runPropExpr (β (SampleableExt.interp candidate)) cfg true e' - if res.isFailure then - if cfg.traceShrink then - slimTrace s!"{var} shrunk to {repr candidate} from {repr x}" - let currentStep := OptionT.lift <| return Sigma.mk candidate (addShrinks (n + 1) res) - let nextStep := minimizeAux αExp βExpr cfg var candidate (n + 1) - return ← (nextStep <|> currentStep) + match αExp?, βExp? with + | some αExp, some βExpr => + for candidate in candidates do + if cfg.traceShrinkCandidates then + slimTrace s!"Trying {var} := {repr candidate}" + let xExpr := toExpr candidate + let .sort u := ← inferType αExp | throwError m!"Expected a sort, got {αExp}" + let v ← mkFreshLevelMVar + let instType := mkApp (mkConst ``SampleableExt [u, v]) αExp + let samp ← synthInstance instType + let xInterp ← mkAppOptM ``SampleableExt.interp #[αExp, samp, xExpr] + let e' ← mkAppM' βExpr #[xInterp] + let res ← OptionT.lift <| runPropExpr (β (SampleableExt.interp candidate)) cfg true e' + if res.isFailure then + if cfg.traceShrink then + slimTrace s!"{var} shrunk to {repr candidate} from {repr x}" + let currentStep := OptionT.lift <| return Sigma.mk candidate (addShrinks (n + 1) res) + let nextStep := minimizeAux αExp βExpr cfg var candidate (n + 1) + return ← (nextStep <|> currentStep) + | _, _ => + for candidate in candidates do + if cfg.traceShrinkCandidates then + slimTrace s!"Trying {var} := {repr candidate}" + let res ← OptionT.lift <| runPropExpr (β (SampleableExt.interp candidate)) cfg true none + if res.isFailure then + if cfg.traceShrink then + slimTrace s!"{var} shrunk to {repr candidate} from {repr x}" + let currentStep := OptionT.lift <| return Sigma.mk candidate (addShrinks (n + 1) res) + let nextStep := minimizeAux none none cfg var candidate (n + 1) + return ← (nextStep <|> currentStep) if cfg.traceShrink then slimTrace s!"No shrinking possible for {var} := {repr x}" failure /-- Once a property fails to hold on an example, look for smaller counter-examples to show the user. -/ -def minimize [SampleableExt α] [ProxyExpr α] {β : α → Prop} [∀ x, MetaTestable (β x)] (αExp βExpr: Expr)(cfg : Configuration) +def minimize [SampleableExt α] [ProxyExpr α] {β : α → Prop} [∀ x, MetaTestable (β x)] (αExp βExpr: Option Expr)(cfg : Configuration) (var : String) (x : SampleableExt.proxy α) (r : MetaTestResult (β <| SampleableExt.interp x)) : MGen (Σ x, MetaTestResult (β <| SampleableExt.interp x)) := do if cfg.traceShrink then @@ -468,7 +513,28 @@ open Lean Meta Elab Term Tactic in bound variable with it. -/ instance varTestable [SampleableExt α] [ProxyExpr α] {β : α → Prop} [∀ x, MetaTestable (β x)] : MetaTestable (NamedBinder var <| ∀ x : α, β x) where - runExpr := fun cfg min e => do + runExpr := fun cfg min e? => do + match e? with + | none => + let x ← SampleableExt.sample + if cfg.traceSuccesses || cfg.traceDiscarded then + slimTrace s!"{var} := {repr x}" + let r ← runPropExpr (β <| SampleableExt.interp x) cfg false none + let ⟨finalX, finalR⟩ ← + if isFailure r then + if cfg.traceSuccesses then + slimTrace s!"{var} := {repr x} is a failure" + if min then + minimize none none cfg var x r + else + pure ⟨x, r⟩ + else + pure ⟨x, r⟩ + let xExpr := toExpr finalX + let h := (· <| SampleableExt.interp finalX) + addVarInfo var finalX h none finalR + + | some e => let (some βExp, some αExp) ← forallProp? e | throwError m!"Expected a `Forall` proposition, got {← ppExpr e}" let x ← SampleableExt.sample let xExpr := toExpr x @@ -534,27 +600,39 @@ theorem bool_to_prop_fmly (β : Prop → Prop): NamedBinder var (∀ (p : Prop), instance propVarTestable {β : Prop → Prop} [h: ∀ b : Bool, MetaTestable (β b)] : MetaTestable (NamedBinder var <| ∀ p : Prop, β p) where - runExpr := fun cfg min e => do - let (some βExpr, _) ← forallProp? e | throwError m!"Expected a `Forall` proposition, got {← ppExpr e}" - let p ← runPropExpr (NamedBinder var <| ∀ b : Bool, β b) cfg min e - let e' ← mkAppM ``bool_to_prop_fmly #[βExpr] + runExpr := fun cfg min e? => do + let p ← runPropExpr (NamedBinder var <| ∀ b : Bool, β b) cfg min e? + let e' ← + match e? with + | none => pure none + | some e => do + let (some βExpr, _) ← forallProp? e | throwError m!"Expected a `Forall` proposition, got {← ppExpr e}" + let res ← mkAppM ``bool_to_prop_fmly #[βExpr] + pure (some res) imp (bool_to_prop_fmly β) e' p instance (priority := high) unusedVarTestable {β : Prop} [Nonempty α] [MetaTestable β] : MetaTestable (NamedBinder var (α → β)) where - runExpr := fun cfg min e => do + runExpr := fun cfg min e? => do if cfg.traceDiscarded || cfg.traceSuccesses then slimTrace s!"{var} is unused" - let (some aExp, some e') ← impProp? e | throwError m!"Expected an `Imp` proposition, got {← ppExpr e}" - let r ← runPropExpr β cfg min e' - let hExp ← mkAppOptM ``id #[e'] - let finalR ← addInfo s!"{var} is irrelevant (unused)" id hExp r - let h := (· <| Classical.ofNonempty) - let nInst ← synthInstance <| ← mkAppM ``Nonempty #[aExp] - let hExp ← withLocalDeclD `h e fun h => do - mkLambdaFVars #[h] <| mkApp h nInst - imp h hExp finalR (PSum.inr <| fun x _ => x) + match e? with + | none => + let r ← runPropExpr β cfg min none + let finalR ← addInfo s!"{var} is irrelevant (unused)" id none r + let h := (· <| Classical.ofNonempty) + imp h none finalR (PSum.inr <| fun x _ => x) + | some e => + let (some aExp, some e') ← impProp? e | throwError m!"Expected an `Imp` proposition, got {← ppExpr e}" + let r ← runPropExpr β cfg min e' + let hExp ← mkAppOptM ``id #[e'] + let finalR ← addInfo s!"{var} is irrelevant (unused)" id hExp r + let nInst ← synthInstance <| ← mkAppM ``Nonempty #[aExp] + let hExp ← withLocalDeclD `h e fun h => do + mkLambdaFVars #[h] <| mkApp h nInst + let h := (· <| Classical.ofNonempty) + imp h hExp finalR (PSum.inr <| fun x _ => x) theorem prop_iff_subtype (p : α → Prop) (β : α → Prop) : NamedBinder var (∀ (x : α), NamedBinder var' (p x → β x)) ↔ ∀ (x : Subtype p), β x.val := by simp [Subtype.forall, NamedBinder] @@ -564,7 +642,18 @@ instance (priority := 2000) subtypeVarTestable {p : α → Prop} {β : α → Pr [∀ x, MetaTestable (β x)][ToExpr α] [SampleableExt (Subtype p)] [ProxyExpr (Subtype p)] {var'} : MetaTestable (NamedBinder var <| (x : α) → NamedBinder var' <| p x → β x) where - runExpr cfg min e := do + runExpr cfg min e? := + match e? with + | none => do + letI (x : Subtype p) : MetaTestable (β x) := + { runExpr := fun cfg min e => do + let r ← runPropExpr (β x.val) cfg min none + let idExp ← mkAppOptM ``id #[e] + addInfo s!"guard: {printProp (p x)} (by construction)" id idExp r (PSum.inr id) } + do + let r ← @runExpr (∀ x : Subtype p, β x.val) (@varTestable var _ _ _ _ _) cfg min none + iff (prop_iff_subtype p β) none r + | some e => do let (some pExp, some βExp, some _) ← forallPropProp? e | throwError m!"Expected a `Forall` proposition with arrow, got {← ppExpr e}" let subType ← mkAppM ``Subtype #[pExp] letI (x : Subtype p) : MetaTestable (β x) := @@ -585,15 +674,19 @@ instance (priority := 2000) subtypeVarTestable {p : α → Prop} {β : α → Pr instance (priority := low) decidableTestable {p : Prop} [PrintableProp p] [Decidable p] : MetaTestable p where - runExpr := fun _ _ e _ => do + runExpr := fun _ _ e? _ => do if h : p then return success (PSum.inr h) else let s := printProp p - let inst ← synthInstance <| ← mkAppM ``Decidable #[e] - let falseRefl ← mkAppM ``Eq.refl #[mkConst ``false] - let pf' ← mkAppOptM ``of_decide_eq_false #[e, inst, falseRefl] - checkDisproof pf' e + let pf' ← do + match e? with + | none => pure none + | some e => + let inst ← synthInstance <| ← mkAppM ``Decidable #[e] + let falseRefl ← mkAppM ``Eq.refl #[mkConst ``false] + mkAppOptM ``of_decide_eq_false #[e, inst, falseRefl] + -- checkDisproof pf' e -- logInfo "Decidable proof checked" return failure h pf' [s!"issue: {s} does not hold"] 0 From cd513bdcc84cf45ba7e588a10cb54d0ef1d8fbc2 Mon Sep 17 00:00:00 2001 From: Siddhartha Gadgil Date: Tue, 26 Nov 2024 15:16:17 +0530 Subject: [PATCH 24/30] removed requirement for `ToExpr` in MetaTestable --- Plausible/Functions.lean | 25 ++++++---- Plausible/MetaTestable.lean | 93 +++++++++++++++---------------------- Plausible/Sampleable.lean | 60 ++++++++++++++++++++++-- Test/Testable.lean | 2 +- 4 files changed, 111 insertions(+), 69 deletions(-) diff --git a/Plausible/Functions.lean b/Plausible/Functions.lean index 882dbc0..f74f98c 100644 --- a/Plausible/Functions.lean +++ b/Plausible/Functions.lean @@ -126,16 +126,19 @@ where variable [Repr α] -instance Pi.sampleableExt : SampleableExt (α → β) where - proxy := TotalFunction α (SampleableExt.proxy β) - interp f := SampleableExt.interp ∘ f.apply - sample := do - let xs : List (_ × _) ← (SampleableExt.sample (α := List (α × β))) - let ⟨x⟩ ← Gen.up <| (SampleableExt.sample : Gen (SampleableExt.proxy β)) - pure <| TotalFunction.withDefault (List.toFinmap' <| xs.map <| - Prod.map SampleableExt.interp id) x - -- note: no way of shrinking the domain without an inverse to `interp` - shrink := { shrink := letI : Shrinkable α := {}; TotalFunction.shrink } +-- instance Pi.sampleableExt : SampleableExt (α → β) where +-- proxy := TotalFunction α (SampleableExt.proxy β) +-- interp f := SampleableExt.interp ∘ f.apply +-- sample := do +-- let xs : List (_ × _) ← (SampleableExt.sample (α := List (α × β))) +-- let ys := (SampleableExt.sample : Gen (SampleableExt.proxy β)) +-- let ys' := Gen.up ys +-- let ⟨x⟩ ← ys' +-- pure <| TotalFunction.withDefault (List.toFinmap' <| xs.map <| +-- Prod.map SampleableExt.interp id) x +-- -- note: no way of shrinking the domain without an inverse to `interp` +-- shrink := { shrink := letI : Shrinkable α := {}; TotalFunction.shrink } +-- proxyExpr? := none end @@ -149,6 +152,7 @@ instance (priority := 2000) PiPred.sampleableExt [SampleableExt (α → Bool)] : interp m x := interp m x sample := sample shrink := SampleableExt.shrink + proxyExpr? := none instance (priority := 2000) PiUncurry.sampleableExt [SampleableExt (α × β → γ)] : SampleableExt.{imax (u + 1) (v + 1) w} (α → β → γ) where @@ -156,6 +160,7 @@ instance (priority := 2000) PiUncurry.sampleableExt [SampleableExt (α × β → interp m x y := interp m (x, y) sample := sample shrink := SampleableExt.shrink + proxyExpr? := none end SampleableExt diff --git a/Plausible/MetaTestable.lean b/Plausible/MetaTestable.lean index 248116c..b50aab2 100644 --- a/Plausible/MetaTestable.lean +++ b/Plausible/MetaTestable.lean @@ -334,8 +334,6 @@ namespace MetaTestable open MetaTestResult -abbrev ProxyExpr α [SampleableExt α] := ToExpr (SampleableExt.proxy α) - def runPropExpr (p : Prop) [MetaTestable p] : Configuration → Bool → Option Expr → MGen (MetaTestResult p) := fun cfg b e => do runExpr cfg b e @@ -457,14 +455,14 @@ candidate that falsifies a property and recursively shrinking that one. The process is guaranteed to terminate because `shrink x` produces a proof that all the values it produces are smaller (according to `SizeOf`) than `x`. -/ -partial def minimizeAux [SampleableExt α] [ProxyExpr α] {β : α → Prop} [∀ x, MetaTestable (β x)] (αExp? βExp?: Option Expr) (cfg : Configuration) +partial def minimizeAux [SampleableExt α] {β : α → Prop} [∀ x, MetaTestable (β x)] (αExp? βExp?: Option Expr) (cfg : Configuration) (var : String) (x : SampleableExt.proxy α) (n : Nat) : OptionT MGen (Σ x, MetaTestResult (β (SampleableExt.interp x))) := do let candidates := SampleableExt.shrink.shrink x if cfg.traceShrinkCandidates then slimTrace s!"Candidates for {var} := {repr x}:\n {repr candidates}" - match αExp?, βExp? with - | some αExp, some βExpr => + match αExp?, βExp?, getProxyExpr? α with + | some αExp, some βExpr, some _ => for candidate in candidates do if cfg.traceShrinkCandidates then slimTrace s!"Trying {var} := {repr candidate}" @@ -482,7 +480,7 @@ partial def minimizeAux [SampleableExt α] [ProxyExpr α] {β : α → Prop} [ let currentStep := OptionT.lift <| return Sigma.mk candidate (addShrinks (n + 1) res) let nextStep := minimizeAux αExp βExpr cfg var candidate (n + 1) return ← (nextStep <|> currentStep) - | _, _ => + | _, _, _ => for candidate in candidates do if cfg.traceShrinkCandidates then slimTrace s!"Trying {var} := {repr candidate}" @@ -499,7 +497,7 @@ partial def minimizeAux [SampleableExt α] [ProxyExpr α] {β : α → Prop} [ /-- Once a property fails to hold on an example, look for smaller counter-examples to show the user. -/ -def minimize [SampleableExt α] [ProxyExpr α] {β : α → Prop} [∀ x, MetaTestable (β x)] (αExp βExpr: Option Expr)(cfg : Configuration) +def minimize [SampleableExt α] {β : α → Prop} [∀ x, MetaTestable (β x)] (αExp βExpr: Option Expr)(cfg : Configuration) (var : String) (x : SampleableExt.proxy α) (r : MetaTestResult (β <| SampleableExt.interp x)) : MGen (Σ x, MetaTestResult (β <| SampleableExt.interp x)) := do if cfg.traceShrink then @@ -511,7 +509,7 @@ def minimize [SampleableExt α] [ProxyExpr α] {β : α → Prop} [∀ x, MetaTe open Lean Meta Elab Term Tactic in /-- Test a universal property by creating a sample of the right type and instantiating the bound variable with it. -/ -instance varTestable [SampleableExt α] [ProxyExpr α] {β : α → Prop} [∀ x, MetaTestable (β x)] : +instance varTestable [SampleableExt α] {β : α → Prop} [∀ x, MetaTestable (β x)] : MetaTestable (NamedBinder var <| ∀ x : α, β x) where runExpr := fun cfg min e? => do match e? with @@ -530,25 +528,27 @@ instance varTestable [SampleableExt α] [ProxyExpr α] {β : α → Prop} [∀ x pure ⟨x, r⟩ else pure ⟨x, r⟩ - let xExpr := toExpr finalX let h := (· <| SampleableExt.interp finalX) addVarInfo var finalX h none finalR - | some e => let (some βExp, some αExp) ← forallProp? e | throwError m!"Expected a `Forall` proposition, got {← ppExpr e}" let x ← SampleableExt.sample - let xExpr := toExpr x - let αExp ← instantiateMVars αExp - let .sort u := ← inferType αExp | throwError m!"Expected a sort, got {αExp}" - let v ← mkFreshLevelMVar - let instType := mkApp (mkConst ``SampleableExt [u, v]) αExp - let samp ← synthInstance instType - let xInterp ← mkAppOptM ``SampleableExt.interp #[αExp, samp, xExpr] - let e' ← mkAppM' βExp #[xInterp] - let (e', _) ← dsimp e' {} + let e'? ← match getProxyExpr? α with + | none => pure none + | some inst => + let xExpr := toExpr x + let αExp ← instantiateMVars αExp + let .sort u := ← inferType αExp | throwError m!"Expected a sort, got {αExp}" + let v ← mkFreshLevelMVar + let instType := mkApp (mkConst ``SampleableExt [u, v]) αExp + let samp ← synthInstance instType + let xInterp ← mkAppOptM ``SampleableExt.interp #[αExp, samp, xExpr] + let e' ← mkAppM' βExp #[xInterp] + let (e', _) ← dsimp e' {} + pure (some e') if cfg.traceSuccesses || cfg.traceDiscarded then slimTrace s!"{var} := {repr x}" - let r ← runPropExpr (β <| SampleableExt.interp x) cfg false e' + let r ← runPropExpr (β <| SampleableExt.interp x) cfg false e'? let ⟨finalX, finalR⟩ ← if isFailure r then if cfg.traceSuccesses then @@ -559,39 +559,22 @@ instance varTestable [SampleableExt α] [ProxyExpr α] {β : α → Prop} [∀ x pure ⟨x, r⟩ else pure ⟨x, r⟩ - let xExpr := toExpr finalX - let .sort u ← inferType αExp | throwError m!"Expected a sort, got {αExp}" - let v ← mkFreshLevelMVar - let instType := mkApp (mkConst ``SampleableExt [u, v]) αExp - let samp ← synthInstance instType - let xInterp ← mkAppOptM ``SampleableExt.interp #[αExp, samp, xExpr] + let hExpr? ← match getProxyExpr? α with + | none => pure none + | some inst => + let xExpr := toExpr finalX + let .sort u ← inferType αExp | throwError m!"Expected a sort, got {αExp}" + let v ← mkFreshLevelMVar + let instType := mkApp (mkConst ``SampleableExt [u, v]) αExp + let samp ← synthInstance instType + let xInterp ← mkAppOptM ``SampleableExt.interp #[αExp, samp, xExpr] + let hExpr ← withLocalDeclD `x e fun x => do + mkLambdaFVars #[x] (mkApp x xInterp) + pure (some hExpr) let h := (· <| SampleableExt.interp finalX) - let hExpr ← withLocalDeclD `x e fun x => do - mkLambdaFVars #[x] (mkApp x xInterp) - addVarInfo var finalX h hExpr finalR - - --- Typeclass inference does not seem to work across `mkSelfContained`, so we need to provide instances for the basic types -instance : ProxyExpr Bool := (inferInstance : ToExpr Bool) -instance : ProxyExpr String := (inferInstance : ToExpr String) -instance : ProxyExpr Nat := (inferInstance : ToExpr Nat) -instance : ProxyExpr Int := (inferInstance : ToExpr Int) -instance : ProxyExpr Char := (inferInstance : ToExpr Char) -instance : ProxyExpr Unit := (inferInstance : ToExpr Unit) -instance : ProxyExpr USize := (inferInstance : ToExpr USize) -instance : ProxyExpr UInt32 := (inferInstance : ToExpr UInt32) -instance : ProxyExpr UInt64 := (inferInstance : ToExpr UInt64) -instance : ProxyExpr UInt16 := (inferInstance : ToExpr UInt16) -instance {n : Nat} : ProxyExpr (Fin n.succ) := (inferInstance : ToExpr (Fin n.succ)) - -instance : ProxyExpr Prop := (inferInstance : ToExpr Bool) - -instance [SampleableExt α][ProxyExpr α] : ProxyExpr (List α) := - (inferInstance : ToExpr (List (SampleableExt.proxy α))) -instance [SampleableExt α][ProxyExpr α] : ProxyExpr (Array α) := - (inferInstance : ToExpr (Array (SampleableExt.proxy α))) -instance [SampleableExt α][SampleableExt β] [ProxyExpr α][ProxyExpr β] : ProxyExpr (Prod α β) := - (inferInstance : ToExpr (Prod (SampleableExt.proxy α) (SampleableExt.proxy β) )) + addVarInfo var finalX h hExpr? finalR + + theorem bool_to_prop_fmly (β : Prop → Prop): NamedBinder var (∀ (p : Prop), β p) → ∀ (b : Bool), β (b = true) := fun h b => h (b = true) @@ -640,7 +623,7 @@ theorem prop_iff_subtype (p : α → Prop) (β : α → Prop) : NamedBinder var instance (priority := 2000) subtypeVarTestable {p : α → Prop} {β : α → Prop} [∀ x, PrintableProp (p x)] [∀ x, MetaTestable (β x)][ToExpr α] - [SampleableExt (Subtype p)] [ProxyExpr (Subtype p)] {var'} : + [SampleableExt (Subtype p)] {var'} : MetaTestable (NamedBinder var <| (x : α) → NamedBinder var' <| p x → β x) where runExpr cfg min e? := match e? with @@ -651,7 +634,7 @@ instance (priority := 2000) subtypeVarTestable {p : α → Prop} {β : α → Pr let idExp ← mkAppOptM ``id #[e] addInfo s!"guard: {printProp (p x)} (by construction)" id idExp r (PSum.inr id) } do - let r ← @runExpr (∀ x : Subtype p, β x.val) (@varTestable var _ _ _ _ _) cfg min none + let r ← @runExpr (∀ x : Subtype p, β x.val) (@varTestable var _ _ _ _) cfg min none iff (prop_iff_subtype p β) none r | some e => do let (some pExp, some βExp, some _) ← forallPropProp? e | throwError m!"Expected a `Forall` proposition with arrow, got {← ppExpr e}" @@ -668,7 +651,7 @@ instance (priority := 2000) subtypeVarTestable {p : α → Prop} {β : α → Pr let x' ← mkAppM ``Subtype.val #[x] let y ← mkAppM' βExp #[x'] mkForallFVars #[x] y - let r ← @runExpr (∀ x : Subtype p, β x.val) (@varTestable var _ _ _ _ _) cfg min e' + let r ← @runExpr (∀ x : Subtype p, β x.val) (@varTestable var _ _ _ _) cfg min e' let hExp ← mkAppM ``prop_iff_subtype #[pExp, βExp] iff (prop_iff_subtype p β) hExp r diff --git a/Plausible/Sampleable.lean b/Plausible/Sampleable.lean index f12090e..715b169 100644 --- a/Plausible/Sampleable.lean +++ b/Plausible/Sampleable.lean @@ -3,8 +3,10 @@ Copyright (c) 2022 Henrik Böving. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Henrik Böving, Simon Hudon -/ +import Lean.Meta import Lean.Elab.Command import Plausible.Gen +open Lean /-! # `SampleableExt` Class @@ -109,23 +111,39 @@ class SampleableExt (α : Sort u) where proxy : Type v [proxyRepr : Repr proxy] [shrink : Shrinkable proxy] + proxyExpr? : Option (ToExpr proxy) sample : Gen proxy interp : proxy → α attribute [instance] SampleableExt.proxyRepr attribute [instance] SampleableExt.shrink +def getProxyExpr? (α : Sort u) [inst: SampleableExt α] : + Option (ToExpr (inst.proxy)) := + inst.proxyExpr? namespace SampleableExt /-- Use to generate instance whose purpose is to simply generate values of a type directly using the `Gen` monad -/ -def mkSelfContained [Repr α] [Shrinkable α] (sample : Gen α) : SampleableExt α where +def mkSelfContained [Repr α] [Shrinkable α] [Lean.ToExpr α] (sample : Gen α) : SampleableExt α where proxy := α proxyRepr := inferInstance shrink := inferInstance + proxyExpr? := some inferInstance sample := sample interp := id +/-- Use to generate instance whose purpose is to simply generate values +of a type directly using the `Gen` monad -/ +def mkSelfContainedSimple [Repr α] [Shrinkable α] (sample : Gen α) : SampleableExt α where + proxy := α + proxyRepr := inferInstance + shrink := inferInstance + proxyExpr? := none + sample := sample + interp := id + + /-- First samples a proxy value and interprets it. Especially useful if the proxy and target type are the same. -/ def interpSample (α : Type u) [SampleableExt α] : Gen α := @@ -242,6 +260,18 @@ section Samplers open SampleableExt +instance instToExprSum [ToExpr α] [ToExpr β] : ToExpr (α ⊕ β) := + let αType := toTypeExpr α + let βType := toTypeExpr β + { toExpr := fun a => + match a with + | .inl a => + mkApp3 (mkConst ``Sum.inl [levelZero, levelZero]) αType βType (toExpr a) + | .inr a => + mkApp3 (mkConst ``Sum.inr [levelZero, levelZero]) αType βType (toExpr a), + toTypeExpr := mkApp2 (mkConst ``Sum [levelZero, levelZero]) αType βType } + + instance Sum.SampleableExt [SampleableExt α] [SampleableExt β] : SampleableExt (Sum α β) where proxy := Sum (proxy α) (proxy β) sample := do @@ -252,6 +282,10 @@ instance Sum.SampleableExt [SampleableExt α] [SampleableExt β] : SampleableExt match s with | .inl l => .inl (interp l) | .inr r => .inr (interp r) + proxyExpr? := do + let inst₁ ← getProxyExpr? α + let inst₂ ← getProxyExpr? β + pure <| @instToExprSum _ _ inst₁ inst₂ instance Unit.sampleableExt : SampleableExt Unit := mkSelfContained (return ()) @@ -262,6 +296,9 @@ instance [SampleableExt α] [SampleableExt β] : SampleableExt ((_ : α) × β) let p ← prodOf sample sample return ⟨p.fst, p.snd⟩ interp s := ⟨interp s.fst, interp s.snd⟩ + proxyExpr? := none -- did not understand the difference with products + + instance Nat.sampleableExt : SampleableExt Nat := mkSelfContained (do choose Nat 0 (← getSize) (Nat.zero_le _)) @@ -331,6 +368,7 @@ instance Option.sampleableExt [SampleableExt α] : SampleableExt (Option α) whe | true => return none | false => return some (← sample) interp o := o.map interp + proxyExpr? := getProxyExpr? α |>.map fun p => (@instToExprOption _ p) instance Prod.sampleableExt {α : Type u} {β : Type v} [SampleableExt α] [SampleableExt β] : SampleableExt (α × β) where @@ -339,23 +377,32 @@ instance Prod.sampleableExt {α : Type u} {β : Type v} [SampleableExt α] [Samp shrink := inferInstance sample := prodOf sample sample interp := Prod.map interp interp + proxyExpr? := do + let inst₁ ← getProxyExpr? α + let inst₂ ← getProxyExpr? β + pure <| @instToExprProd _ _ inst₁ inst₂ + instance Prop.sampleableExt : SampleableExt Prop where proxy := Bool proxyRepr := inferInstance sample := interpSample Bool shrink := inferInstance + proxyExpr? := some inferInstance interp := Coe.coe instance List.sampleableExt [SampleableExt α] : SampleableExt (List α) where proxy := List (proxy α) sample := Gen.listOf sample interp := List.map interp + proxyExpr? := getProxyExpr? α |>.map fun p => (@instToExprList _ p) instance ULift.sampleableExt [SampleableExt α] : SampleableExt (ULift α) where proxy := proxy α sample := sample interp a := ⟨interp a⟩ + proxyExpr? := none + instance String.sampleableExt : SampleableExt String := mkSelfContained do return String.mk (← Gen.listOf (Char.sampleableDefault.sample)) @@ -364,12 +411,19 @@ instance Array.sampleableExt [SampleableExt α] : SampleableExt (Array α) where proxy := Array (proxy α) sample := Gen.arrayOf sample interp := Array.map interp + proxyExpr? := getProxyExpr? α |>.map fun p => (@instToExprArray _ p) + end Samplers /-- An annotation for values that should never get shrunk. -/ def NoShrink (α : Type u) := α - +open Lean in +instance [inst: ToExpr α] : ToExpr (NoShrink α) := + { + toExpr := fun x => inst.toExpr x + toTypeExpr := inst.toTypeExpr + } namespace NoShrink def mk (x : α) : NoShrink α := x @@ -381,7 +435,7 @@ instance repr [inst : Repr α] : Repr (NoShrink α) := inst instance shrinkable : Shrinkable (NoShrink α) where shrink := fun _ => [] -instance sampleableExt [SampleableExt α] [Repr α] : SampleableExt (NoShrink α) := +instance sampleableExt [SampleableExt α] [Repr α] [Lean.ToExpr α] : SampleableExt (NoShrink α) := SampleableExt.mkSelfContained <| (NoShrink.mk ∘ SampleableExt.interp) <$> SampleableExt.sample end NoShrink diff --git a/Test/Testable.lean b/Test/Testable.lean index 6df0b4b..a5a044a 100644 --- a/Test/Testable.lean +++ b/Test/Testable.lean @@ -19,7 +19,7 @@ instance : Shrinkable MyType where proxy.map (fun (fst, snd) => ⟨fst, fst + snd, by omega⟩) instance : SampleableExt MyType := - SampleableExt.mkSelfContained do + SampleableExt.mkSelfContainedSimple do let x ← SampleableExt.interpSample Nat let xyDiff ← SampleableExt.interpSample Nat return ⟨x, x + xyDiff, by omega⟩ From 09bbbef47eb23c99ebc47468a1b156a9d115da7a Mon Sep 17 00:00:00 2001 From: Siddhartha Gadgil Date: Tue, 26 Nov 2024 17:24:25 +0530 Subject: [PATCH 25/30] making APIs more similar --- Plausible/MetaTestable.lean | 31 ++++++++++++++++++++++++------- Test/MetaTestable.lean | 8 ++++---- Test/Testable.lean | 3 +++ 3 files changed, 31 insertions(+), 11 deletions(-) diff --git a/Plausible/MetaTestable.lean b/Plausible/MetaTestable.lean index b50aab2..a837139 100644 --- a/Plausible/MetaTestable.lean +++ b/Plausible/MetaTestable.lean @@ -700,7 +700,7 @@ def giveUp (x : Nat) : MetaTestResult p → MetaTestResult p end MetaTestable /-- Try `n` times to find a counter-example for `p`. -/ -def MetaTestable.runSuiteAux (p : Prop) [MetaTestable p] (propExpr: Expr) (cfg : Configuration) : +def MetaTestable.runSuiteAux (p : Prop) [MetaTestable p] (propExpr: Option Expr) (cfg : Configuration) : MetaTestResult p → Nat → MRand (MetaTestResult p) | r, 0 => return r | r, n+1 => do @@ -715,11 +715,11 @@ def MetaTestable.runSuiteAux (p : Prop) [MetaTestable p] (propExpr: Expr) (cfg : | _ => return x /-- Try to find a counter-example of `p`. -/ -def MetaTestable.runSuite (p : Prop) [MetaTestable p] (propExpr: Expr) (cfg : Configuration := {}) : MRand (MetaTestResult p) := +def MetaTestable.runSuite (p : Prop) [MetaTestable p] (propExpr: Option Expr) (cfg : Configuration := {}) : MRand (MetaTestResult p) := MetaTestable.runSuiteAux p propExpr cfg (success <| PSum.inl ()) cfg.numInst /-- Run a test suite for `p` in `MetaM` using the global RNG in `stdGenRef`. -/ -def MetaTestable.checkM (p : Prop) [MetaTestable p] (cfg : Configuration := {}) (propExpr: Expr) : MetaM (MetaTestResult p) := +def MetaTestable.seekM (p : Prop) [MetaTestable p] (cfg : Configuration := {}) (propExpr: Option Expr) : MetaM (MetaTestResult p) := match cfg.randomSeed with | none => runRand (MetaTestable.runSuite p propExpr cfg) | some seed => runRandWith seed (MetaTestable.runSuite p propExpr cfg) @@ -728,9 +728,9 @@ end Meta open Decorations in /-- Run a test suite for `p` and throw an exception if `p` does not hold. -/ -def MetaTestable.check (p : Prop) (cfg : Configuration := {}) +def MetaTestable.seek (p : Prop) (cfg : Configuration := {}) (p' : Decorations.DecorationsOf p := by mk_decorations) [MetaTestable p'](propExpr: Expr) : Lean.MetaM (Option Expr) := do - match ← MetaTestable.checkM p' cfg propExpr with + match ← MetaTestable.seekM p' cfg propExpr with | MetaTestResult.success _ => if !cfg.quiet then Lean.logInfo "Unable to find a counter-example" return none @@ -753,7 +753,7 @@ def disproveM? (cfg : Configuration) (tgt: Expr) : MetaM <| Option Expr := do synthInstance (← mkAppM ``MetaTestable #[tgt']) catch _ => throwError "Failed to create a `testable` instance for `{tgt}`." - let e ← mkAppOptM ``MetaTestable.check #[tgt, toExpr cfg, tgt', inst] + let e ← mkAppOptM ``MetaTestable.seek #[tgt, toExpr cfg, tgt', inst] let expectedType := Lean.Expr.forallE `a (Lean.Expr.const `Lean.Expr []) (Lean.Expr.app @@ -765,7 +765,24 @@ def disproveM? (cfg : Configuration) (tgt: Expr) : MetaM <| Option Expr := do let code ← unsafe evalExpr (Expr → MetaM (Option Expr)) expectedType e code tgt --- Negate +open Decorations in +/-- Run a test suite for `p` and throw an exception if `p` does not hold. -/ +def MetaTestable.check (p : Prop) (cfg : Configuration := {}) + (p' : Decorations.DecorationsOf p := by mk_decorations) [MetaTestable p'](propExpr: Option Expr := none) : Lean.MetaM Unit := do + match ← MetaTestable.seekM p' cfg propExpr with + | MetaTestResult.success _ => + if !cfg.quiet then Lean.logInfo "Unable to find a counter-example" + | MetaTestResult.gaveUp n => + if !cfg.quiet then + let msg := s!"Gave up after failing to generate values that fulfill the preconditions {n} times." + Lean.logWarning msg + | MetaTestResult.failure _ _ xs n => + let msg := "Found a counter-example!" + if cfg.quiet then + throwError msg + else + throwError Testable.formatFailure msg xs n + end Plausible diff --git a/Test/MetaTestable.lean b/Test/MetaTestable.lean index 802ea4b..4287ed3 100644 --- a/Test/MetaTestable.lean +++ b/Test/MetaTestable.lean @@ -149,12 +149,12 @@ example : MetaTestable (NamedBinder "a" (∀ (a : Nat), a ≤ 1)) := inferInstan example : MetaTestable (NamedBinder "a" (∀ (a : Nat), NamedBinder "b" (∀ (b : Nat), a ≤ b))) := inferInstance -- Main tests: finding counterexamples -#eval MetaTestable.check (∀ (_a : Nat), False) (propExpr := Lean.Expr.forallE `a (Lean.Expr.const `Nat []) (Lean.Expr.const `False []) (Lean.BinderInfo.default)) +#eval MetaTestable.seek (∀ (_a : Nat), False) (propExpr := Lean.Expr.forallE `a (Lean.Expr.const `Nat []) (Lean.Expr.const `False []) (Lean.BinderInfo.default)) -#eval MetaTestable.check (∀ (a : Nat), a < 1) (propExpr := Lean.Expr.forallE `a (Lean.Expr.const `Nat []) (Lean.Expr.const `False []) (Lean.BinderInfo.default)) +#eval MetaTestable.seek (∀ (a : Nat), a < 1) (propExpr := Lean.Expr.forallE `a (Lean.Expr.const `Nat []) (Lean.Expr.const `False []) (Lean.BinderInfo.default)) -#eval MetaTestable.check (∀ (a b : Nat), a < b) (propExpr := (Lean.Expr.forallE +#eval MetaTestable.seek (∀ (a b : Nat), a < b) (propExpr := (Lean.Expr.forallE `a (Lean.Expr.const `Nat []) (Lean.Expr.forallE @@ -172,7 +172,7 @@ example : MetaTestable (NamedBinder "a" (∀ (a : Nat), NamedBinder "b" (∀ (b (Lean.BinderInfo.default)) (Lean.BinderInfo.default))) -#eval MetaTestable.check (∀ (a _b : Nat), a < 0) (propExpr := (Lean.Expr.forallE +#eval MetaTestable.seek (∀ (a _b : Nat), a < 0) (propExpr := (Lean.Expr.forallE `a (Lean.Expr.const `Nat []) (Lean.Expr.forallE diff --git a/Test/Testable.lean b/Test/Testable.lean index a5a044a..504c469 100644 --- a/Test/Testable.lean +++ b/Test/Testable.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Henrik Böving -/ import Plausible.Testable +import Plausible.MetaTestable open Plausible @@ -25,3 +26,5 @@ instance : SampleableExt MyType := return ⟨x, x + xyDiff, by omega⟩ #eval Testable.check <| ∀ a b : MyType, a.y ≤ b.x → a.x ≤ b.y + +#eval MetaTestable.check <| ∀ a b : MyType, a.y ≤ b.x → a.x ≤ b.y From 713defb6dc1b9ad9881a151918be5dddbb90b353 Mon Sep 17 00:00:00 2001 From: Siddhartha Gadgil Date: Thu, 28 Nov 2024 19:46:32 +0530 Subject: [PATCH 26/30] removed product expression, fixes function case --- Plausible/Functions.lean | 24 +++++++++++------------- Plausible/Sampleable.lean | 13 ++++++++----- Plausible/tmp_data.md | 23 +++++++++++++++++++++++ 3 files changed, 42 insertions(+), 18 deletions(-) create mode 100644 Plausible/tmp_data.md diff --git a/Plausible/Functions.lean b/Plausible/Functions.lean index f74f98c..f02ab2e 100644 --- a/Plausible/Functions.lean +++ b/Plausible/Functions.lean @@ -126,19 +126,17 @@ where variable [Repr α] --- instance Pi.sampleableExt : SampleableExt (α → β) where --- proxy := TotalFunction α (SampleableExt.proxy β) --- interp f := SampleableExt.interp ∘ f.apply --- sample := do --- let xs : List (_ × _) ← (SampleableExt.sample (α := List (α × β))) --- let ys := (SampleableExt.sample : Gen (SampleableExt.proxy β)) --- let ys' := Gen.up ys --- let ⟨x⟩ ← ys' --- pure <| TotalFunction.withDefault (List.toFinmap' <| xs.map <| --- Prod.map SampleableExt.interp id) x --- -- note: no way of shrinking the domain without an inverse to `interp` --- shrink := { shrink := letI : Shrinkable α := {}; TotalFunction.shrink } --- proxyExpr? := none +instance Pi.sampleableExt : SampleableExt (α → β) where + proxy := TotalFunction α (SampleableExt.proxy β) + interp f := SampleableExt.interp ∘ f.apply + sample := do + let xs : List (_ × _) ← (SampleableExt.sample (α := List (α × β))) + let ⟨x⟩ ← Gen.up (SampleableExt.sample : Gen (SampleableExt.proxy β)) + pure <| TotalFunction.withDefault (List.toFinmap' <| xs.map <| + Prod.map SampleableExt.interp id) x + -- note: no way of shrinking the domain without an inverse to `interp` + shrink := { shrink := letI : Shrinkable α := {}; TotalFunction.shrink } + proxyExpr? := none end diff --git a/Plausible/Sampleable.lean b/Plausible/Sampleable.lean index 715b169..d3c0d09 100644 --- a/Plausible/Sampleable.lean +++ b/Plausible/Sampleable.lean @@ -377,11 +377,14 @@ instance Prod.sampleableExt {α : Type u} {β : Type v} [SampleableExt α] [Samp shrink := inferInstance sample := prodOf sample sample interp := Prod.map interp interp - proxyExpr? := do - let inst₁ ← getProxyExpr? α - let inst₂ ← getProxyExpr? β - pure <| @instToExprProd _ _ inst₁ inst₂ - + proxyExpr? := none -- an expression causes a problem with Function + -- do + -- let inst₁ ← getProxyExpr? α + -- let inst₂ ← getProxyExpr? β + -- pure <| @instToExprProd _ _ inst₁ inst₂ + +set_option pp.universes true +#check instToExprProd instance Prop.sampleableExt : SampleableExt Prop where proxy := Bool diff --git a/Plausible/tmp_data.md b/Plausible/tmp_data.md new file mode 100644 index 0000000..52e33e1 --- /dev/null +++ b/Plausible/tmp_data.md @@ -0,0 +1,23 @@ +def Plausible.TotalFunction.Pi.sampleableExt.{u, v, ub} : {α : Type u} → + {β : Type v} → + [inst : SampleableExt.{u + 1, u} α] → + [inst : SampleableExt.{v + 1, ub} β] → + [inst : DecidableEq.{u + 1} α] → [inst : Repr.{u} α] → SampleableExt.{max (u + 1) (v + 1), max ub u} (α → β) := +fun {α} {β} [SampleableExt.{u + 1, u} α] [SampleableExt.{v + 1, ub} β] [DecidableEq.{u + 1} α] [Repr.{u} α] => + SampleableExt.mk.{max (u + 1) (v + 1), max ub u} (TotalFunction.{u, ub} α (SampleableExt.proxy.{v + 1, ub} β)) + (do + let xs ← SampleableExt.sample.{(max v u) + 1, max u ub} + let __discr ← Gen.up.{ub, u} SampleableExt.sample.{v + 1, ub} + match __discr with + | ULift.up.{u, ub} x => + pure.{max u ub, max u ub} + (withDefault.{u, ub} + (List.toFinmap'.{u, ub} + (List.map.{max ub u, max u ub} (Prod.map.{u, u, ub, ub} SampleableExt.interp.{u + 1, u} id.{ub + 1}) + xs)) + x)) + fun f => Function.comp.{u + 1, ub + 1, v + 1} SampleableExt.interp.{v + 1, ub} (apply.{u, ub} f) + +```lean +xs : List.{max ub u} (Prod.{u, ub} (SampleableExt.proxy.{u + 1, u} α) (SampleableExt.proxy.{v + 1, ub} β)) +``` \ No newline at end of file From fb1fa9ff8aa20ec4efcac5cc5af6567c7d9357f2 Mon Sep 17 00:00:00 2001 From: Siddhartha Gadgil Date: Fri, 29 Nov 2024 19:00:34 +0530 Subject: [PATCH 27/30] cleaner definition fixes regression --- Plausible/Sampleable.lean | 11 +++++------ 1 file changed, 5 insertions(+), 6 deletions(-) diff --git a/Plausible/Sampleable.lean b/Plausible/Sampleable.lean index d3c0d09..9f82c96 100644 --- a/Plausible/Sampleable.lean +++ b/Plausible/Sampleable.lean @@ -370,18 +370,17 @@ instance Option.sampleableExt [SampleableExt α] : SampleableExt (Option α) whe interp o := o.map interp proxyExpr? := getProxyExpr? α |>.map fun p => (@instToExprOption _ p) -instance Prod.sampleableExt {α : Type u} {β : Type v} [SampleableExt α] [SampleableExt β] : +instance Prod.sampleableExt {α : Type u} {β : Type v} [inst₁ : SampleableExt α] [inst₂ : SampleableExt β] : SampleableExt (α × β) where proxy := Prod (proxy α) (proxy β) proxyRepr := inferInstance shrink := inferInstance sample := prodOf sample sample interp := Prod.map interp interp - proxyExpr? := none -- an expression causes a problem with Function - -- do - -- let inst₁ ← getProxyExpr? α - -- let inst₂ ← getProxyExpr? β - -- pure <| @instToExprProd _ _ inst₁ inst₂ + proxyExpr? := + match inst₁.proxyExpr?, inst₂.proxyExpr? with + | some p₁, some p₂ => some <| @instToExprProd _ _ p₁ p₂ + | _ , _ => none set_option pp.universes true #check instToExprProd From 2217b44d355f131ff73875ddea577d440838f0a6 Mon Sep 17 00:00:00 2001 From: Siddhartha Gadgil Date: Fri, 29 Nov 2024 19:01:43 +0530 Subject: [PATCH 28/30] purge debug data --- Plausible/tmp_data.md | 23 ----------------------- 1 file changed, 23 deletions(-) delete mode 100644 Plausible/tmp_data.md diff --git a/Plausible/tmp_data.md b/Plausible/tmp_data.md deleted file mode 100644 index 52e33e1..0000000 --- a/Plausible/tmp_data.md +++ /dev/null @@ -1,23 +0,0 @@ -def Plausible.TotalFunction.Pi.sampleableExt.{u, v, ub} : {α : Type u} → - {β : Type v} → - [inst : SampleableExt.{u + 1, u} α] → - [inst : SampleableExt.{v + 1, ub} β] → - [inst : DecidableEq.{u + 1} α] → [inst : Repr.{u} α] → SampleableExt.{max (u + 1) (v + 1), max ub u} (α → β) := -fun {α} {β} [SampleableExt.{u + 1, u} α] [SampleableExt.{v + 1, ub} β] [DecidableEq.{u + 1} α] [Repr.{u} α] => - SampleableExt.mk.{max (u + 1) (v + 1), max ub u} (TotalFunction.{u, ub} α (SampleableExt.proxy.{v + 1, ub} β)) - (do - let xs ← SampleableExt.sample.{(max v u) + 1, max u ub} - let __discr ← Gen.up.{ub, u} SampleableExt.sample.{v + 1, ub} - match __discr with - | ULift.up.{u, ub} x => - pure.{max u ub, max u ub} - (withDefault.{u, ub} - (List.toFinmap'.{u, ub} - (List.map.{max ub u, max u ub} (Prod.map.{u, u, ub, ub} SampleableExt.interp.{u + 1, u} id.{ub + 1}) - xs)) - x)) - fun f => Function.comp.{u + 1, ub + 1, v + 1} SampleableExt.interp.{v + 1, ub} (apply.{u, ub} f) - -```lean -xs : List.{max ub u} (Prod.{u, ub} (SampleableExt.proxy.{u + 1, u} α) (SampleableExt.proxy.{v + 1, ub} β)) -``` \ No newline at end of file From deaaf3b34b22ebac4df471717d88a2914f512173 Mon Sep 17 00:00:00 2001 From: Siddhartha Gadgil Date: Sat, 30 Nov 2024 09:59:19 +0530 Subject: [PATCH 29/30] profiling --- Plausible/MetaTestable.lean | 3 ++- Test/Testable.lean | 1 + 2 files changed, 3 insertions(+), 1 deletion(-) diff --git a/Plausible/MetaTestable.lean b/Plausible/MetaTestable.lean index a837139..6c02550 100644 --- a/Plausible/MetaTestable.lean +++ b/Plausible/MetaTestable.lean @@ -399,7 +399,8 @@ instance decGuardTestable [PrintableProp p] [Decidable p] {β : p → Prop} [∀ | none => let res := runPropExpr (β h) cfg min none let s := printProp p - ← (fun r => addInfo s!"guard: {s}" (· <| h) none r (PSum.inr <| fun q _ => q)) <$> res | some e => + ← (fun r => addInfo s!"guard: {s}" (· <| h) none r (PSum.inr <| fun q _ => q)) <$> res + | some e => let (some βExp, some pExp) ← forallProp? e | throwError m!"Expected a `Forall` proposition, got {← ppExpr e}" let h' := (· <| h) let yExp ← diff --git a/Test/Testable.lean b/Test/Testable.lean index 504c469..516df02 100644 --- a/Test/Testable.lean +++ b/Test/Testable.lean @@ -27,4 +27,5 @@ instance : SampleableExt MyType := #eval Testable.check <| ∀ a b : MyType, a.y ≤ b.x → a.x ≤ b.y +set_option trace.profiler true in #eval MetaTestable.check <| ∀ a b : MyType, a.y ≤ b.x → a.x ≤ b.y From 9decec18f4d47356ee2fe19364745c3032009efe Mon Sep 17 00:00:00 2001 From: Siddhartha Gadgil Date: Sat, 30 Nov 2024 17:26:47 +0530 Subject: [PATCH 30/30] some fixes --- Plausible/Sampleable.lean | 8 +++++++- 1 file changed, 7 insertions(+), 1 deletion(-) diff --git a/Plausible/Sampleable.lean b/Plausible/Sampleable.lean index 9f82c96..45c5487 100644 --- a/Plausible/Sampleable.lean +++ b/Plausible/Sampleable.lean @@ -106,11 +106,15 @@ If we test a quantification over functions the counter-examples cannot be shrunken or printed meaningfully. For that purpose, `SampleableExt` provides a proxy representation `proxy` that can be printed and shrunken as well -as interpreted (using `interp`) as an object of the right type. -/ +as interpreted (using `interp`) as an object of the right type. + +SampleableExt can also be used to constuct expressions representing disproofs in some cases. For this, the `proxy` type should be a type that can be represented as an expression and the field `proxyExpr?` should be set to such a function. Expressions for disproofs will not be generated for sampling involving instances where `proxyExpr?` is `none`. +-/ class SampleableExt (α : Sort u) where proxy : Type v [proxyRepr : Repr proxy] [shrink : Shrinkable proxy] + /-- Expressions representing terms of type proxy if available.-/ proxyExpr? : Option (ToExpr proxy) sample : Gen proxy interp : proxy → α @@ -258,6 +262,8 @@ end Shrinkers section Samplers +variable {α : Type u} {β : Type v} + open SampleableExt instance instToExprSum [ToExpr α] [ToExpr β] : ToExpr (α ⊕ β) :=