├── .github └── workflows │ └── lean_action_ci.yml ├── .gitignore ├── LICENSE ├── Plausible.lean ├── Plausible ├── Attr.lean ├── Functions.lean ├── Gen.lean ├── Random.lean ├── Sampleable.lean ├── Tactic.lean └── Testable.lean ├── README.md ├── Test.lean ├── Test ├── Tactic.lean └── Testable.lean ├── lake-manifest.json ├── lakefile.toml └── lean-toolchain /.github/workflows/lean_action_ci.yml: -------------------------------------------------------------------------------- 1 | name: Lean Action CI 2 | 3 | on: 4 | push: 5 | pull_request: 6 | workflow_dispatch: 7 | 8 | jobs: 9 | build: 10 | runs-on: ubuntu-latest 11 | 12 | steps: 13 | - uses: actions/checkout@v4 14 | - uses: leanprover/lean-action@v1 15 | -------------------------------------------------------------------------------- /.gitignore: -------------------------------------------------------------------------------- 1 | /.lake 2 | -------------------------------------------------------------------------------- /LICENSE: -------------------------------------------------------------------------------- 1 | Apache License 2 | Version 2.0, January 2004 3 | http://www.apache.org/licenses/ 4 | 5 | TERMS AND CONDITIONS FOR USE, REPRODUCTION, AND DISTRIBUTION 6 | 7 | 1. Definitions. 8 | 9 | "License" shall mean the terms and conditions for use, reproduction, 10 | and distribution as defined by Sections 1 through 9 of this document. 11 | 12 | "Licensor" shall mean the copyright owner or entity authorized by 13 | the copyright owner that is granting the License. 14 | 15 | "Legal Entity" shall mean the union of the acting entity and all 16 | other entities that control, are controlled by, or are under common 17 | control with that entity. For the purposes of this definition, 18 | "control" means (i) the power, direct or indirect, to cause the 19 | direction or management of such entity, whether by contract or 20 | otherwise, or (ii) ownership of fifty percent (50%) or more of the 21 | outstanding shares, or (iii) beneficial ownership of such entity. 22 | 23 | "You" (or "Your") shall mean an individual or Legal Entity 24 | exercising permissions granted by this License. 25 | 26 | "Source" form shall mean the preferred form for making modifications, 27 | including but not limited to software source code, documentation 28 | source, and configuration files. 29 | 30 | "Object" form shall mean any form resulting from mechanical 31 | transformation or translation of a Source form, including but 32 | not limited to compiled object code, generated documentation, 33 | and conversions to other media types. 34 | 35 | "Work" shall mean the work of authorship, whether in Source or 36 | Object form, made available under the License, as indicated by a 37 | copyright notice that is included in or attached to the work 38 | (an example is provided in the Appendix below). 39 | 40 | "Derivative Works" shall mean any work, whether in Source or Object 41 | form, that is based on (or derived from) the Work and for which the 42 | editorial revisions, annotations, elaborations, or other modifications 43 | represent, as a whole, an original work of authorship. For the purposes 44 | of this License, Derivative Works shall not include works that remain 45 | separable from, or merely link (or bind by name) to the interfaces of, 46 | the Work and Derivative Works thereof. 47 | 48 | "Contribution" shall mean any work of authorship, including 49 | the original version of the Work and any modifications or additions 50 | to that Work or Derivative Works thereof, that is intentionally 51 | submitted to Licensor for inclusion in the Work by the copyright owner 52 | or by an individual or Legal Entity authorized to submit on behalf of 53 | the copyright owner. For the purposes of this definition, "submitted" 54 | means any form of electronic, verbal, or written communication sent 55 | to the Licensor or its representatives, including but not limited to 56 | communication on electronic mailing lists, source code control systems, 57 | and issue tracking systems that are managed by, or on behalf of, the 58 | Licensor for the purpose of discussing and improving the Work, but 59 | excluding communication that is conspicuously marked or otherwise 60 | designated in writing by the copyright owner as "Not a Contribution." 61 | 62 | "Contributor" shall mean Licensor and any individual or Legal Entity 63 | on behalf of whom a Contribution has been received by Licensor and 64 | subsequently incorporated within the Work. 65 | 66 | 2. Grant of Copyright License. Subject to the terms and conditions of 67 | this License, each Contributor hereby grants to You a perpetual, 68 | worldwide, non-exclusive, no-charge, royalty-free, irrevocable 69 | copyright license to reproduce, prepare Derivative Works of, 70 | publicly display, publicly perform, sublicense, and distribute the 71 | Work and such Derivative Works in Source or Object form. 72 | 73 | 3. Grant of Patent License. Subject to the terms and conditions of 74 | this License, each Contributor hereby grants to You a perpetual, 75 | worldwide, non-exclusive, no-charge, royalty-free, irrevocable 76 | (except as stated in this section) patent license to make, have made, 77 | use, offer to sell, sell, import, and otherwise transfer the Work, 78 | where such license applies only to those patent claims licensable 79 | by such Contributor that are necessarily infringed by their 80 | Contribution(s) alone or by combination of their Contribution(s) 81 | with the Work to which such Contribution(s) was submitted. If You 82 | institute patent litigation against any entity (including a 83 | cross-claim or counterclaim in a lawsuit) alleging that the Work 84 | or a Contribution incorporated within the Work constitutes direct 85 | or contributory patent infringement, then any patent licenses 86 | granted to You under this License for that Work shall terminate 87 | as of the date such litigation is filed. 88 | 89 | 4. Redistribution. You may reproduce and distribute copies of the 90 | Work or Derivative Works thereof in any medium, with or without 91 | modifications, and in Source or Object form, provided that You 92 | meet the following conditions: 93 | 94 | (a) You must give any other recipients of the Work or 95 | Derivative Works a copy of this License; and 96 | 97 | (b) You must cause any modified files to carry prominent notices 98 | stating that You changed the files; and 99 | 100 | (c) You must retain, in the Source form of any Derivative Works 101 | that You distribute, all copyright, patent, trademark, and 102 | attribution notices from the Source form of the Work, 103 | excluding those notices that do not pertain to any part of 104 | the Derivative Works; and 105 | 106 | (d) If the Work includes a "NOTICE" text file as part of its 107 | distribution, then any Derivative Works that You distribute must 108 | include a readable copy of the attribution notices contained 109 | within such NOTICE file, excluding those notices that do not 110 | pertain to any part of the Derivative Works, in at least one 111 | of the following places: within a NOTICE text file distributed 112 | as part of the Derivative Works; within the Source form or 113 | documentation, if provided along with the Derivative Works; or, 114 | within a display generated by the Derivative Works, if and 115 | wherever such third-party notices normally appear. The contents 116 | of the NOTICE file are for informational purposes only and 117 | do not modify the License. You may add Your own attribution 118 | notices within Derivative Works that You distribute, alongside 119 | or as an addendum to the NOTICE text from the Work, provided 120 | that such additional attribution notices cannot be construed 121 | as modifying the License. 122 | 123 | You may add Your own copyright statement to Your modifications and 124 | may provide additional or different license terms and conditions 125 | for use, reproduction, or distribution of Your modifications, or 126 | for any such Derivative Works as a whole, provided Your use, 127 | reproduction, and distribution of the Work otherwise complies with 128 | the conditions stated in this License. 129 | 130 | 5. Submission of Contributions. Unless You explicitly state otherwise, 131 | any Contribution intentionally submitted for inclusion in the Work 132 | by You to the Licensor shall be under the terms and conditions of 133 | this License, without any additional terms or conditions. 134 | Notwithstanding the above, nothing herein shall supersede or modify 135 | the terms of any separate license agreement you may have executed 136 | with Licensor regarding such Contributions. 137 | 138 | 6. Trademarks. This License does not grant permission to use the trade 139 | names, trademarks, service marks, or product names of the Licensor, 140 | except as required for reasonable and customary use in describing the 141 | origin of the Work and reproducing the content of the NOTICE file. 142 | 143 | 7. Disclaimer of Warranty. Unless required by applicable law or 144 | agreed to in writing, Licensor provides the Work (and each 145 | Contributor provides its Contributions) on an "AS IS" BASIS, 146 | WITHOUT WARRANTIES OR CONDITIONS OF ANY KIND, either express or 147 | implied, including, without limitation, any warranties or conditions 148 | of TITLE, NON-INFRINGEMENT, MERCHANTABILITY, or FITNESS FOR A 149 | PARTICULAR PURPOSE. You are solely responsible for determining the 150 | appropriateness of using or redistributing the Work and assume any 151 | risks associated with Your exercise of permissions under this License. 152 | 153 | 8. Limitation of Liability. In no event and under no legal theory, 154 | whether in tort (including negligence), contract, or otherwise, 155 | unless required by applicable law (such as deliberate and grossly 156 | negligent acts) or agreed to in writing, shall any Contributor be 157 | liable to You for damages, including any direct, indirect, special, 158 | incidental, or consequential damages of any character arising as a 159 | result of this License or out of the use or inability to use the 160 | Work (including but not limited to damages for loss of goodwill, 161 | work stoppage, computer failure or malfunction, or any and all 162 | other commercial damages or losses), even if such Contributor 163 | has been advised of the possibility of such damages. 164 | 165 | 9. Accepting Warranty or Additional Liability. While redistributing 166 | the Work or Derivative Works thereof, You may choose to offer, 167 | and charge a fee for, acceptance of support, warranty, indemnity, 168 | or other liability obligations and/or rights consistent with this 169 | License. However, in accepting such obligations, You may act only 170 | on Your own behalf and on Your sole responsibility, not on behalf 171 | of any other Contributor, and only if You agree to indemnify, 172 | defend, and hold each Contributor harmless for any liability 173 | incurred by, or claims asserted against, such Contributor by reason 174 | of your accepting any such warranty or additional liability. 175 | 176 | END OF TERMS AND CONDITIONS 177 | 178 | APPENDIX: How to apply the Apache License to your work. 179 | 180 | To apply the Apache License to your work, attach the following 181 | boilerplate notice, with the fields enclosed by brackets "[]" 182 | replaced with your own identifying information. (Don't include 183 | the brackets!) The text should be enclosed in the appropriate 184 | comment syntax for the file format. We also recommend that a 185 | file or class name and description of purpose be included on the 186 | same "printed page" as the copyright notice for easier 187 | identification within third-party archives. 188 | 189 | Copyright [yyyy] [name of copyright owner] 190 | 191 | Licensed under the Apache License, Version 2.0 (the "License"); 192 | you may not use this file except in compliance with the License. 193 | You may obtain a copy of the License at 194 | 195 | http://www.apache.org/licenses/LICENSE-2.0 196 | 197 | Unless required by applicable law or agreed to in writing, software 198 | distributed under the License is distributed on an "AS IS" BASIS, 199 | WITHOUT WARRANTIES OR CONDITIONS OF ANY KIND, either express or implied. 200 | See the License for the specific language governing permissions and 201 | limitations under the License. 202 | -------------------------------------------------------------------------------- /Plausible.lean: -------------------------------------------------------------------------------- 1 | /- 2 | Copyright (c) 2024 Lean FRO, LLC. All rights reserved. 3 | Released under Apache 2.0 license as described in the file LICENSE. 4 | Authors: Henrik Böving 5 | -/ 6 | import Plausible.Random 7 | import Plausible.Gen 8 | import Plausible.Sampleable 9 | import Plausible.Testable 10 | import Plausible.Functions 11 | import Plausible.Attr 12 | import Plausible.Tactic 13 | -------------------------------------------------------------------------------- /Plausible/Attr.lean: -------------------------------------------------------------------------------- 1 | /- 2 | Copyright (c) 2020 Simon Hudon. All rights reserved. 3 | Released under Apache 2.0 license as described in the file LICENSE. 4 | Authors: Simon Hudon, Kim Morrison 5 | -/ 6 | import Lean.Util.Trace 7 | 8 | open Lean 9 | 10 | initialize registerTraceClass `plausible.instance 11 | initialize registerTraceClass `plausible.decoration 12 | initialize registerTraceClass `plausible.discarded 13 | initialize registerTraceClass `plausible.success 14 | initialize registerTraceClass `plausible.shrink.steps 15 | initialize registerTraceClass `plausible.shrink.candidates 16 | -------------------------------------------------------------------------------- /Plausible/Functions.lean: -------------------------------------------------------------------------------- 1 | /- 2 | Copyright (c) 2020 Simon Hudon. All rights reserved. 3 | Released under Apache 2.0 license as described in the file LICENSE. 4 | Authors: Simon Hudon 5 | -/ 6 | import Plausible.Sampleable 7 | import Plausible.Testable 8 | 9 | /-! 10 | ## `plausible`: generators for functions 11 | 12 | This file defines `Sampleable` instances for `α → β` functions and 13 | `Int → Int` injective functions. 14 | 15 | Functions are generated by creating a list of pairs and one more value 16 | using the list as a lookup table and resorting to the additional value 17 | when a value is not found in the table. 18 | 19 | Injective functions are generated by creating a list of numbers and 20 | a permutation of that list. The permutation insures that every input 21 | is mapped to a unique output. When an input is not found in the list 22 | the input itself is used as an output. 23 | 24 | Injective functions `f : α → α` could be generated easily instead of 25 | `Int → Int` by generating a `List α`, removing duplicates and creating a 26 | permutation. One has to be careful when generating the domain to make 27 | it vast enough that, when generating arguments to apply `f` to, 28 | they argument should be likely to lie in the domain of `f`. This is 29 | the reason that injective functions `f : Int → Int` are generated by 30 | fixing the domain to the range `[-2*size .. 2*size]`, with `size` 31 | the size parameter of the `gen` monad. 32 | 33 | Much of the machinery provided in this file is applicable to generate 34 | injective functions of type `α → α` and new instances should be easy 35 | to define. 36 | 37 | Other classes of functions such as monotone functions can generated using 38 | similar techniques. For monotone functions, generating two lists, sorting them 39 | and matching them should suffice, with appropriate default values. 40 | Some care must be taken for shrinking such functions to make sure 41 | their defining property is invariant through shrinking. Injective 42 | functions are an example of how complicated it can get. 43 | -/ 44 | 45 | 46 | universe u v w 47 | 48 | variable {α : Type u} {β : Type v} {γ : Sort w} 49 | 50 | namespace Plausible 51 | 52 | /-- Data structure specifying a total function using a list of pairs 53 | and a default value returned when the input is not in the domain of 54 | the partial function. 55 | 56 | `withDefault f y` encodes `x => f x` when `x ∈ f` and `x => y` 57 | otherwise. 58 | 59 | We use `Σ` to encode mappings instead of `×` because we 60 | rely on the association list API defined in `Mathlib/Data/List/Sigma.lean`. 61 | -/ 62 | inductive TotalFunction (α : Type u) (β : Type v) : Type max u v 63 | | withDefault : List (Σ _ : α, β) → β → TotalFunction α β 64 | 65 | instance TotalFunction.inhabited [Inhabited β] : Inhabited (TotalFunction α β) := 66 | ⟨TotalFunction.withDefault ∅ default⟩ 67 | 68 | namespace TotalFunction 69 | 70 | -- Porting note: new 71 | /-- Compose a total function with a regular function on the left -/ 72 | def comp {γ : Type w} (f : β → γ) : TotalFunction α β → TotalFunction α γ 73 | | TotalFunction.withDefault m y => 74 | TotalFunction.withDefault (m.map fun ⟨a, b⟩ => ⟨a, f b⟩) (f y) 75 | 76 | /-- Apply a total function to an argument. -/ 77 | def apply [DecidableEq α] : TotalFunction α β → α → β 78 | | TotalFunction.withDefault m y, x => (m.find? fun ⟨a, _⟩ => a = x).map Sigma.snd |>.getD y 79 | 80 | /-- Implementation of `Repr (TotalFunction α β)`. 81 | 82 | Creates a string for a given `Finmap` and output, `x₀ => y₀, .. xₙ => yₙ` 83 | for each of the entries. The brackets are provided by the calling function. 84 | -/ 85 | def reprAux [Repr α] [Repr β] (m : List (Σ _ : α, β)) : String := 86 | String.join <| 87 | -- Porting note: No `List.qsort`, so convert back and forth to an `Array`. 88 | Array.toList <| Array.qsort (lt := fun x y => x < y) 89 | (m.map fun x => s!"{(repr <| Sigma.fst x)} => {repr <| Sigma.snd x}, ").toArray 90 | 91 | /-- Produce a string for a given `TotalFunction`. 92 | The output is of the form `[x₀ => f x₀, .. xₙ => f xₙ, _ => y]`. 93 | -/ 94 | protected def repr [Repr α] [Repr β] : TotalFunction α β → String 95 | | TotalFunction.withDefault m y => s!"[{(reprAux m)}_ => {repr y}]" 96 | 97 | instance (α : Type u) (β : Type v) [Repr α] [Repr β] : Repr (TotalFunction α β) where 98 | reprPrec f _ := TotalFunction.repr f 99 | 100 | /-- Create a `Finmap` from a list of pairs. -/ 101 | def List.toFinmap' (xs : List (α × β)) : List (Σ _ : α, β) := 102 | xs.map (fun ⟨a, b⟩ => ⟨a, b⟩) 103 | 104 | section 105 | 106 | universe ua ub 107 | variable [SampleableExt.{_,u} α] [SampleableExt.{_,ub} β] 108 | 109 | variable [DecidableEq α] 110 | 111 | /-- Shrink a total function by shrinking the lists that represent it. -/ 112 | def shrink {α β} [DecidableEq α] [Shrinkable α] [Shrinkable β] : 113 | TotalFunction α β → List (TotalFunction α β) 114 | | ⟨m, x⟩ => (Shrinkable.shrink (m, x)).map fun ⟨m', x'⟩ => ⟨dedup m', x'⟩ 115 | where 116 | dedup (m' : List ((_ : α) × β)) : List ((_ : α) × β) := 117 | let rec insertKey (xs : List ((_ : α) × β)) (pair : (_ : α) × β) : List ((_ : α) × β) := 118 | match xs with 119 | | [] => [pair] 120 | | x :: xs => 121 | if pair.fst = x.fst then 122 | pair :: xs 123 | else 124 | x :: insertKey xs pair 125 | m'.foldl (init := []) insertKey 126 | 127 | variable [Repr α] 128 | 129 | instance Pi.sampleableExt : SampleableExt (α → β) where 130 | proxy := TotalFunction α (SampleableExt.proxy β) 131 | interp f := SampleableExt.interp ∘ f.apply 132 | sample := do 133 | let xs : List (_ × _) ← (SampleableExt.sample (α := List (α × β))) 134 | let ⟨x⟩ ← Gen.up <| (SampleableExt.sample : Gen (SampleableExt.proxy β)) 135 | pure <| TotalFunction.withDefault (List.toFinmap' <| xs.map <| 136 | Prod.map SampleableExt.interp id) x 137 | -- note: no way of shrinking the domain without an inverse to `interp` 138 | shrink := { shrink := letI : Shrinkable α := {}; TotalFunction.shrink } 139 | 140 | end 141 | 142 | section SampleableExt 143 | 144 | open SampleableExt 145 | 146 | instance (priority := 2000) PiPred.sampleableExt [SampleableExt (α → Bool)] : 147 | SampleableExt.{u + 1} (α → Prop) where 148 | proxy := proxy (α → Bool) 149 | interp m x := interp m x 150 | sample := sample 151 | shrink := SampleableExt.shrink 152 | 153 | instance (priority := 2000) PiUncurry.sampleableExt [SampleableExt (α × β → γ)] : 154 | SampleableExt.{imax (u + 1) (v + 1) w} (α → β → γ) where 155 | proxy := proxy (α × β → γ) 156 | interp m x y := interp m (x, y) 157 | sample := sample 158 | shrink := SampleableExt.shrink 159 | 160 | end SampleableExt 161 | 162 | end TotalFunction 163 | 164 | end Plausible 165 | -------------------------------------------------------------------------------- /Plausible/Gen.lean: -------------------------------------------------------------------------------- 1 | /- 2 | Copyright (c) 2021 Henrik Böving. All rights reserved. 3 | Released under Apache 2.0 license as described in the file LICENSE. 4 | Authors: Henrik Böving, Simon Hudon 5 | -/ 6 | import Plausible.Random 7 | 8 | /-! 9 | # `Gen` Monad 10 | 11 | This monad is used to formulate randomized computations with a parameter 12 | to specify the desired size of the result. 13 | 14 | ## Main definitions 15 | 16 | * `Gen` monad 17 | 18 | ## References 19 | 20 | * https://hackage.haskell.org/package/QuickCheck 21 | -/ 22 | 23 | universe u v 24 | 25 | namespace Plausible 26 | 27 | open Random 28 | 29 | /-- Monad to generate random examples to test properties with. 30 | It has a `Nat` parameter so that the caller can decide on the 31 | size of the examples. -/ 32 | abbrev Gen (α : Type u) := ReaderT (ULift Nat) Rand α 33 | 34 | namespace Gen 35 | 36 | @[inline] 37 | def up (x : Gen.{u} α) : Gen (ULift.{v} α) := do 38 | let size ← read 39 | Rand.up <| x.run ⟨size.down⟩ 40 | 41 | @[inline] 42 | def down (x : Gen (ULift.{v} α)) : Gen α := do 43 | let size ← read 44 | Rand.down <| x.run ⟨size.down⟩ 45 | 46 | /-- Lift `Random.random` to the `Gen` monad. -/ 47 | def chooseAny (α : Type u) [Random Id α] : Gen α := 48 | fun _ => rand α 49 | 50 | /-- Lift `BoundedRandom.randomR` to the `Gen` monad. -/ 51 | def choose (α : Type u) [LE α] [BoundedRandom Id α] (lo hi : α) (h : lo ≤ hi) : 52 | Gen {a // lo ≤ a ∧ a ≤ hi} := 53 | fun _ => randBound α lo hi h 54 | 55 | /-- Generate a `Nat` example between `lo` and `hi` (exclusively). -/ 56 | def chooseNatLt (lo hi : Nat) (h : lo < hi) : Gen {a // lo ≤ a ∧ a < hi} := do 57 | let ⟨val, h⟩ ← choose Nat (lo + 1) hi (by omega) 58 | return ⟨val - 1, by omega⟩ 59 | 60 | /-- Get access to the size parameter of the `Gen` monad. -/ 61 | def getSize : Gen Nat := 62 | return (← read).down 63 | 64 | /-- Apply a function to the size parameter. -/ 65 | def resize {α : Type _} (f : Nat → Nat) (x : Gen α) : Gen α := 66 | withReader (ULift.up ∘ f ∘ ULift.down) x 67 | 68 | /-- 69 | Choose a `Nat` between `0` and `getSize`. 70 | -/ 71 | def chooseNat : Gen Nat := do choose Nat 0 (← getSize) (by omega) 72 | 73 | variable {α : Type u} 74 | 75 | /-- Create an `Array` of examples using `x`. The size is controlled 76 | by the size parameter of `Gen`. -/ 77 | def arrayOf (x : Gen α) : Gen (Array α) := do 78 | let ⟨sz⟩ ← up chooseNat 79 | let mut res := Array.mkEmpty sz 80 | for _ in [0:sz] do 81 | res := res.push (← x) 82 | return res 83 | 84 | /-- Create a `List` of examples using `x`. The size is controlled 85 | by the size parameter of `Gen`. -/ 86 | def listOf (x : Gen α) : Gen (List α) := do 87 | let arr ← arrayOf x 88 | return arr.toList 89 | 90 | /-- Given a list of example generators, choose one to create an example. -/ 91 | def oneOf (xs : Array (Gen α)) (pos : 0 < xs.size := by decide) : Gen α := do 92 | let ⟨x, _, h2⟩ ← up <| chooseNatLt 0 xs.size pos 93 | xs[x] 94 | 95 | /-- Given a list of examples, choose one to create an example. -/ 96 | def elements (xs : List α) (pos : 0 < xs.length) : Gen α := do 97 | let ⟨x, _, h2⟩ ← up <| chooseNatLt 0 xs.length pos 98 | return xs[x] 99 | 100 | 101 | open List in 102 | /-- Generate a random permutation of a given list. -/ 103 | def permutationOf : (xs : List α) → Gen { ys // xs ~ ys } 104 | | [] => pure ⟨[], Perm.nil⟩ 105 | | x::xs => do 106 | let ⟨ys, h1⟩ ← permutationOf xs 107 | let ⟨n, _, h3⟩ ← up <| choose Nat 0 ys.length (by omega) 108 | return ⟨ys.insertIdx n x, Perm.trans (Perm.cons _ h1) (List.perm_insertIdx _ _ h3).symm⟩ 109 | 110 | /-- Given two generators produces a tuple consisting out of the result of both -/ 111 | def prodOf {α : Type u} {β : Type v} (x : Gen α) (y : Gen β) : Gen (α × β) := do 112 | let ⟨a⟩ ← up x 113 | let ⟨b⟩ ← up y 114 | return (a, b) 115 | 116 | end Gen 117 | 118 | /-- Execute a `Gen` inside the `IO` monad using `size` as the example size -/ 119 | def Gen.run {α : Type} (x : Gen α) (size : Nat) : BaseIO α := 120 | letI : MonadLift Id BaseIO := ⟨fun f => pure <| Id.run f⟩ 121 | runRand (ReaderT.run x ⟨size⟩:) 122 | 123 | end Plausible 124 | -------------------------------------------------------------------------------- /Plausible/Random.lean: -------------------------------------------------------------------------------- 1 | /- 2 | Copyright (c) 2022 Henrik Böving. All rights reserved. 3 | Released under Apache 2.0 license as described in the file LICENSE. 4 | Authors: Henrik Böving 5 | -/ 6 | 7 | /-! 8 | # Rand Monad and Random Class 9 | 10 | This module provides tools for formulating computations guided by randomness and for 11 | defining objects that can be created randomly. 12 | 13 | ## Main definitions 14 | 15 | * `RandT` and `RandGT` monad transformers for computations guided by randomness; 16 | * `Rand` and `RandG` monads as special cases of the above 17 | * `Random` class for objects that can be generated randomly; 18 | * `random` to generate one object; 19 | * `BoundedRandom` class for objects that can be generated randomly inside a range; 20 | * `randomR` to generate one object inside a range; 21 | * `runRand` to run a randomized computation inside any monad that has access to `stdGenRef`. 22 | 23 | ## References 24 | 25 | * Similar library in Haskell: https://hackage.haskell.org/package/MonadRandom 26 | 27 | -/ 28 | 29 | set_option linter.missingDocs true 30 | 31 | namespace Plausible 32 | 33 | /-- A monad transformer to generate random objects using the generic generator type `g` -/ 34 | abbrev RandGT (g : Type) := StateT (ULift g) 35 | /-- A monad to generate random objects using the generator type `g`. -/ 36 | abbrev RandG (g : Type) := RandGT g Id 37 | 38 | /-- A monad transformer to generate random objects using the generator type `StdGen`. 39 | `RandT m α` should be thought of a random value in `m α`. -/ 40 | abbrev RandT := RandGT StdGen 41 | 42 | /-- A monad to generate random objects using the generator type `StdGen`. -/ 43 | abbrev Rand := RandG StdGen 44 | 45 | instance [MonadLift m n] : MonadLiftT (RandGT g m) (RandGT g n) where 46 | monadLift x := fun s => x s 47 | 48 | /-- `Random m α` gives us machinery to generate values of type `α` in the monad `m`. 49 | 50 | Note that `m` is a parameter as some types may only be sampleable with access to a certain monad. -/ 51 | class Random (m) (α : Type u) where 52 | /-- 53 | Generate a value of type `α` randomly using generator `g`. 54 | -/ 55 | random [RandomGen g] : RandGT g m α 56 | 57 | /-- `BoundedRandom m α` gives us machinery to generate values of type `α` between certain bounds in 58 | the monad `m`. -/ 59 | class BoundedRandom (m) (α : Type u) [LE α] where 60 | /-- 61 | Generate a value of type `α` between `lo` and `hi` randomly using generator `g`. 62 | -/ 63 | randomR {g : Type} (lo hi : α) (h : lo ≤ hi) [RandomGen g] : RandGT g m {a // lo ≤ a ∧ a ≤ hi} 64 | 65 | /-- Given a random generator for `α`, we can convert it to a random generator for `ULift α`. -/ 66 | @[inline] 67 | protected def RandT.up {α : Type u} {m : Type u → Type w} {m' : Type (max u v) → Type w'} 68 | {g : Type} [RandomGen g] [Monad m] [Monad m'] 69 | (m_up : ∀ {α}, m α → m' (ULift α)) (x : RandGT g m α) : 70 | RandGT g m' (ULift.{v} α) := do 71 | let ⟨val, gen⟩ ← m_up <| x.run ⟨(← get).down⟩ 72 | set <| ULift.up gen.down 73 | return ⟨val⟩ 74 | 75 | /-- Given a random generator for `ULift α`, we can convert it to a random generator for `α`. -/ 76 | @[inline] 77 | protected def RandT.down {α : Type u} {m : Type (max u v) → Type w} {m' : Type u → Type w'} 78 | {g : Type} [RandomGen g] [Monad m] [Monad m'] 79 | (m_down : ∀ {α}, m (ULift α) → m' α) (x : RandGT g m (ULift.{v} α) ) : 80 | RandGT g m' α := do 81 | let gen := (← get).down 82 | let ⟨val, gen⟩ ← m_down do 83 | let ⟨⟨val⟩, ⟨gen⟩⟩ ← x.run ⟨gen⟩ 84 | pure <| .up (val, gen) 85 | set <| ULift.up gen 86 | return val 87 | 88 | namespace Rand 89 | /-- Generate one more `Nat` -/ 90 | def next [RandomGen g] [Monad m] : RandGT g m Nat := do 91 | let rng := (← get).down 92 | let (res, new) := RandomGen.next rng 93 | set (ULift.up new) 94 | return res 95 | 96 | /-- Create a new random number generator distinct from the one stored in the state -/ 97 | def split {g : Type} [RandomGen g] [Monad m] : RandGT g m g := do 98 | let rng := (← get).down 99 | let (r1, r2) := RandomGen.split rng 100 | set (ULift.up r1) 101 | return r2 102 | 103 | /-- Get the range of Nat that can be generated by the generator `g` -/ 104 | def range {g : Type} [RandomGen g] [Monad m] : RandGT g m (Nat × Nat) := do 105 | let rng := (← get).down 106 | return RandomGen.range rng 107 | 108 | /-- Given a random generator for `α`, we can convert it to a random generator for `ULift α`. -/ 109 | @[inline] 110 | protected def up {α : Type u} {g : Type} [RandomGen g] (x : RandG g α) : 111 | RandG g (ULift.{v} α) := do 112 | RandT.up (fun x => pure ⟨Id.run x⟩) x 113 | 114 | /-- Given a random generator for `ULift α`, we can convert it to a random generator for `α`. -/ 115 | @[inline] 116 | protected def down {α : Type u} {g : Type} [RandomGen g] (x : RandG g (ULift.{v} α)) : 117 | RandG g α := 118 | RandT.down (fun x => pure (Id.run x).down) x 119 | 120 | end Rand 121 | 122 | namespace Random 123 | 124 | open Rand 125 | 126 | variable [Monad m] 127 | 128 | /-- Generate a random value of type `α`. -/ 129 | def rand (α : Type u) [Random m α] [RandomGen g] : RandGT g m α := Random.random 130 | 131 | /-- Generate a random value of type `α` between `x` and `y` inclusive. -/ 132 | def randBound (α : Type u) [LE α] [BoundedRandom m α] (lo hi : α) (h : lo ≤ hi) [RandomGen g] : 133 | RandGT g m {a // lo ≤ a ∧ a ≤ hi} := 134 | BoundedRandom.randomR lo hi h 135 | 136 | /-- 137 | Generate a random `Fin`. 138 | -/ 139 | def randFin {n : Nat} [RandomGen g] : RandGT g m (Fin n.succ) := 140 | fun ⟨g⟩ => return randNat g 0 n |>.map (Fin.ofNat _) ULift.up 141 | 142 | instance {n : Nat} : Random m (Fin n.succ) where 143 | random := randFin 144 | 145 | /-- 146 | Generate a random `Bool`. 147 | -/ 148 | def randBool [RandomGen g] : RandGT g m Bool := 149 | return (← rand (Fin 2)) == 1 150 | 151 | instance : Random m Bool where 152 | random := randBool 153 | 154 | instance : BoundedRandom m Nat where 155 | randomR lo hi h _ := do 156 | let z ← rand (Fin (hi - lo).succ) 157 | return ⟨lo + z.val, by omega, by omega⟩ 158 | 159 | instance : BoundedRandom m Int where 160 | randomR lo hi h _ := do 161 | let ⟨z, _, h2⟩ ← randBound Nat 0 (Int.natAbs <| hi - lo) (by omega) 162 | return ⟨z + lo, by omega, by omega⟩ 163 | 164 | instance {n : Nat} : BoundedRandom m (Fin n) where 165 | randomR lo hi h _ := do 166 | let ⟨r, h1, h2⟩ ← randBound Nat lo.val hi.val h 167 | return ⟨⟨r, by omega⟩, h1, h2⟩ 168 | 169 | instance {n : Nat} : BoundedRandom m (BitVec n) where 170 | randomR lo hi h _ := do 171 | let ⟨r, h1, h2⟩ ← randBound Nat lo.toNat hi.toNat h 172 | return ⟨⟨r, by omega⟩, h1, h2⟩ 173 | 174 | end Random 175 | 176 | open IO 177 | variable {m : Type _ → Type _} 178 | variable [Monad m] [MonadLiftT (ST RealWorld) m] 179 | 180 | /-- 181 | Computes a `RandT m α` using the global `stdGenRef` as RNG. 182 | 183 | Note that: 184 | - `stdGenRef` is not necessarily properly seeded on program startup 185 | as of now and will therefore be deterministic. 186 | - `stdGenRef` is not thread local, hence two threads accessing it 187 | at the same time will get the exact same generator. 188 | -/ 189 | def runRand (cmd : RandT m α) : m α := do 190 | let stdGen ← stdGenRef.get 191 | let (res, new) ← StateT.run cmd (ULift.up stdGen) 192 | let _ ← stdGenRef.set new.down 193 | return res 194 | 195 | /-- 196 | Run the random computaton `cmd` with `seed` for the RNG. 197 | -/ 198 | def runRandWith (seed : Nat) (cmd : RandT m α) : m α := do 199 | return (← cmd.run (ULift.up <| mkStdGen seed)).1 200 | 201 | end Plausible 202 | -------------------------------------------------------------------------------- /Plausible/Sampleable.lean: -------------------------------------------------------------------------------- 1 | /- 2 | Copyright (c) 2022 Henrik Böving. All rights reserved. 3 | Released under Apache 2.0 license as described in the file LICENSE. 4 | Authors: Henrik Böving, Simon Hudon 5 | -/ 6 | import Lean.Elab.Command 7 | import Lean.Meta.Eval 8 | import Plausible.Gen 9 | 10 | /-! 11 | # `SampleableExt` Class 12 | 13 | This class permits the creation samples of a given type 14 | controlling the size of those values using the `Gen` monad. 15 | 16 | # `Shrinkable` Class 17 | 18 | This class helps minimize examples by creating smaller versions of 19 | given values. 20 | 21 | When testing a proposition like `∀ n : Nat, Prime n → n ≤ 100`, 22 | `Plausible` requires that `Nat` have an instance of `SampleableExt` and for 23 | `Prime n` to be decidable. `Plausible` will then use the instance of 24 | `SampleableExt` to generate small examples of Nat and progressively increase 25 | in size. For each example `n`, `Prime n` is tested. If it is false, 26 | the example will be rejected (not a test success nor a failure) and 27 | `Plausible` will move on to other examples. If `Prime n` is true, 28 | `n ≤ 100` will be tested. If it is false, `n` is a counter-example of 29 | `∀ n : Nat, Prime n → n ≤ 100` and the test fails. If `n ≤ 100` is true, 30 | the test passes and `Plausible` moves on to trying more examples. 31 | 32 | ## Main definitions 33 | 34 | * `SampleableExt` class 35 | * `Shrinkable` class 36 | 37 | ### `SampleableExt` 38 | 39 | `SampleableExt` can be used in two ways. The first (and most common) 40 | is to simply generate values of a type directly using the `Gen` monad, 41 | if this is what you want to do then `SampleableExt.mkSelfContained` is 42 | the way to go. 43 | 44 | Furthermore it makes it possible to express generators for types that 45 | do not lend themselves to introspection, such as `Nat → Nat`. 46 | If we test a quantification over functions the 47 | counter-examples cannot be shrunken or printed meaningfully. 48 | For that purpose, `SampleableExt` provides a proxy representation 49 | `proxy` that can be printed and shrunken as well 50 | as interpreted (using `interp`) as an object of the right type. If you 51 | are using it in the first way, this proxy type will simply be the type 52 | itself and the `interp` function `id`. 53 | 54 | ### `Shrinkable` 55 | 56 | Given an example `x : α`, `Shrinkable α` gives us a way to shrink it 57 | and suggest simpler examples. 58 | 59 | ## Shrinking 60 | 61 | Shrinking happens when `Plausible` finds a counter-example to a 62 | property. It is likely that the example will be more complicated than 63 | necessary so `Plausible` proceeds to shrink it as much as 64 | possible. Although equally valid, a smaller counter-example is easier 65 | for a user to understand and use. 66 | 67 | The `Shrinkable` class, , has a `shrink` function so that we can use 68 | specialized knowledge while shrinking a value. It is not responsible 69 | for the whole shrinking process however. It only has to take one step 70 | in the shrinking process. `Plausible` will repeatedly call `shrink` 71 | until no more steps can be taken. Because `shrink` guarantees that the 72 | size of the candidates it produces is strictly smaller than the 73 | argument, we know that `Plausible` is guaranteed to terminate. 74 | 75 | ## Tags 76 | 77 | random testing 78 | 79 | ## References 80 | 81 | * https://hackage.haskell.org/package/QuickCheck 82 | 83 | -/ 84 | 85 | namespace Plausible 86 | 87 | open Random Gen 88 | 89 | universe u v 90 | variable {α β : Type _} 91 | 92 | /-- Given an example `x : α`, `Shrinkable α` gives us a way to shrink it 93 | and suggest simpler examples. -/ 94 | class Shrinkable (α : Type u) where 95 | shrink : (x : α) → List α := fun _ => [] 96 | 97 | /-- `SampleableExt` can be used in two ways. The first (and most common) 98 | is to simply generate values of a type directly using the `Gen` monad, 99 | if this is what you want to do then `SampleableExt.mkSelfContained` is 100 | the way to go. 101 | 102 | Furthermore it makes it possible to express generators for types that 103 | do not lend themselves to introspection, such as `Nat → Nat`. 104 | If we test a quantification over functions the 105 | counter-examples cannot be shrunken or printed meaningfully. 106 | For that purpose, `SampleableExt` provides a proxy representation 107 | `proxy` that can be printed and shrunken as well 108 | as interpreted (using `interp`) as an object of the right type. -/ 109 | class SampleableExt (α : Sort u) where 110 | proxy : Type v 111 | [proxyRepr : Repr proxy] 112 | [shrink : Shrinkable proxy] 113 | sample : Gen proxy 114 | interp : proxy → α 115 | 116 | attribute [instance] SampleableExt.proxyRepr 117 | attribute [instance] SampleableExt.shrink 118 | 119 | namespace SampleableExt 120 | 121 | /-- Use to generate instance whose purpose is to simply generate values 122 | of a type directly using the `Gen` monad -/ 123 | def mkSelfContained [Repr α] [Shrinkable α] (sample : Gen α) : SampleableExt α where 124 | proxy := α 125 | proxyRepr := inferInstance 126 | shrink := inferInstance 127 | sample := sample 128 | interp := id 129 | 130 | /-- First samples a proxy value and interprets it. Especially useful if 131 | the proxy and target type are the same. -/ 132 | def interpSample (α : Type u) [SampleableExt α] : Gen α := 133 | SampleableExt.interp <$> SampleableExt.sample 134 | 135 | end SampleableExt 136 | 137 | section Shrinkers 138 | 139 | instance [Shrinkable α] [Shrinkable β] : Shrinkable (Sum α β) where 140 | shrink s := 141 | match s with 142 | | .inl l => Shrinkable.shrink l |>.map .inl 143 | | .inr r => Shrinkable.shrink r |>.map .inr 144 | 145 | instance Unit.shrinkable : Shrinkable Unit where 146 | shrink _ := [] 147 | 148 | /-- `Nat.shrink' n` creates a list of smaller natural numbers by 149 | successively dividing `n` by 2 . For example, `Nat.shrink 5 = [2, 1, 0]`. -/ 150 | def Nat.shrink (n : Nat) : List Nat := 151 | if h : 0 < n then 152 | let m := n/2 153 | m :: shrink m 154 | else 155 | [] 156 | 157 | instance Nat.shrinkable : Shrinkable Nat where 158 | shrink := Nat.shrink 159 | 160 | instance Fin.shrinkable {n : Nat} : Shrinkable (Fin n.succ) where 161 | shrink m := Nat.shrink m |>.map (Fin.ofNat _) 162 | 163 | instance BitVec.shrinkable {n : Nat} : Shrinkable (BitVec n) where 164 | shrink m := Nat.shrink m.toNat |>.map (BitVec.ofNat n) 165 | 166 | instance UInt8.shrinkable : Shrinkable UInt8 where 167 | shrink m := Nat.shrink m.toNat |>.map UInt8.ofNat 168 | 169 | instance UInt16.shrinkable : Shrinkable UInt16 where 170 | shrink m := Nat.shrink m.toNat |>.map UInt16.ofNat 171 | 172 | instance UInt32.shrinkable : Shrinkable UInt32 where 173 | shrink m := Nat.shrink m.toNat |>.map UInt32.ofNat 174 | 175 | instance UInt64.shrinkable : Shrinkable UInt64 where 176 | shrink m := Nat.shrink m.toNat |>.map UInt64.ofNat 177 | 178 | instance USize.shrinkable : Shrinkable USize where 179 | shrink m := Nat.shrink m.toNat |>.map USize.ofNat 180 | 181 | /-- `Int.shrinkable` operates like `Nat.shrinkable` but also includes the negative variants. -/ 182 | instance Int.shrinkable : Shrinkable Int where 183 | shrink n := 184 | let converter n := 185 | let int := Int.ofNat n 186 | [int, -int] 187 | Nat.shrink n.natAbs |>.flatMap converter 188 | 189 | instance Bool.shrinkable : Shrinkable Bool := {} 190 | instance Char.shrinkable : Shrinkable Char := {} 191 | 192 | instance Option.shrinkable [Shrinkable α] : Shrinkable (Option α) where 193 | shrink o := 194 | match o with 195 | | some x => Shrinkable.shrink x |>.map .some 196 | | none => [] 197 | 198 | instance Prod.shrinkable [shrA : Shrinkable α] [shrB : Shrinkable β] : 199 | Shrinkable (Prod α β) where 200 | shrink := fun (fst,snd) => 201 | let shrink1 := shrA.shrink fst |>.map fun x => (x, snd) 202 | let shrink2 := shrB.shrink snd |>.map fun x => (fst, x) 203 | shrink1 ++ shrink2 204 | 205 | instance Sigma.shrinkable [shrA : Shrinkable α] [shrB : Shrinkable β] : 206 | Shrinkable ((_ : α) × β) where 207 | shrink := fun ⟨fst,snd⟩ => 208 | let shrink1 := shrA.shrink fst |>.map fun x => ⟨x, snd⟩ 209 | let shrink2 := shrB.shrink snd |>.map fun x => ⟨fst, x⟩ 210 | shrink1 ++ shrink2 211 | 212 | open Shrinkable 213 | 214 | /-- Shrink a list of a shrinkable type, either by discarding an element or shrinking an element. -/ 215 | instance List.shrinkable [Shrinkable α] : Shrinkable (List α) where 216 | shrink := fun L => 217 | (L.mapIdx fun i _ => L.eraseIdx i) ++ 218 | (L.mapIdx fun i a => (shrink a).map fun a' => L.modify i fun _ => a').flatten 219 | 220 | instance ULift.shrinkable [Shrinkable α] : Shrinkable (ULift α) where 221 | shrink u := (shrink u.down).map ULift.up 222 | 223 | instance String.shrinkable : Shrinkable String where 224 | shrink s := (shrink s.toList).map String.mk 225 | 226 | instance Array.shrinkable [Shrinkable α] : Shrinkable (Array α) where 227 | shrink xs := (shrink xs.toList).map Array.mk 228 | 229 | instance Subtype.shrinkable {α : Type u} {β : α → Prop} [Shrinkable α] [∀ x, Decidable (β x)] : Shrinkable {x : α // β x} where 230 | shrink x := 231 | let val := x.val 232 | let candidates := shrink val 233 | let filter x := do 234 | if h : β x then 235 | some ⟨x, h⟩ 236 | else 237 | none 238 | candidates.filterMap filter 239 | 240 | end Shrinkers 241 | 242 | section Samplers 243 | 244 | open SampleableExt 245 | 246 | instance Sum.SampleableExt [SampleableExt α] [SampleableExt β] : SampleableExt (Sum α β) where 247 | proxy := Sum (proxy α) (proxy β) 248 | sample := do 249 | match ← chooseAny Bool with 250 | | true => return .inl (← sample) 251 | | false => return .inr (← sample) 252 | interp s := 253 | match s with 254 | | .inl l => .inl (interp l) 255 | | .inr r => .inr (interp r) 256 | 257 | instance Unit.sampleableExt : SampleableExt Unit := 258 | mkSelfContained (return ()) 259 | 260 | instance [SampleableExt α] [SampleableExt β] : SampleableExt ((_ : α) × β) where 261 | proxy := (_ : proxy α) × proxy β 262 | sample := do 263 | let p ← prodOf sample sample 264 | return ⟨p.fst, p.snd⟩ 265 | interp s := ⟨interp s.fst, interp s.snd⟩ 266 | 267 | instance Nat.sampleableExt : SampleableExt Nat := 268 | mkSelfContained (do choose Nat 0 (← getSize) (Nat.zero_le _)) 269 | 270 | instance Fin.sampleableExt {n : Nat} : SampleableExt (Fin (n.succ)) := 271 | mkSelfContained do 272 | let m ← choose Nat 0 (min (← getSize) n) (Nat.zero_le _) 273 | return (Fin.ofNat _ m) 274 | 275 | instance BitVec.sampleableExt {n : Nat} : SampleableExt (BitVec n) := 276 | mkSelfContained do 277 | let m ← choose Nat 0 (min (← getSize) (2^n)) (Nat.zero_le _) 278 | return BitVec.ofNat _ m 279 | 280 | instance UInt8.SampleableExt : SampleableExt UInt8 := 281 | mkSelfContained do 282 | let n ← choose Nat 0 (min (← getSize) UInt8.size) (Nat.zero_le _) 283 | return UInt8.ofNat n 284 | 285 | instance UInt16.SampleableExt : SampleableExt UInt16 := 286 | mkSelfContained do 287 | let n ← choose Nat 0 (min (← getSize) UInt16.size) (Nat.zero_le _) 288 | return UInt16.ofNat n 289 | 290 | instance UInt32.SampleableExt : SampleableExt UInt32 := 291 | mkSelfContained do 292 | let n ← choose Nat 0 (min (← getSize) UInt32.size) (Nat.zero_le _) 293 | return UInt32.ofNat n 294 | 295 | instance UInt64.SampleableExt : SampleableExt UInt64 := 296 | mkSelfContained do 297 | let n ← choose Nat 0 (min (← getSize) UInt64.size) (Nat.zero_le _) 298 | return UInt64.ofNat n 299 | 300 | instance USize.SampleableExt : SampleableExt USize := 301 | mkSelfContained do 302 | let n ← choose Nat 0 (min (← getSize) USize.size) (Nat.zero_le _) 303 | return USize.ofNat n 304 | 305 | instance Int.sampleableExt : SampleableExt Int := 306 | mkSelfContained do 307 | choose Int (-(← getSize)) (← getSize) (by omega) 308 | 309 | instance Bool.sampleableExt : SampleableExt Bool := 310 | mkSelfContained <| chooseAny Bool 311 | 312 | /-- This can be specialized into customized `SampleableExt Char` instances. 313 | The resulting instance has `1 / length` chances of making an unrestricted choice of characters 314 | and it otherwise chooses a character from `chars` with uniform probabilities. -/ 315 | def Char.sampleable (length : Nat) (chars : List Char) (pos : 0 < chars.length) : 316 | SampleableExt Char := 317 | mkSelfContained do 318 | let x ← choose Nat 0 length (Nat.zero_le _) 319 | if x.val == 0 then 320 | let n ← interpSample Nat 321 | pure <| Char.ofNat n 322 | else 323 | elements chars pos 324 | 325 | instance Char.sampleableDefault : SampleableExt Char := 326 | Char.sampleable 3 " 0123abcABC:,;`\\/".toList (by decide) 327 | 328 | instance Option.sampleableExt [SampleableExt α] : SampleableExt (Option α) where 329 | proxy := Option (proxy α) 330 | sample := do 331 | match ← chooseAny Bool with 332 | | true => return none 333 | | false => return some (← sample) 334 | interp o := o.map interp 335 | 336 | instance Prod.sampleableExt {α : Type u} {β : Type v} [SampleableExt α] [SampleableExt β] : 337 | SampleableExt (α × β) where 338 | proxy := Prod (proxy α) (proxy β) 339 | proxyRepr := inferInstance 340 | shrink := inferInstance 341 | sample := prodOf sample sample 342 | interp := Prod.map interp interp 343 | 344 | instance Prop.sampleableExt : SampleableExt Prop where 345 | proxy := Bool 346 | proxyRepr := inferInstance 347 | sample := interpSample Bool 348 | shrink := inferInstance 349 | interp := Coe.coe 350 | 351 | instance List.sampleableExt [SampleableExt α] : SampleableExt (List α) where 352 | proxy := List (proxy α) 353 | sample := Gen.listOf sample 354 | interp := List.map interp 355 | 356 | instance ULift.sampleableExt [SampleableExt α] : SampleableExt (ULift α) where 357 | proxy := proxy α 358 | sample := sample 359 | interp a := ⟨interp a⟩ 360 | 361 | instance String.sampleableExt : SampleableExt String := 362 | mkSelfContained do return String.mk (← Gen.listOf (Char.sampleableDefault.sample)) 363 | 364 | instance Array.sampleableExt [SampleableExt α] : SampleableExt (Array α) where 365 | proxy := Array (proxy α) 366 | sample := Gen.arrayOf sample 367 | interp := Array.map interp 368 | 369 | end Samplers 370 | 371 | /-- An annotation for values that should never get shrunk. -/ 372 | def NoShrink (α : Type u) := α 373 | 374 | namespace NoShrink 375 | 376 | def mk (x : α) : NoShrink α := x 377 | def get (x : NoShrink α) : α := x 378 | 379 | instance inhabited [inst : Inhabited α] : Inhabited (NoShrink α) := inst 380 | instance repr [inst : Repr α] : Repr (NoShrink α) := inst 381 | 382 | instance shrinkable : Shrinkable (NoShrink α) where 383 | shrink := fun _ => [] 384 | 385 | instance sampleableExt [SampleableExt α] [Repr α] : SampleableExt (NoShrink α) := 386 | SampleableExt.mkSelfContained <| (NoShrink.mk ∘ SampleableExt.interp) <$> SampleableExt.sample 387 | 388 | end NoShrink 389 | 390 | /-- 391 | Print (at most) 10 samples of a given type to stdout for debugging. 392 | -/ 393 | def printSamples {t : Type u} [Repr t] (g : Gen t) : IO PUnit := do 394 | -- TODO: this should be a global instance 395 | letI : MonadLift Id IO := ⟨fun f => pure <| Id.run f⟩ 396 | do 397 | -- we can't convert directly from `Rand (List t)` to `RandT IO (List Std.Format)` 398 | -- (and `RandT IO (List t)` isn't type-correct without 399 | -- https://github.com/leanprover/lean4/issues/3011), so go via an intermediate 400 | let xs : List Std.Format ← Plausible.runRand <| Rand.down <| do 401 | let xs : List t ← (List.range 10).mapM (ReaderT.run g ∘ ULift.up) 402 | pure <| ULift.up (xs.map repr) 403 | for x in xs do 404 | IO.println s!"{x}" 405 | 406 | open Lean Meta Elab 407 | 408 | /-- 409 | `e` is a type to sample from, this can either be a type that implements `SampleableExt` or `Gen α` 410 | directly. For this return: 411 | - the universe level of the `Type u` that the relevant type to sample lives in. 412 | - the actual type `α` to sample from 413 | - a `Repr α` instance 414 | - a `Gen α` generator to run in order to sample 415 | -/ 416 | private def mkGenerator (e : Expr) : MetaM (Level × Expr × Expr × Expr) := do 417 | let exprTyp ← inferType e 418 | let .sort u ← whnf (← inferType exprTyp) | throwError m!"{exprTyp} is not a type" 419 | let .succ u := u | throwError m!"{exprTyp} is not a type with computational content" 420 | match_expr exprTyp with 421 | | Gen α => 422 | let reprInst ← synthInstance (mkApp (mkConst ``Repr [u]) α) 423 | return ⟨u, α, reprInst, e⟩ 424 | | _ => 425 | let v ← mkFreshLevelMVar 426 | let sampleableExtInst ← synthInstance (mkApp (mkConst ``SampleableExt [u, v]) e) 427 | let v ← instantiateLevelMVars v 428 | let reprInst := mkApp2 (mkConst ``SampleableExt.proxyRepr [u, v]) e sampleableExtInst 429 | let gen := mkApp2 (mkConst ``SampleableExt.sample [u, v]) e sampleableExtInst 430 | let typ := mkApp2 (mkConst ``SampleableExt.proxy [u, v]) e sampleableExtInst 431 | return ⟨v, typ, reprInst, gen⟩ 432 | 433 | /-- 434 | `#sample type`, where `type` has an instance of `SampleableExt`, prints ten random 435 | values of type `type` using an increasing size parameter. 436 | 437 | ```lean 438 | #sample Nat 439 | -- prints 440 | -- 0 441 | -- 0 442 | -- 2 443 | -- 24 444 | -- 64 445 | -- 76 446 | -- 5 447 | -- 132 448 | -- 8 449 | -- 449 450 | -- or some other sequence of numbers 451 | 452 | #sample List Int 453 | -- prints 454 | -- [] 455 | -- [1, 1] 456 | -- [-7, 9, -6] 457 | -- [36] 458 | -- [-500, 105, 260] 459 | -- [-290] 460 | -- [17, 156] 461 | -- [-2364, -7599, 661, -2411, -3576, 5517, -3823, -968] 462 | -- [-643] 463 | -- [11892, 16329, -15095, -15461] 464 | -- or whatever 465 | ``` 466 | -/ 467 | elab "#sample " e:term : command => 468 | Command.runTermElabM fun _ => do 469 | let e ← Elab.Term.elabTermAndSynthesize e none 470 | let ⟨u, α, repr, gen⟩ ← mkGenerator e 471 | let printSamples := mkApp3 (mkConst ``printSamples [u]) α repr gen 472 | let code ← unsafe evalExpr (IO PUnit) (mkApp (mkConst ``IO) (mkConst ``PUnit [1])) printSamples 473 | _ ← code 474 | 475 | end Plausible 476 | -------------------------------------------------------------------------------- /Plausible/Tactic.lean: -------------------------------------------------------------------------------- 1 | /- 2 | Copyright (c) 2020 Simon Hudon. All rights reserved. 3 | Released under Apache 2.0 license as described in the file LICENSE. 4 | Authors: Simon Hudon, Kim Morrison 5 | -/ 6 | 7 | import Plausible.Testable 8 | import Plausible.Attr 9 | 10 | /-! 11 | ## Finding counterexamples automatically using `plausible` 12 | 13 | A proposition can be tested by writing it out as: 14 | 15 | ```lean 16 | example (xs : List Nat) (w : ∃ x ∈ xs, x < 3) : ∀ y ∈ xs, y < 5 := by plausible 17 | -- =================== 18 | -- Found problems! 19 | 20 | -- xs := [0, 5] 21 | -- x := 0 22 | -- y := 5 23 | -- ------------------- 24 | 25 | example (x : Nat) (h : 2 ∣ x) : x < 100 := by plausible 26 | -- =================== 27 | -- Found problems! 28 | 29 | -- x := 258 30 | -- ------------------- 31 | 32 | example (α : Type) (xs ys : List α) : xs ++ ys = ys ++ xs := by plausible 33 | -- =================== 34 | -- Found problems! 35 | 36 | -- α := Int 37 | -- xs := [-4] 38 | -- ys := [1] 39 | -- ------------------- 40 | 41 | example : ∀ x ∈ [1,2,3], x < 4 := by plausible 42 | -- Success 43 | ``` 44 | 45 | In the first example, `plausible` is called on the following goal: 46 | 47 | ```lean 48 | xs : List Nat, 49 | h : ∃ (x : Nat) (H : x ∈ xs), x < 3 50 | ⊢ ∀ (y : Nat), y ∈ xs → y < 5 51 | ``` 52 | 53 | The local constants are reverted and an instance is found for 54 | `Testable (∀ (xs : List Nat), (∃ x ∈ xs, x < 3) → (∀ y ∈ xs, y < 5))`. 55 | The `Testable` instance is supported by instances of `Sampleable (List Nat)`, 56 | `Decidable (x < 3)` and `Decidable (y < 5)`. `plausible` builds a 57 | `Testable` instance step by step with: 58 | 59 | ``` 60 | - Testable (∀ (xs : List Nat), (∃ x ∈ xs, x < 3) → (∀ y ∈ xs, y < 5)) 61 | -: Sampleable (List xs) 62 | - Testable ((∃ x ∈ xs, x < 3) → (∀ y ∈ xs, y < 5)) 63 | - Testable (∀ x ∈ xs, x < 3 → (∀ y ∈ xs, y < 5)) 64 | - Testable (x < 3 → (∀ y ∈ xs, y < 5)) 65 | -: Decidable (x < 3) 66 | - Testable (∀ y ∈ xs, y < 5) 67 | -: Decidable (y < 5) 68 | ``` 69 | 70 | `Sampleable (List Nat)` lets us create random data of type `List Nat` in a way that 71 | helps find small counter-examples. Next, the test of the proposition 72 | hinges on `x < 3` and `y < 5` to both be decidable. The 73 | implication between the two could be tested as a whole but it would be 74 | less informative. Indeed, if we generate lists that only contain numbers 75 | greater than `3`, the implication will always trivially hold but we should 76 | conclude that we haven't found meaningful examples. Instead, when `x < 3` 77 | does not hold, we reject the example (i.e. we do not count it toward 78 | the 100 required positive examples) and we start over. Therefore, when 79 | `plausible` prints `Success`, it means that a hundred suitable lists 80 | were found and successfully tested. 81 | 82 | If no counter-examples are found, `plausible` behaves like `admit`. 83 | 84 | `plausible` can also be invoked using `#eval`: 85 | 86 | ```lean 87 | #eval Plausible.Testable.check (∀ (α : Type) (xs ys : List α), xs ++ ys = ys ++ xs) 88 | -- =================== 89 | -- Found problems! 90 | 91 | -- α := Int 92 | -- xs := [-4] 93 | -- ys := [1] 94 | -- ------------------- 95 | ``` 96 | 97 | For more information on writing your own `Sampleable` and `Testable` 98 | instances, see `Testing.Plausible.Testable`. 99 | -/ 100 | 101 | open Lean Elab Meta Tactic 102 | open Parser.Tactic 103 | open Plausible Decorations 104 | 105 | /-- 106 | `plausible` considers a proof goal and tries to generate examples 107 | that would contradict the statement. 108 | 109 | Let's consider the following proof goal. 110 | 111 | ```lean 112 | xs : List Nat, 113 | h : ∃ (x : Nat) (H : x ∈ xs), x < 3 114 | ⊢ ∀ (y : Nat), y ∈ xs → y < 5 115 | ``` 116 | 117 | The local constants will be reverted and an instance will be found for 118 | `Testable (∀ (xs : List Nat), (∃ x ∈ xs, x < 3) → (∀ y ∈ xs, y < 5))`. 119 | The `Testable` instance is supported by an instance of `Sampleable (List Nat)`, 120 | `Decidable (x < 3)` and `Decidable (y < 5)`. 121 | 122 | Examples will be created in ascending order of size (more or less) 123 | 124 | The first counter-examples found will be printed and will result in an error: 125 | 126 | ``` 127 | =================== 128 | Found problems! 129 | xs := [1, 28] 130 | x := 1 131 | y := 28 132 | ------------------- 133 | ``` 134 | 135 | If `plausible` successfully tests 100 examples, it acts like 136 | admit. If it gives up or finds a counter-example, it reports an error. 137 | 138 | For more information on writing your own `Sampleable` and `Testable` 139 | instances, see `Testing.Plausible.Testable`. 140 | 141 | Optional arguments given with `plausible (config : { ... })` 142 | * `numInst` (default 100): number of examples to test properties with 143 | * `maxSize` (default 100): final size argument 144 | 145 | Options: 146 | * `set_option trace.plausible.decoration true`: print the proposition with quantifier annotations 147 | * `set_option trace.plausible.discarded true`: print the examples discarded because they do not 148 | satisfy assumptions 149 | * `set_option trace.plausible.shrink.steps true`: trace the shrinking of counter-example 150 | * `set_option trace.plausible.shrink.candidates true`: print the lists of candidates considered 151 | when shrinking each variable 152 | * `set_option trace.plausible.instance true`: print the instances of `testable` being used to test 153 | the proposition 154 | * `set_option trace.plausible.success true`: print the tested samples that satisfy a property 155 | -/ 156 | syntax (name := plausibleSyntax) "plausible" (config)? : tactic 157 | 158 | elab_rules : tactic | `(tactic| plausible $[$cfg]?) => withMainContext do 159 | let cfg ← elabConfig (mkOptionalNode cfg) 160 | let (_, g) ← (← getMainGoal).revert ((← getLocalHyps).map (Expr.fvarId!)) 161 | g.withContext do 162 | let tgt ← g.getType 163 | let tgt' ← addDecorations tgt 164 | let cfg := { cfg with 165 | traceDiscarded := cfg.traceDiscarded || (← isTracingEnabledFor `plausible.discarded), 166 | traceSuccesses := cfg.traceSuccesses || (← isTracingEnabledFor `plausible.success), 167 | traceShrink := cfg.traceShrink || (← isTracingEnabledFor `plausible.shrink.steps), 168 | traceShrinkCandidates := cfg.traceShrinkCandidates 169 | || (← isTracingEnabledFor `plausible.shrink.candidates) } 170 | let inst ← try 171 | synthInstance (← mkAppM ``Testable #[tgt']) 172 | catch _ => throwError "\ 173 | Failed to create a `testable` instance for `{tgt}`.\ 174 | \nWhat to do:\ 175 | \n1. make sure that the types you are using have `Plausible.SampleableExt` instances\ 176 | \n (you can use `#sample my_type` if you are unsure);\ 177 | \n2. make sure that the relations and predicates that your proposition use are decidable;\ 178 | \n3. make sure that instances of `Plausible.Testable` exist that, when combined,\ 179 | \n apply to your decorated proposition:\ 180 | \n```\ 181 | \n{tgt'}\ 182 | \n```\ 183 | \n\ 184 | \nUse `set_option trace.Meta.synthInstance true` to understand what instances are missing.\ 185 | \n\ 186 | \nTry this:\ 187 | \nset_option trace.Meta.synthInstance true\ 188 | \n#synth Plausible.Testable ({tgt'})" 189 | let e ← mkAppOptM ``Testable.check #[tgt, toExpr cfg, tgt', inst] 190 | trace[plausible.decoration] "[testable decoration]\n {tgt'}" 191 | -- Porting note: I have not ported support for `trace.plausible.instance`. 192 | -- See the commented out code below from mathlib3 if you would like to implement this. 193 | -- when_tracing `plausible.instance <| do 194 | -- { inst ← summarize_instance inst >>= pp, 195 | -- trace!"\n[testable instance]{format.indent inst 2}" }, 196 | let code ← unsafe evalExpr (CoreM PUnit) (mkApp (mkConst ``CoreM) (mkConst ``PUnit [1])) e 197 | _ ← code 198 | admitGoal g 199 | 200 | -- Porting note: below is the remaining code from mathlib3 which supports the 201 | -- `trace.plausible.instance` trace option, and which has not been ported. 202 | 203 | -- namespace tactic.interactive 204 | -- open tactic plausible 205 | 206 | -- open expr 207 | 208 | -- /-- Tree structure representing a `testable` instance. -/ 209 | -- meta inductive instance_tree 210 | -- | node : name → expr → list instance_tree → instance_tree 211 | 212 | -- /-- Gather information about a `testable` instance. Given 213 | -- an expression of type `testable ?p`, gather the 214 | -- name of the `testable` instances that it is built from 215 | -- and the proposition that they test. -/ 216 | -- meta def summarize_instance : expr → tactic instance_tree 217 | -- | (lam n bi d b) := do 218 | -- v ← mk_local' n bi d, 219 | -- summarize_instance <| b.instantiate_var v 220 | -- | e@(app f x) := do 221 | -- `(testable %%p) ← infer_type e, 222 | -- xs ← e.get_app_args.mmap_filter (try_core ∘ summarize_instance), 223 | -- pure <| instance_tree.node e.get_app_fn.const_name p xs 224 | -- | e := do 225 | -- failed 226 | 227 | -- /-- format an `instance_tree` -/ 228 | -- meta def instance_tree.to_format : instance_tree → tactic format 229 | -- | (instance_tree.node n p xs) := do 230 | -- xs ← format.join <$> (xs.mmap <| λ t, flip format.indent 2 <$> instance_tree.to_format t), 231 | -- ys ← pformat!"testable ({p})", 232 | -- pformat!"+ {n} :{format.indent ys 2}\n{xs}" 233 | 234 | -- meta instance instance_tree.has_to_tactic_format : has_to_tactic_format instance_tree := 235 | -- ⟨ instance_tree.to_format ⟩ 236 | -------------------------------------------------------------------------------- /Plausible/Testable.lean: -------------------------------------------------------------------------------- 1 | /- 2 | Copyright (c) 2022 Henrik Böving. All rights reserved. 3 | Released under Apache 2.0 license as described in the file LICENSE. 4 | Authors: Henrik Böving, Simon Hudon 5 | -/ 6 | import Lean.Elab.Tactic.Config 7 | import Plausible.Sampleable 8 | 9 | 10 | /-! 11 | # `Testable` Class 12 | 13 | Testable propositions have a procedure that can generate counter-examples 14 | together with a proof that they invalidate the proposition. 15 | 16 | This is a port of the Haskell QuickCheck library. 17 | 18 | ## Creating Customized Instances 19 | 20 | The type classes `Testable`, `SampleableExt` and `Shrinkable` are the 21 | means by which `Plausible` creates samples and tests them. For instance, 22 | the proposition `∀ i j : Nat, i ≤ j` has a `Testable` instance because `Nat` 23 | is sampleable and `i ≤ j` is decidable. Once `Plausible` finds the `Testable` 24 | instance, it can start using the instance to repeatedly creating samples 25 | and checking whether they satisfy the property. Once it has found a 26 | counter-example it will then use a `Shrinkable` instance to reduce the 27 | example. This allows the user to create new instances and apply 28 | `Plausible` to new situations. 29 | 30 | ### What do I do if I'm testing a property about my newly defined type? 31 | 32 | Let us consider a type made for a new formalization: 33 | 34 | ```lean 35 | structure MyType where 36 | x : Nat 37 | y : Nat 38 | h : x ≤ y 39 | deriving Repr 40 | ``` 41 | 42 | How do we test a property about `MyType`? For instance, let us consider 43 | `Testable.check <| ∀ a b : MyType, a.y ≤ b.x → a.x ≤ b.y`. Writing this 44 | property as is will give us an error because we do not have an instance 45 | of `Shrinkable MyType` and `SampleableExt MyType`. We can define one as follows: 46 | 47 | ```lean 48 | instance : Shrinkable MyType where 49 | shrink := fun ⟨x, y, _⟩ => 50 | let proxy := Shrinkable.shrink (x, y - x) 51 | proxy.map (fun (fst, snd) => ⟨fst, fst + snd, by omega⟩) 52 | 53 | instance : SampleableExt MyType := 54 | SampleableExt.mkSelfContained do 55 | let x ← SampleableExt.interpSample Nat 56 | let xyDiff ← SampleableExt.interpSample Nat 57 | return ⟨x, x + xyDiff, by omega⟩ 58 | ``` 59 | 60 | Again, we take advantage of the fact that other types have useful 61 | `Shrinkable` implementations, in this case `Prod`. 62 | 63 | ## Main definitions 64 | 65 | * `Testable` class 66 | * `Testable.check`: a way to test a proposition using random examples 67 | 68 | ## References 69 | 70 | * https://hackage.haskell.org/package/QuickCheck 71 | 72 | -/ 73 | 74 | namespace Plausible 75 | 76 | /-- Result of trying to disprove `p` -/ 77 | inductive TestResult (p : Prop) where 78 | /-- 79 | Succeed when we find another example satisfying `p`. 80 | In `success h`, `h` is an optional proof of the proposition. 81 | Without the proof, all we know is that we found one example 82 | where `p` holds. With a proof, the one test was sufficient to 83 | prove that `p` holds and we do not need to keep finding examples. 84 | -/ 85 | | success : Unit ⊕' p → TestResult p 86 | /-- 87 | Give up when a well-formed example cannot be generated. 88 | `gaveUp n` tells us that `n` invalid examples were tried. 89 | -/ 90 | | gaveUp : Nat → TestResult p 91 | /-- 92 | A counter-example to `p`; the strings specify values for the relevant variables. 93 | `failure h vs n` also carries a proof that `p` does not hold. This way, we can 94 | guarantee that there will be no false positive. The last component, `n`, 95 | is the number of times that the counter-example was shrunk. 96 | -/ 97 | | failure : ¬ p → List String → Nat → TestResult p 98 | deriving Inhabited 99 | 100 | /-- Configuration for testing a property. -/ 101 | structure Configuration where 102 | /-- 103 | How many test instances to generate. 104 | -/ 105 | numInst : Nat := 100 106 | /-- 107 | The maximum size of the values to generate. 108 | -/ 109 | maxSize : Nat := 100 110 | numRetries : Nat := 10 111 | /-- 112 | Enable tracing of values that didn't fulfill preconditions and were thus discarded. 113 | -/ 114 | traceDiscarded : Bool := false 115 | /-- 116 | Enable tracing of values that fulfilled the property and were thus discarded. 117 | -/ 118 | traceSuccesses : Bool := false 119 | /-- 120 | Enable basic tracing of shrinking. 121 | -/ 122 | traceShrink : Bool := false 123 | /-- 124 | Enable tracing of all attempted values during shrinking. 125 | -/ 126 | traceShrinkCandidates : Bool := false 127 | /-- 128 | Hard code the seed to use for the RNG 129 | -/ 130 | randomSeed : Option Nat := none 131 | /-- 132 | Disable output. 133 | -/ 134 | quiet : Bool := false 135 | deriving Inhabited 136 | 137 | open Lean in 138 | instance : ToExpr Configuration where 139 | toTypeExpr := mkConst `Configuration 140 | toExpr cfg := mkApp9 (mkConst ``Configuration.mk) 141 | (toExpr cfg.numInst) (toExpr cfg.maxSize) (toExpr cfg.numRetries) (toExpr cfg.traceDiscarded) 142 | (toExpr cfg.traceSuccesses) (toExpr cfg.traceShrink) (toExpr cfg.traceShrinkCandidates) 143 | (toExpr cfg.randomSeed) (toExpr cfg.quiet) 144 | 145 | /-- 146 | Allow elaboration of `Configuration` arguments to tactics. 147 | -/ 148 | declare_config_elab elabConfig Configuration 149 | 150 | /-- 151 | `PrintableProp p` allows one to print a proposition so that 152 | `Plausible` can indicate how values relate to each other. 153 | It's basically a poor man's delaborator. 154 | -/ 155 | class PrintableProp (p : Prop) where 156 | printProp : String 157 | 158 | export PrintableProp (printProp) 159 | 160 | variable {p q : Prop} 161 | 162 | instance (priority := low) : PrintableProp p where 163 | printProp := "⋯" 164 | 165 | /-- `Testable p` uses random examples to try to disprove `p`. -/ 166 | class Testable (p : Prop) where 167 | run (cfg : Configuration) (minimize : Bool) : Gen (TestResult p) 168 | 169 | def NamedBinder (_n : String) (p : Prop) : Prop := p 170 | 171 | namespace TestResult 172 | 173 | def toString : TestResult p → String 174 | | success (PSum.inl _) => "success (no proof)" 175 | | success (PSum.inr _) => "success (proof)" 176 | | gaveUp n => s!"gave {n} times" 177 | | failure _ counters _ => s!"failed {counters}" 178 | 179 | instance : ToString (TestResult p) := ⟨toString⟩ 180 | 181 | /-- Applicative combinator proof carrying test results. -/ 182 | def combine {p q : Prop} : Unit ⊕' (p → q) → Unit ⊕' p → Unit ⊕' q 183 | | PSum.inr f, PSum.inr proof => PSum.inr <| f proof 184 | | _, _ => PSum.inl () 185 | 186 | /-- Combine the test result for properties `p` and `q` to create a test for their conjunction. -/ 187 | def and : TestResult p → TestResult q → TestResult (p ∧ q) 188 | | failure h xs n, _ => failure (fun h2 => h h2.left) xs n 189 | | _, failure h xs n => failure (fun h2 => h h2.right) xs n 190 | | success h1, success h2 => success <| combine (combine (PSum.inr And.intro) h1) h2 191 | | gaveUp n, gaveUp m => gaveUp <| n + m 192 | | gaveUp n, _ => gaveUp n 193 | | _, gaveUp n => gaveUp n 194 | 195 | /-- Combine the test result for properties `p` and `q` to create a test for their disjunction. -/ 196 | def or : TestResult p → TestResult q → TestResult (p ∨ q) 197 | | failure h1 xs n, failure h2 ys m => 198 | let h3 := fun h => 199 | match h with 200 | | Or.inl h3 => h1 h3 201 | | Or.inr h3 => h2 h3 202 | failure h3 (xs ++ ys) (n + m) 203 | | success h, _ => success <| combine (PSum.inr Or.inl) h 204 | | _, success h => success <| combine (PSum.inr Or.inr) h 205 | | gaveUp n, gaveUp m => gaveUp <| n + m 206 | | gaveUp n, _ => gaveUp n 207 | | _, gaveUp n => gaveUp n 208 | 209 | /-- If `q → p`, then `¬ p → ¬ q` which means that testing `p` can allow us 210 | to find counter-examples to `q`. -/ 211 | def imp (h : q → p) (r : TestResult p) 212 | (p : Unit ⊕' (p → q) := PSum.inl ()) : TestResult q := 213 | match r with 214 | | failure h2 xs n => failure (mt h h2) xs n 215 | | success h2 => success <| combine p h2 216 | | gaveUp n => gaveUp n 217 | 218 | /-- Test `q` by testing `p` and proving the equivalence between the two. -/ 219 | def iff (h : q ↔ p) (r : TestResult p) : TestResult q := 220 | imp h.mp r (PSum.inr h.mpr) 221 | 222 | /-- When we assign a value to a universally quantified variable, 223 | we record that value using this function so that our counter-examples 224 | can be informative. -/ 225 | def addInfo (x : String) (h : q → p) (r : TestResult p) 226 | (p : Unit ⊕' (p → q) := PSum.inl ()) : TestResult q := 227 | if let failure h2 xs n := r then 228 | failure (mt h h2) (x :: xs) n 229 | else 230 | imp h r p 231 | 232 | /-- Add some formatting to the information recorded by `addInfo`. -/ 233 | def addVarInfo {γ : Type _} [Repr γ] (var : String) (x : γ) (h : q → p) (r : TestResult p) 234 | (p : Unit ⊕' (p → q) := PSum.inl ()) : TestResult q := 235 | addInfo s!"{var} := {repr x}" h r p 236 | 237 | def isFailure : TestResult p → Bool 238 | | failure _ _ _ => true 239 | | _ => false 240 | 241 | end TestResult 242 | 243 | namespace Configuration 244 | 245 | /-- A configuration with all the trace options enabled, useful for debugging. -/ 246 | def verbose : Configuration where 247 | traceDiscarded := true 248 | traceSuccesses := true 249 | traceShrink := true 250 | traceShrinkCandidates := true 251 | 252 | end Configuration 253 | 254 | namespace Testable 255 | 256 | open TestResult 257 | 258 | def runProp (p : Prop) [Testable p] : Configuration → Bool → Gen (TestResult p) := Testable.run 259 | 260 | /-- A `dbgTrace` with special formatting -/ 261 | def slimTrace {m : Type → Type _} [Pure m] (s : String) : m PUnit := 262 | dbgTrace s!"[Plausible: {s}]" (fun _ => pure ()) 263 | 264 | instance andTestable [Testable p] [Testable q] : Testable (p ∧ q) where 265 | run := fun cfg min => do 266 | let xp ← runProp p cfg min 267 | let xq ← runProp q cfg min 268 | return and xp xq 269 | 270 | instance orTestable [Testable p] [Testable q] : Testable (p ∨ q) where 271 | run := fun cfg min => do 272 | let xp ← runProp p cfg min 273 | -- As a little performance optimization we can just not run the second 274 | -- test if the first succeeds 275 | match xp with 276 | | success (PSum.inl h) => return success (PSum.inl h) 277 | | success (PSum.inr h) => return success (PSum.inr <| Or.inl h) 278 | | _ => 279 | let xq ← runProp q cfg min 280 | return or xp xq 281 | 282 | instance iffTestable [Testable ((p ∧ q) ∨ (¬ p ∧ ¬ q))] : Testable (p ↔ q) where 283 | run := fun cfg min => do 284 | let h ← runProp ((p ∧ q) ∨ (¬ p ∧ ¬ q)) cfg min 285 | have := by 286 | constructor 287 | · intro h 288 | simp [h, Classical.em] 289 | · intro h 290 | rcases h with ⟨hleft, hright⟩ | ⟨hleft, hright⟩ <;> simp [hleft, hright] 291 | return iff this h 292 | 293 | variable {var : String} 294 | 295 | instance decGuardTestable [PrintableProp p] [Decidable p] {β : p → Prop} [∀ h, Testable (β h)] : 296 | Testable (NamedBinder var <| ∀ h, β h) where 297 | run := fun cfg min => do 298 | if h : p then 299 | let res := runProp (β h) cfg min 300 | let s := printProp p 301 | (fun r => addInfo s!"guard: {s}" (· <| h) r (PSum.inr <| fun q _ => q)) <$> res 302 | else if cfg.traceDiscarded || cfg.traceSuccesses then 303 | let res := fun _ => return gaveUp 1 304 | let s := printProp p 305 | slimTrace s!"discard: Guard {s} does not hold"; res 306 | else 307 | return gaveUp 1 308 | 309 | instance forallTypesTestable {f : Type → Prop} [Testable (f Int)] : 310 | Testable (NamedBinder var <| ∀ x, f x) where 311 | run := fun cfg min => do 312 | let r ← runProp (f Int) cfg min 313 | return addVarInfo var "Int" (· <| Int) r 314 | 315 | -- TODO: only in mathlib: @[pp_with_univ] 316 | instance (priority := 100) forallTypesULiftTestable.{u} 317 | {f : Type u → Prop} [Testable (f (ULift.{u} Int))] : 318 | Testable (NamedBinder var <| ∀ x, f x) where 319 | run cfg min := do 320 | let r ← runProp (f (ULift Int)) cfg min 321 | pure <| addVarInfo var "ULift Int" (· <| ULift Int) r 322 | 323 | /-- 324 | Format the counter-examples found in a test failure. 325 | -/ 326 | def formatFailure (s : String) (xs : List String) (n : Nat) : String := 327 | let counter := String.intercalate "\n" xs 328 | let parts := [ 329 | "\n===================", 330 | s, 331 | counter, 332 | s!"({n} shrinks)", 333 | "-------------------" 334 | ] 335 | String.intercalate "\n" parts 336 | 337 | /-- 338 | Increase the number of shrinking steps in a test result. 339 | -/ 340 | def addShrinks (n : Nat) : TestResult p → TestResult p 341 | | TestResult.failure p xs m => TestResult.failure p xs (m + n) 342 | | p => p 343 | 344 | universe u in 345 | instance {α : Type u} {m : Type u → Type _} [Pure m] : Inhabited (OptionT m α) := 346 | ⟨(pure none : m (Option α))⟩ 347 | 348 | variable {α : Sort _} 349 | 350 | /-- Shrink a counter-example `x` by using `Shrinkable.shrink x`, picking the first 351 | candidate that falsifies a property and recursively shrinking that one. 352 | The process is guaranteed to terminate because `shrink x` produces 353 | a proof that all the values it produces are smaller (according to `SizeOf`) 354 | than `x`. -/ 355 | partial def minimizeAux [SampleableExt α] {β : α → Prop} [∀ x, Testable (β x)] (cfg : Configuration) 356 | (var : String) (x : SampleableExt.proxy α) (n : Nat) : 357 | OptionT Gen (Σ x, TestResult (β (SampleableExt.interp x))) := do 358 | let candidates := SampleableExt.shrink.shrink x 359 | if cfg.traceShrinkCandidates then 360 | slimTrace s!"Candidates for {var} := {repr x}:\n {repr candidates}" 361 | for candidate in candidates do 362 | if cfg.traceShrinkCandidates then 363 | slimTrace s!"Trying {var} := {repr candidate}" 364 | let res ← OptionT.lift <| Testable.runProp (β (SampleableExt.interp candidate)) cfg true 365 | if res.isFailure then 366 | if cfg.traceShrink then 367 | slimTrace s!"{var} shrunk to {repr candidate} from {repr x}" 368 | let currentStep := OptionT.lift <| return Sigma.mk candidate (addShrinks (n + 1) res) 369 | let nextStep := minimizeAux cfg var candidate (n + 1) 370 | return ← (nextStep <|> currentStep) 371 | if cfg.traceShrink then 372 | slimTrace s!"No shrinking possible for {var} := {repr x}" 373 | failure 374 | 375 | /-- Once a property fails to hold on an example, look for smaller counter-examples 376 | to show the user. -/ 377 | def minimize [SampleableExt α] {β : α → Prop} [∀ x, Testable (β x)] (cfg : Configuration) 378 | (var : String) (x : SampleableExt.proxy α) (r : TestResult (β <| SampleableExt.interp x)) : 379 | Gen (Σ x, TestResult (β <| SampleableExt.interp x)) := do 380 | if cfg.traceShrink then 381 | slimTrace "Shrink" 382 | slimTrace s!"Attempting to shrink {var} := {repr x}" 383 | let res ← OptionT.run <| minimizeAux cfg var x 0 384 | return res.getD ⟨x, r⟩ 385 | 386 | /-- Test a universal property by creating a sample of the right type and instantiating the 387 | bound variable with it. -/ 388 | instance varTestable [SampleableExt α] {β : α → Prop} [∀ x, Testable (β x)] : 389 | Testable (NamedBinder var <| ∀ x : α, β x) where 390 | run := fun cfg min => do 391 | let x ← SampleableExt.sample 392 | if cfg.traceSuccesses || cfg.traceDiscarded then 393 | slimTrace s!"{var} := {repr x}" 394 | let r ← Testable.runProp (β <| SampleableExt.interp x) cfg false 395 | let ⟨finalX, finalR⟩ ← 396 | if isFailure r then 397 | if cfg.traceSuccesses then 398 | slimTrace s!"{var} := {repr x} is a failure" 399 | if min then 400 | minimize cfg var x r 401 | else 402 | pure ⟨x, r⟩ 403 | else 404 | pure ⟨x, r⟩ 405 | return addVarInfo var finalX (· <| SampleableExt.interp finalX) finalR 406 | 407 | /-- Test a universal property about propositions -/ 408 | instance propVarTestable {β : Prop → Prop} [∀ b : Bool, Testable (β b)] : 409 | Testable (NamedBinder var <| ∀ p : Prop, β p) 410 | where 411 | run := fun cfg min => 412 | imp (fun h (b : Bool) => h b) <$> Testable.runProp (NamedBinder var <| ∀ b : Bool, β b) cfg min 413 | 414 | instance (priority := high) unusedVarTestable {β : Prop} [Nonempty α] [Testable β] : 415 | Testable (NamedBinder var (α → β)) 416 | where 417 | run := fun cfg min => do 418 | if cfg.traceDiscarded || cfg.traceSuccesses then 419 | slimTrace s!"{var} is unused" 420 | let r ← Testable.runProp β cfg min 421 | let finalR := addInfo s!"{var} is irrelevant (unused)" id r 422 | return imp (· <| Classical.ofNonempty) finalR (PSum.inr <| fun x _ => x) 423 | 424 | instance (priority := 2000) subtypeVarTestable {p : α → Prop} {β : α → Prop} 425 | [∀ x, PrintableProp (p x)] 426 | [∀ x, Testable (β x)] 427 | [SampleableExt (Subtype p)] {var'} : 428 | Testable (NamedBinder var <| (x : α) → NamedBinder var' <| p x → β x) where 429 | run cfg min := 430 | letI (x : Subtype p) : Testable (β x) := 431 | { run := fun cfg min => do 432 | let r ← Testable.runProp (β x.val) cfg min 433 | return addInfo s!"guard: {printProp (p x)} (by construction)" id r (PSum.inr id) } 434 | do 435 | let r ← @Testable.run (∀ x : Subtype p, β x.val) (@varTestable var _ _ _ _) cfg min 436 | have := by simp [Subtype.forall, NamedBinder] 437 | return iff this r 438 | 439 | instance (priority := low) decidableTestable {p : Prop} [PrintableProp p] [Decidable p] : 440 | Testable p where 441 | run := fun _ _ => 442 | if h : p then 443 | return success (PSum.inr h) 444 | else 445 | let s := printProp p 446 | return failure h [s!"issue: {s} does not hold"] 0 447 | 448 | end Testable 449 | 450 | section PrintableProp 451 | 452 | variable {α : Type _} 453 | 454 | instance Eq.printableProp [Repr α] {x y : α} : PrintableProp (x = y) where 455 | printProp := s!"{repr x} = {repr y}" 456 | 457 | instance Ne.printableProp [Repr α] {x y : α} : PrintableProp (x ≠ y) where 458 | printProp := s!"{repr x} ≠ {repr y}" 459 | 460 | instance LE.printableProp [Repr α] [LE α] {x y : α} : PrintableProp (x ≤ y) where 461 | printProp := s!"{repr x} ≤ {repr y}" 462 | 463 | instance LT.printableProp [Repr α] [LT α] {x y : α} : PrintableProp (x < y) where 464 | printProp := s!"{repr x} < {repr y}" 465 | 466 | variable {x y : Prop} 467 | 468 | instance And.printableProp [PrintableProp x] [PrintableProp y] : PrintableProp (x ∧ y) where 469 | printProp := s!"{printProp x} ∧ {printProp y}" 470 | 471 | instance Or.printableProp [PrintableProp x] [PrintableProp y] : PrintableProp (x ∨ y) where 472 | printProp := s!"{printProp x} ∨ {printProp y}" 473 | 474 | instance Iff.printableProp [PrintableProp x] [PrintableProp y] : PrintableProp (x ↔ y) where 475 | printProp := s!"{printProp x} ↔ {printProp y}" 476 | 477 | instance Imp.printableProp [PrintableProp x] [PrintableProp y] : PrintableProp (x → y) where 478 | printProp := s!"{printProp x} → {printProp y}" 479 | 480 | instance Not.printableProp [PrintableProp x] : PrintableProp (¬x) where 481 | printProp := s!"¬{printProp x}" 482 | 483 | instance True.printableProp : PrintableProp True where 484 | printProp := "True" 485 | 486 | instance False.printableProp : PrintableProp False where 487 | printProp := "False" 488 | 489 | instance Bool.printableProp {b : Bool} : PrintableProp b where 490 | printProp := if b then "true" else "false" 491 | 492 | end PrintableProp 493 | 494 | section IO 495 | open TestResult 496 | 497 | /-- Execute `cmd` and repeat every time the result is `gaveUp` (at most `n` times). -/ 498 | def retry (cmd : Rand (TestResult p)) : Nat → Rand (TestResult p) 499 | | 0 => return TestResult.gaveUp 1 500 | | n+1 => do 501 | let r ← cmd 502 | match r with 503 | | .success hp => return success hp 504 | | .failure h xs n => return failure h xs n 505 | | .gaveUp _ => retry cmd n 506 | 507 | /-- Count the number of times the test procedure gave up. -/ 508 | def giveUp (x : Nat) : TestResult p → TestResult p 509 | | success (PSum.inl ()) => gaveUp x 510 | | success (PSum.inr p) => success <| (PSum.inr p) 511 | | gaveUp n => gaveUp <| n + x 512 | | TestResult.failure h xs n => failure h xs n 513 | 514 | /-- Try `n` times to find a counter-example for `p`. -/ 515 | def Testable.runSuiteAux (p : Prop) [Testable p] (cfg : Configuration) : 516 | TestResult p → Nat → Rand (TestResult p) 517 | | r, 0 => return r 518 | | r, n+1 => do 519 | let size := (cfg.numInst - n - 1) * cfg.maxSize / cfg.numInst 520 | if cfg.traceSuccesses then 521 | slimTrace s!"New sample" 522 | slimTrace s!"Retrying up to {cfg.numRetries} times until guards hold" 523 | let x ← retry (ReaderT.run (Testable.runProp p cfg true) ⟨size⟩) cfg.numRetries 524 | match x with 525 | | success (PSum.inl ()) => runSuiteAux p cfg r n 526 | | gaveUp g => runSuiteAux p cfg (giveUp g r) n 527 | | _ => return x 528 | 529 | /-- Try to find a counter-example of `p`. -/ 530 | def Testable.runSuite (p : Prop) [Testable p] (cfg : Configuration := {}) : Rand (TestResult p) := 531 | Testable.runSuiteAux p cfg (success <| PSum.inl ()) cfg.numInst 532 | 533 | /-- Run a test suite for `p` in `BaseIO` using the global RNG in `stdGenRef`. -/ 534 | def Testable.checkIO (p : Prop) [Testable p] (cfg : Configuration := {}) : BaseIO (TestResult p) := 535 | letI : MonadLift Id BaseIO := ⟨fun f => return Id.run f⟩ 536 | match cfg.randomSeed with 537 | | none => runRand (Testable.runSuite p cfg) 538 | | some seed => runRandWith seed (Testable.runSuite p cfg) 539 | 540 | end IO 541 | 542 | namespace Decorations 543 | 544 | open Lean 545 | 546 | /-- Traverse the syntax of a proposition to find universal quantifiers 547 | quantifiers and add `NamedBinder` annotations next to them. -/ 548 | partial def addDecorations (e : Expr) : MetaM Expr := 549 | Meta.transform e fun expr => do 550 | if not (← Meta.inferType expr).isProp then 551 | return .done expr 552 | else if let Expr.forallE name type body data := expr then 553 | let newType ← addDecorations type 554 | let newBody ← Meta.withLocalDecl name data type fun fvar => do 555 | return (← addDecorations (body.instantiate1 fvar)).abstract #[fvar] 556 | let rest := Expr.forallE name newType newBody data 557 | return .done <| (← Meta.mkAppM `Plausible.NamedBinder #[mkStrLit name.toString, rest]) 558 | else 559 | return .continue 560 | 561 | /-- `DecorationsOf p` is used as a hint to `mk_decorations` to specify 562 | that the goal should be satisfied with a proposition equivalent to `p` 563 | with added annotations. -/ 564 | abbrev DecorationsOf (_p : Prop) := Prop 565 | 566 | open Elab.Tactic 567 | open Meta 568 | 569 | /-- In a goal of the shape `⊢ DecorationsOf p`, `mk_decoration` examines 570 | the syntax of `p` and adds `NamedBinder` around universal quantifications 571 | to improve error messages. This tool can be used in the declaration of a 572 | function as follows: 573 | ```lean 574 | def foo (p : Prop) (p' : Decorations.DecorationsOf p := by mk_decorations) [Testable p'] : ... 575 | ``` 576 | `p` is the parameter given by the user, `p'` is a definitionally equivalent 577 | proposition where the quantifiers are annotated with `NamedBinder`. 578 | -/ 579 | scoped elab "mk_decorations" : tactic => do 580 | let goal ← getMainGoal 581 | let goalType ← goal.getType 582 | if let .app (.const ``Decorations.DecorationsOf _) body := goalType then 583 | closeMainGoal `mk_decorations (← addDecorations body) 584 | 585 | end Decorations 586 | 587 | open Decorations in 588 | /-- Run a test suite for `p` and throw an exception if `p` does not hold. -/ 589 | def Testable.check (p : Prop) (cfg : Configuration := {}) 590 | (p' : Decorations.DecorationsOf p := by mk_decorations) [Testable p'] : Lean.CoreM PUnit := do 591 | match ← Testable.checkIO p' cfg with 592 | | TestResult.success _ => if !cfg.quiet then Lean.logInfo "Unable to find a counter-example" 593 | | TestResult.gaveUp n => 594 | if !cfg.quiet then 595 | let msg := s!"Gave up after failing to generate values that fulfill the preconditions {n} times." 596 | Lean.logWarning msg 597 | | TestResult.failure _ xs n => 598 | let msg := "Found a counter-example!" 599 | if cfg.quiet then 600 | Lean.throwError msg 601 | else 602 | Lean.throwError <| formatFailure msg xs n 603 | 604 | -- #eval Testable.check (∀ (x y z a : Nat) (h1 : 3 < x) (h2 : 3 < y), x - y = y - x) 605 | -- Configuration.verbose 606 | -- #eval Testable.check (∀ x : Nat, ∀ y : Nat, x + y = y + x) Configuration.verbose 607 | -- #eval Testable.check (∀ (x : (Nat × Nat)), x.fst - x.snd - 10 = x.snd - x.fst - 10) 608 | -- Configuration.verbose 609 | -- #eval Testable.check (∀ (x : Nat) (h : 10 < x), 5 < x) Configuration.verbose 610 | 611 | macro tk:"#test " e:term : command => `(command| #eval%$tk Testable.check $e) 612 | 613 | -- #test ∀ (x : Nat) (h : 5 < x), 10 < x 614 | -- #test ∀ (x : Nat) (h : 10 < x), 5 < x 615 | 616 | end Plausible 617 | -------------------------------------------------------------------------------- /README.md: -------------------------------------------------------------------------------- 1 | # Plausible 2 | A property testing framework for Lean 4 that integrates into the tactic framework. 3 | 4 | ## Usage 5 | If you are using built in types Plausible is usually able to handle them already: 6 | ```lean 7 | import Plausible 8 | 9 | example (xs ys : Array Nat) : xs.size = ys.size → xs = ys := by 10 | /-- 11 | =================== 12 | Found a counter-example! 13 | xs := #[0] 14 | ys := #[1] 15 | guard: 1 = 1 16 | issue: #[0] = #[1] does not hold 17 | (0 shrinks) 18 | ------------------- 19 | -/ 20 | plausible 21 | 22 | #eval Plausible.Testable.check <| ∀ (xs ys : Array Nat), xs.size = ys.size → xs = ys 23 | ``` 24 | 25 | If you are defining your own type it needs instances of `Repr`, `Plausible.Shrinkable` and 26 | `Plausible.SampleableExt`: 27 | ```lean 28 | import Plausible 29 | 30 | open Plausible 31 | 32 | structure MyType where 33 | x : Nat 34 | y : Nat 35 | h : x ≤ y 36 | deriving Repr 37 | 38 | instance : Shrinkable MyType where 39 | shrink := fun ⟨x, y, _⟩ => 40 | let proxy := Shrinkable.shrink (x, y - x) 41 | proxy.map (fun (fst, snd) => ⟨fst, fst + snd, by omega⟩) 42 | 43 | instance : SampleableExt MyType := 44 | SampleableExt.mkSelfContained do 45 | let x ← SampleableExt.interpSample Nat 46 | let xyDiff ← SampleableExt.interpSample Nat 47 | return ⟨x, x + xyDiff, by omega⟩ 48 | 49 | -- No counter example found 50 | #eval Testable.check <| ∀ a b : MyType, a.y ≤ b.x → a.x ≤ b.y 51 | ``` 52 | For more documentation refer to the module docs. 53 | -------------------------------------------------------------------------------- /Test.lean: -------------------------------------------------------------------------------- 1 | /- 2 | Copyright (c) 2024 Lean FRO, LLC. All rights reserved. 3 | Released under Apache 2.0 license as described in the file LICENSE. 4 | Authors: Henrik Böving 5 | -/ 6 | import Test.Tactic 7 | import Test.Testable 8 | -------------------------------------------------------------------------------- /Test/Tactic.lean: -------------------------------------------------------------------------------- 1 | /- 2 | Copyright (c) 2024 Lean FRO, LLC. All rights reserved. 3 | Released under Apache 2.0 license as described in the file LICENSE. 4 | Authors: Henrik Böving 5 | -/ 6 | import Plausible 7 | 8 | /-! 9 | Demonstrate that Plausible can handle the basic types from core: 10 | - Sum 11 | - Sigma 12 | - Unit 13 | - Prod 14 | - Bool 15 | - Nat 16 | - Fin 17 | - UIntX 18 | - BitVec 19 | - Char 20 | - Option 21 | - List 22 | - String 23 | - Array 24 | -/ 25 | 26 | 27 | /-- error: Found a counter-example! -/ 28 | #guard_msgs in 29 | example (a b : Sum Nat Nat) : a = b := by 30 | plausible (config := {quiet := true}) 31 | 32 | /-- error: Found a counter-example! -/ 33 | #guard_msgs in 34 | example (a b : Σ n : Nat, Nat) : a.fst = b.snd := by 35 | plausible (config := {quiet := true}) 36 | 37 | /-- error: Found a counter-example! -/ 38 | #guard_msgs in 39 | example (a b : Unit) : a ≠ b := by 40 | plausible (config := {quiet := true}) 41 | 42 | /-- error: Found a counter-example! -/ 43 | #guard_msgs in 44 | example (x y : Nat × Unit) : x = y := by 45 | plausible (config := {quiet := true}) 46 | 47 | /-- error: Found a counter-example! -/ 48 | #guard_msgs in 49 | example (a b : Bool) : a = b := by 50 | plausible (config := {quiet := true}) 51 | 52 | /-- error: Found a counter-example! -/ 53 | #guard_msgs in 54 | example (a b c : Nat) : a + (b - c) = (a + b) - c := by 55 | plausible (config := {quiet := true}) 56 | 57 | /-- error: Found a counter-example! -/ 58 | #guard_msgs in 59 | example (a : Fin (n + 1)) : a + 1 > a := by 60 | plausible (config := {quiet := true}) 61 | 62 | /-- error: Found a counter-example! -/ 63 | #guard_msgs in 64 | example (a : BitVec n) : a + 1 > a := by 65 | plausible (config := {quiet := true}) 66 | 67 | /-- error: Found a counter-example! -/ 68 | #guard_msgs in 69 | example (a : UInt8) : a - 1 < a := by 70 | plausible (config := {quiet := true}) 71 | 72 | /-- error: Found a counter-example! -/ 73 | #guard_msgs in 74 | example (a : UInt16) : a - 1 < a := by 75 | plausible (config := {quiet := true}) 76 | 77 | /-- error: Found a counter-example! -/ 78 | #guard_msgs in 79 | example (a : UInt32) : a - 1 < a := by 80 | plausible (config := {quiet := true}) 81 | 82 | /-- error: Found a counter-example! -/ 83 | #guard_msgs in 84 | example (a : UInt64) : a - 1 < a := by 85 | plausible (config := {quiet := true}) 86 | 87 | /-- error: Found a counter-example! -/ 88 | #guard_msgs in 89 | example (a : USize) : a - 1 < a := by 90 | plausible (config := {quiet := true}) 91 | 92 | /-- error: Found a counter-example! -/ 93 | #guard_msgs in 94 | example (a : Char) : a ≠ a := by 95 | plausible (config := {quiet := true}) 96 | 97 | /-- error: Found a counter-example! -/ 98 | #guard_msgs in 99 | example (a : Option Char) : a ≠ a := by 100 | plausible (config := {quiet := true}) 101 | 102 | /-- error: Found a counter-example! -/ 103 | #guard_msgs in 104 | example (xs ys : List Nat) : xs.length = ys.length → xs = ys := by 105 | plausible (config := {quiet := true}) 106 | 107 | /-- error: Found a counter-example! -/ 108 | #guard_msgs in 109 | example (xs ys : String) : xs.length = ys.length → xs = ys := by 110 | plausible (config := {quiet := true}) 111 | 112 | /-- error: Found a counter-example! -/ 113 | #guard_msgs in 114 | example (xs ys : Array Nat) : xs.size = ys.size → xs = ys := by 115 | plausible (config := {quiet := true}) 116 | 117 | /-- error: Found a counter-example! -/ 118 | #guard_msgs in 119 | example (xs : List Int) (f : Int → Int) : xs.map f = xs := by 120 | plausible (config := {quiet := true}) 121 | 122 | /-- 123 | info: Unable to find a counter-example 124 | --- 125 | warning: declaration uses 'sorry' 126 | -/ 127 | #guard_msgs in 128 | example (a : Sum Nat Nat) : a = a := by 129 | plausible 130 | 131 | /-- 132 | warning: Gave up after failing to generate values that fulfill the preconditions 100 times. 133 | --- 134 | warning: declaration uses 'sorry' 135 | -/ 136 | #guard_msgs in 137 | example (a b : Sum Nat Nat) : a ≠ a → a = b := by 138 | plausible (config := {numInst := 100}) 139 | 140 | -- https://leanprover.zulipchat.com/#narrow/stream/287929-mathlib4/topic/slim_check.20giving.20wrong.20counterexamples.3F/near/420008365 141 | open Nat in 142 | /-- 143 | info: Unable to find a counter-example 144 | --- 145 | warning: declaration uses 'sorry' 146 | -/ 147 | #guard_msgs in 148 | theorem testBit_pred : 149 | testBit (pred x) i = (decide (0 < x) && 150 | (Bool.xor ((List.range i).all fun j => ! testBit x j) (testBit x i))) := by 151 | plausible 152 | 153 | /-- error: Found a counter-example! -/ 154 | #guard_msgs in 155 | theorem ulift_nat (f : ULift.{1} Nat) : f = ⟨0⟩ := by 156 | plausible (config := {quiet := true}) 157 | 158 | /-- error: Found a counter-example! -/ 159 | #guard_msgs in 160 | theorem type_u (α : Type u) (l : List α) : l = l ++ l := by 161 | plausible (config := {quiet := true}) 162 | -------------------------------------------------------------------------------- /Test/Testable.lean: -------------------------------------------------------------------------------- 1 | /- 2 | Copyright (c) 2024 Lean FRO, LLC. All rights reserved. 3 | Released under Apache 2.0 license as described in the file LICENSE. 4 | Authors: Henrik Böving 5 | -/ 6 | import Plausible.Testable 7 | 8 | open Plausible 9 | 10 | structure MyType where 11 | x : Nat 12 | y : Nat 13 | h : x ≤ y 14 | deriving Repr 15 | 16 | instance : Shrinkable MyType where 17 | shrink := fun ⟨x, y, _⟩ => 18 | let proxy := Shrinkable.shrink (x, y - x) 19 | proxy.map (fun (fst, snd) => ⟨fst, fst + snd, by omega⟩) 20 | 21 | instance : SampleableExt MyType := 22 | SampleableExt.mkSelfContained do 23 | let x ← SampleableExt.interpSample Nat 24 | let xyDiff ← SampleableExt.interpSample Nat 25 | return ⟨x, x + xyDiff, by omega⟩ 26 | 27 | -- TODO: this is a noisy test. 28 | -- We can't use `#guard_msgs` because the number of attempts to non-deterministic. 29 | #eval Testable.check <| ∀ a b : MyType, a.y ≤ b.x → a.x ≤ b.y 30 | -------------------------------------------------------------------------------- /lake-manifest.json: -------------------------------------------------------------------------------- 1 | {"version": "1.1.0", 2 | "packagesDir": ".lake/packages", 3 | "packages": [], 4 | "name": "plausible", 5 | "lakeDir": ".lake"} 6 | -------------------------------------------------------------------------------- /lakefile.toml: -------------------------------------------------------------------------------- 1 | name = "plausible" 2 | version = "0.1.0" 3 | defaultTargets = ["Plausible"] 4 | testRunner = "Test" 5 | 6 | [[lean_lib]] 7 | name = "Test" 8 | 9 | [[lean_lib]] 10 | name = "Plausible" 11 | -------------------------------------------------------------------------------- /lean-toolchain: -------------------------------------------------------------------------------- 1 | leanprover/lean4:v4.21.0-rc3 2 | --------------------------------------------------------------------------------