├── .gitignore ├── lean-toolchain ├── smtcomp └── 2025 │ ├── .gitignore │ ├── makefile │ ├── refs.bib │ └── main.typ ├── Leanwuzla ├── Sexp.lean ├── Aux.lean ├── Basic.lean ├── NoKernel.lean ├── Sexp │ ├── Dsl.lean │ └── Basic.lean └── Parser.lean ├── Test.lean ├── lakefile.toml ├── lake-manifest.json ├── .github └── workflows │ └── lean_action_ci.yml ├── Test ├── Parser.lean ├── BVCompare.lean ├── Sexp.lean └── Bitwuzla.lean ├── README.md ├── Main.lean ├── LICENSE └── Leanwuzla.lean /.gitignore: -------------------------------------------------------------------------------- 1 | /.lake 2 | -------------------------------------------------------------------------------- /lean-toolchain: -------------------------------------------------------------------------------- 1 | nightly-2025-06-27 2 | -------------------------------------------------------------------------------- /smtcomp/2025/.gitignore: -------------------------------------------------------------------------------- 1 | main.pdf 2 | 3 | -------------------------------------------------------------------------------- /smtcomp/2025/makefile: -------------------------------------------------------------------------------- 1 | main.pdf: main.typ 2 | typst compile main.typ 3 | -------------------------------------------------------------------------------- /Leanwuzla/Sexp.lean: -------------------------------------------------------------------------------- 1 | import Leanwuzla.Sexp.Basic 2 | import Leanwuzla.Sexp.Dsl 3 | -------------------------------------------------------------------------------- /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.Bitwuzla 7 | import Test.Parser 8 | -------------------------------------------------------------------------------- /lakefile.toml: -------------------------------------------------------------------------------- 1 | name = "leanwuzla" 2 | version = "0.1.0" 3 | defaultTargets = ["Leanwuzla", "leanwuzla"] 4 | testRunner = "Test" 5 | 6 | [[lean_lib]] 7 | name = "Test" 8 | 9 | [[lean_lib]] 10 | name = "Leanwuzla" 11 | 12 | [[lean_exe]] 13 | name = "leanwuzla" 14 | root = "Main" 15 | supportInterpreter = true 16 | 17 | [[require]] 18 | scope = "leanprover" 19 | name = "Cli" 20 | rev = "main" 21 | -------------------------------------------------------------------------------- /lake-manifest.json: -------------------------------------------------------------------------------- 1 | {"version": "1.1.0", 2 | "packagesDir": ".lake/packages", 3 | "packages": 4 | [{"url": "https://github.com/leanprover/lean4-cli", 5 | "type": "git", 6 | "subDir": null, 7 | "scope": "leanprover", 8 | "rev": "1604206fcd0462da9a241beeac0e2df471647435", 9 | "name": "Cli", 10 | "manifestFile": "lake-manifest.json", 11 | "inputRev": "main", 12 | "inherited": false, 13 | "configFile": "lakefile.toml"}], 14 | "name": "leanwuzla", 15 | "lakeDir": ".lake"} 16 | -------------------------------------------------------------------------------- /.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: DeterminateSystems/nix-installer-action@main 14 | 15 | - name: install bitwuzla 16 | run: nix profile install nixpkgs#bitwuzla 17 | 18 | - uses: actions/checkout@v4 19 | 20 | - uses: leanprover/lean-action@v1 21 | with: 22 | test: true 23 | -------------------------------------------------------------------------------- /Test/Parser.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: Abdalrhman Mohamed 5 | -/ 6 | import Lean 7 | import Leanwuzla.Parser 8 | 9 | open Lean Elab Command 10 | 11 | def parseSexps (cmds : List Sexp) : MetaM Expr := do 12 | let query := Parser.filterCmds cmds 13 | ofExcept <| (Parser.parseQuery query).run' {} 14 | 15 | def elabSexps (cmds : List Sexp) : CommandElabM Unit := do 16 | runTermElabM fun _ => do 17 | let e ← parseSexps cmds 18 | logInfo m!"{e}" 19 | 20 | def query1 : List Sexp := 21 | sexps!{ 22 | (define-sort i32 () (_ BitVec 32)) 23 | (define-sort i64 () (_ BitVec 64)) 24 | (declare-fun f (i32 i64) Bool) 25 | } 26 | 27 | /-- 28 | info: let i32 := BitVec 32; 29 | let i64 := BitVec 64; 30 | ∀ (f : i32 → i64 → Bool), (!true) = true 31 | -/ 32 | #guard_msgs in 33 | #eval elabSexps query1 34 | -------------------------------------------------------------------------------- /Test/BVCompare.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 Leanwuzla 7 | 8 | variable (a b c : Bool) 9 | variable (x y z : BitVec 32) 10 | 11 | example : x + y = y + x := by 12 | bv_compare 13 | sorry 14 | 15 | example : x - 1 ≠ x := by 16 | bv_compare 17 | sorry 18 | 19 | example : ∀ (x : BitVec 32), x.and 0#32 = 0#32 := by 20 | bv_compare 21 | sorry 22 | 23 | example : ∀ (x x_1 : BitVec 16), BitVec.truncate 8 ((x_1.and 255#16).and x) = BitVec.truncate 8 (x_1.and x) := by 24 | bv_compare 25 | sorry 26 | 27 | example : ∀ (x : BitVec 1), BitVec.ofBool (0#1 > x) = 0#1 := by 28 | bv_compare 29 | sorry 30 | 31 | theorem extracted_1 (x y : BitVec 8) : x + y = x + x := by 32 | bv_compare 33 | sorry 34 | 35 | example (x y : BitVec 64) : x * y = y * x := by 36 | bv_compare (config := { timeout := 1 }) 37 | sorry 38 | -------------------------------------------------------------------------------- /README.md: -------------------------------------------------------------------------------- 1 | # leanwuzla 2 | ## Tactic Frontend 3 | This provides the ability to call Bitwuzla with problems generated by the `bv_decide` tactic 4 | frontend. You can call it like this: 5 | ``` 6 | bv_bitwuzla "/path/to/bitwuzla" 7 | ``` 8 | To point the tool towards the solver and if you want to see what bitwuzla is doing: 9 | ``` 10 | set_option diagnostics true 11 | ``` 12 | Alternatively to compare `bv_decide` and `bv_bitwuzla` directly run: 13 | ``` 14 | bv_compare "/path/to/bitwuzla" 15 | ``` 16 | 17 | ## SMT-LIB v2 Terminal Frontend 18 | This provides the ability to call `bv_decide` with QF_BV SMT-LIB v2 problems. You can call it like 19 | this: 20 | ``` 21 | lake exe leanwuzla file.smt2 22 | ``` 23 | For options see `--help` 24 | 25 | ## Quality and Stability 26 | This tool: 27 | - is a development tool for figuring out differences between what Bitwuzla and `bv_decide` can do. 28 | - is not meant to be used productively 29 | - does not necessarily produce a sound reduction to SMT-lib 30 | -------------------------------------------------------------------------------- /Leanwuzla/Aux.lean: -------------------------------------------------------------------------------- 1 | prelude 2 | import Init.Data.BitVec.Basic 3 | 4 | /-! A list of Auxiliary definitions of SMT-LIB functions. -/ 5 | 6 | @[bv_normalize] def implies (a b : Bool) : Bool := 7 | !a || b 8 | 9 | namespace BitVec 10 | 11 | @[bv_normalize] protected def nand {n : Nat} (x y : BitVec n) : BitVec n := 12 | ~~~(x &&& y) 13 | 14 | @[bv_normalize] protected def nor {n : Nat} (x y : BitVec n) : BitVec n := 15 | ~~~(x ||| y) 16 | 17 | @[bv_normalize] protected def xnor {n : Nat} (x y : BitVec n) : BitVec n := 18 | ~~~(x ^^^ y) 19 | 20 | @[bv_normalize] protected def compare {n : Nat} (x y : BitVec n) : BitVec 1 := 21 | bif x == y then 1#1 else 0#1 22 | 23 | @[bv_normalize] protected def ugt {n : Nat} (x y : BitVec n) : Bool := 24 | BitVec.ult y x 25 | 26 | @[bv_normalize] protected def uge {n : Nat} (x y : BitVec n) : Bool := 27 | BitVec.ule y x 28 | 29 | @[bv_normalize] protected def sgt {n : Nat} (x y : BitVec n) : Bool := 30 | BitVec.slt y x 31 | 32 | @[bv_normalize] protected def sge {n : Nat} (x y : BitVec n) : Bool := 33 | BitVec.sle y x 34 | 35 | end BitVec 36 | -------------------------------------------------------------------------------- /Test/Sexp.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 Leanwuzla 7 | import Leanwuzla.Sexp 8 | 9 | open Std.Tactic.BVDecide 10 | 11 | def expr1 : BVLogicalExpr := .const true 12 | 13 | /-- info: true -/ 14 | #guard_msgs in 15 | #eval Sexp.Parser.manySexps! |>.run (Lean.Elab.Tactic.BVDecide.Frontend.toSMT expr1 {}) |>.isOk 16 | 17 | def expr2 : BVLogicalExpr := .not (.const true) 18 | 19 | /-- info: true -/ 20 | #guard_msgs in 21 | #eval Sexp.Parser.manySexps! |>.run (Lean.Elab.Tactic.BVDecide.Frontend.toSMT expr2 {}) |>.isOk 22 | 23 | def expr3 : BVLogicalExpr := .gate .and (.const true) (.const false) 24 | 25 | /-- info: true -/ 26 | #guard_msgs in 27 | #eval Sexp.Parser.manySexps! |>.run (Lean.Elab.Tactic.BVDecide.Frontend.toSMT expr3 {}) |>.isOk 28 | 29 | def expr4 : BVLogicalExpr := .literal (.bin (.const 10#32) .eq (.const 11#32)) 30 | 31 | /-- info: true -/ 32 | #guard_msgs in 33 | #eval Sexp.Parser.manySexps! |>.run (Lean.Elab.Tactic.BVDecide.Frontend.toSMT expr4 {}) |>.isOk 34 | -------------------------------------------------------------------------------- /Leanwuzla/Basic.lean: -------------------------------------------------------------------------------- 1 | import Lean.Elab.Tactic.BVDecide.Frontend.BVDecide 2 | 3 | open Lean 4 | 5 | private partial def getIntrosSize (e : Expr) : Nat := 6 | go 0 e 7 | where 8 | go (size : Nat) : Expr → Nat 9 | | .forallE _ _ b _ => go (size + 1) b 10 | | .mdata _ b => go size b 11 | | _ => size 12 | 13 | /-- 14 | Introduce only forall binders and preserve names. 15 | -/ 16 | def _root_.Lean.MVarId.introsP (mvarId : MVarId) : MetaM (Array FVarId × MVarId) := do 17 | let type ← mvarId.getType 18 | let type ← instantiateMVars type 19 | let n := getIntrosSize type 20 | if n == 0 then 21 | return (#[], mvarId) 22 | else 23 | mvarId.introNP n 24 | 25 | structure Context where 26 | acNf : Bool 27 | parseOnly : Bool 28 | timeout : Nat 29 | input : String 30 | maxSteps : Nat 31 | disableAndFlatten : Bool 32 | disableEmbeddedConstraintSubst : Bool 33 | disableKernel : Bool 34 | 35 | abbrev SolverM := ReaderT Context MetaM 36 | 37 | namespace SolverM 38 | 39 | def getParseOnly : SolverM Bool := return (← read).parseOnly 40 | def getInput : SolverM String := return (← read).input 41 | def getKernelDisabled : SolverM Bool := return (← read).disableKernel 42 | 43 | def getBVDecideConfig : SolverM Elab.Tactic.BVDecide.Frontend.BVDecideConfig := do 44 | let ctx ← read 45 | return { 46 | timeout := ctx.timeout 47 | acNf := ctx.acNf 48 | embeddedConstraintSubst := !ctx.disableEmbeddedConstraintSubst 49 | andFlattening := !ctx.disableAndFlatten 50 | maxSteps := ctx.maxSteps 51 | structures := false, 52 | fixedInt := false, 53 | enums := false 54 | } 55 | 56 | def run (x : SolverM α) (ctx : Context) (coreContext : Core.Context) (coreState : Core.State) : 57 | IO α := do 58 | let (res, _, _) ← ReaderT.run x ctx |> (Meta.MetaM.toIO · coreContext coreState) 59 | return res 60 | 61 | end SolverM 62 | -------------------------------------------------------------------------------- /Test/Bitwuzla.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 Leanwuzla 7 | 8 | variable (a b c : Bool) 9 | variable (x y z : BitVec 32) 10 | 11 | /-- error: Bitwuzla thinks it's right but can't trust the wuzla! -/ 12 | #guard_msgs in 13 | example : x + y = y + x := by 14 | bv_bitwuzla 15 | 16 | /-- error: Bitwuzla thinks it's right but can't trust the wuzla! -/ 17 | #guard_msgs in 18 | example : x - 1 ≠ x := by 19 | bv_bitwuzla 20 | 21 | /-- error: Bitwuzla thinks it's right but can't trust the wuzla! -/ 22 | #guard_msgs in 23 | example : (-x) + y = y - x:= by 24 | bv_bitwuzla 25 | 26 | /-- error: Bitwuzla thinks it's right but can't trust the wuzla! -/ 27 | #guard_msgs in 28 | example : x * y = y * x := by 29 | bv_bitwuzla 30 | 31 | /-- error: Bitwuzla thinks it's right but can't trust the wuzla! -/ 32 | #guard_msgs in 33 | example : x / y ≤ x := by 34 | bv_bitwuzla 35 | 36 | /-- error: Bitwuzla thinks it's right but can't trust the wuzla! -/ 37 | #guard_msgs in 38 | example (hx : BitVec.slt 0 x) (hy : BitVec.slt 0 y) : x.msb = y.msb := by 39 | bv_bitwuzla 40 | 41 | /-- error: Bitwuzla thinks it's right but can't trust the wuzla! -/ 42 | #guard_msgs in 43 | example : (x.zeroExtend 64).msb = false := by 44 | bv_bitwuzla 45 | 46 | /-- error: Bitwuzla thinks it's right but can't trust the wuzla! -/ 47 | #guard_msgs in 48 | example (hx : BitVec.slt 0 x) : (x.signExtend 64).msb = false := by 49 | bv_bitwuzla 50 | 51 | /-- error: Bitwuzla thinks it's right but can't trust the wuzla! -/ 52 | #guard_msgs in 53 | example : x &&& y = y &&& x := by 54 | bv_bitwuzla 55 | 56 | /-- error: Bitwuzla thinks it's right but can't trust the wuzla! -/ 57 | #guard_msgs in 58 | example : x ||| y = y ||| x := by 59 | bv_bitwuzla 60 | 61 | /-- error: Bitwuzla thinks it's right but can't trust the wuzla! -/ 62 | #guard_msgs in 63 | example : ~~~(x ||| y) = (~~~y &&& ~~~x) := by 64 | bv_bitwuzla 65 | 66 | /-- error: Bitwuzla thinks it's right but can't trust the wuzla! -/ 67 | #guard_msgs in 68 | example : x <<< 32 = 0 := by 69 | bv_bitwuzla 70 | 71 | /-- error: Bitwuzla thinks it's right but can't trust the wuzla! -/ 72 | #guard_msgs in 73 | example : x >>> x = 0 := by 74 | bv_bitwuzla 75 | 76 | /-- error: Bitwuzla thinks it's right but can't trust the wuzla! -/ 77 | #guard_msgs in 78 | example : (x ++ x).extractLsb' 0 32 = x := by 79 | bv_bitwuzla 80 | 81 | /-- error: Bitwuzla thinks it's right but can't trust the wuzla! -/ 82 | #guard_msgs in 83 | example (h : a = c) : (a → b) = (!c || b) := by 84 | bv_bitwuzla 85 | 86 | /-- error: Bitwuzla thinks it's right but can't trust the wuzla! -/ 87 | #guard_msgs in 88 | example : ite a b c = ite (!a) c b := by 89 | bv_bitwuzla 90 | 91 | /-- error: Bitwuzla thinks it's right but can't trust the wuzla! -/ 92 | #guard_msgs in 93 | example {x y : BitVec 32} (h : BitVec.slt 0 x) : BitVec.sshiftRight' x y = x >>> y := by 94 | bv_bitwuzla 95 | -------------------------------------------------------------------------------- /smtcomp/2025/refs.bib: -------------------------------------------------------------------------------- 1 | @inproceedings{cadical, 2 | author = {Armin Biere and 3 | Tobias Faller and 4 | Katalin Fazekas and 5 | Mathias Fleury and 6 | Nils Froleyks and 7 | Florian Pollitt}, 8 | editor = {Arie Gurfinkel and 9 | Vijay Ganesh}, 10 | title = {{CaDiCaL 2.0}}, 11 | booktitle = {Computer Aided Verification - 36th International Conference, {CAV} 12 | 2024, Montreal, QC, Canada, July 24-27, 2024, Proceedings, Part {I}}, 13 | series = {Lecture Notes in Computer Science}, 14 | volume = {14681}, 15 | pages = {133--152}, 16 | publisher = {Springer}, 17 | year = {2024}, 18 | doi = {10.1007/978-3-031-65627-9\_7}, 19 | } 20 | @inproceedings{bitwuzla, 21 | author = {Aina Niemetz and 22 | Mathias Preiner}, 23 | editor = {Constantin Enea and 24 | Akash Lal}, 25 | title = {Bitwuzla}, 26 | booktitle = {Computer Aided Verification - 35th International Conference, {CAV} 27 | 2023, Paris, France, July 17-22, 2023, Proceedings, Part {II}}, 28 | series = {Lecture Notes in Computer Science}, 29 | volume = {13965}, 30 | pages = {3--17}, 31 | publisher = {Springer}, 32 | year = {2023}, 33 | url = {https://doi.org/10.1007/978-3-031-37703-7\_1}, 34 | doi = {10.1007/978-3-031-37703-7\_1}, 35 | timestamp = {Fri, 21 Jul 2023 17:55:59 +0200}, 36 | biburl = {https://dblp.org/rec/conf/cav/NiemetzP23.bib}, 37 | bibsource = {dblp computer science bibliography, https://dblp.org} 38 | } 39 | @inproceedings{lean4, 40 | title={The {Lean} 4 Theorem Prover and Programming Language}, 41 | author={de Moura, Leonardo and Ullrich, Sebastian}, 42 | booktitle={International Conference on Automated Deduction}, 43 | pages={625--635}, 44 | year={2021}, 45 | organization={Springer} 46 | } 47 | @inproceedings{lrat, 48 | author="Cruz-Filipe, Lu{\'i}s 49 | and Heule, Marijn J. H. 50 | and Hunt, Warren A. 51 | and Kaufmann, Matt 52 | and Schneider-Kamp, Peter", 53 | editor="de Moura, Leonardo", 54 | title="Efficient Certified RAT Verification", 55 | booktitle="Automated Deduction -- CADE 26", 56 | year="2017", 57 | publisher="Springer International Publishing", 58 | address="Cham", 59 | pages="220--236", 60 | isbn="978-3-319-63046-5" 61 | } 62 | 63 | @online{leanwuzla, 64 | title="Leanwuzla", 65 | url={https://github.com/hargoniX/leanwuzla} 66 | } 67 | 68 | @online{lrat-trim, 69 | title="lrat-trim", 70 | url={https://github.com/arminbiere/lrat-trim} 71 | } 72 | @article{barendregt2005challenge, 73 | title={The challenge of computer mathematics}, 74 | author={Barendregt, Henk and Wiedijk, Freek}, 75 | journal={Philosophical Transactions of the Royal Society A: Mathematical, Physical and Engineering Sciences}, 76 | volume={363}, 77 | number={1835}, 78 | pages={2351--2375}, 79 | year={2005}, 80 | publisher={The Royal Society London} 81 | } 82 | -------------------------------------------------------------------------------- /Leanwuzla/NoKernel.lean: -------------------------------------------------------------------------------- 1 | import Leanwuzla.Basic 2 | 3 | open Lean Std.Sat Std.Tactic.BVDecide 4 | open Elab.Tactic.BVDecide 5 | open Elab.Tactic.BVDecide.Frontend 6 | 7 | def runSolver (cnf : CNF Nat) (solver : System.FilePath) (lratPath : System.FilePath) 8 | (trimProofs : Bool) (timeout : Nat) (binaryProofs : Bool) : 9 | CoreM (Except (Array (Bool × Nat)) (Array LRAT.IntAction)) := do 10 | IO.FS.withTempFile fun cnfHandle cnfPath => do 11 | withTraceNode `sat (fun _ => return "Serializing SAT problem to DIMACS file") do 12 | -- lazyPure to prevent compiler lifting 13 | cnfHandle.putStr (← IO.lazyPure (fun _ => cnf.dimacs)) 14 | cnfHandle.flush 15 | 16 | let res ← 17 | withTraceNode `sat (fun _ => return "Running SAT solver") do 18 | External.satQuery solver cnfPath lratPath timeout binaryProofs 19 | if let .sat assignment := res then 20 | return .error assignment 21 | 22 | let lratProof ← 23 | withTraceNode `sat (fun _ => return "Obtaining LRAT certificate") do 24 | LratCert.load lratPath trimProofs 25 | 26 | return .ok lratProof 27 | 28 | def decideSmtNoKernel (type : Expr) : SolverM UInt8 := do 29 | let solver ← TacticContext.new.determineSolver 30 | let g := (← Meta.mkFreshExprMVar type).mvarId! 31 | let (_, g) ← g.introsP 32 | trace[Meta.Tactic.bv] m!"Working on goal: {g}" 33 | try 34 | g.withContext $ IO.FS.withTempFile fun _ lratPath => do 35 | let cfg ← SolverM.getBVDecideConfig 36 | match ← Normalize.bvNormalize g cfg with 37 | | some g => 38 | let bvExpr := (← M.run <| reflectBV g).bvExpr 39 | 40 | let entry ← 41 | withTraceNode `bv (fun _ => return "Bitblasting BVLogicalExpr to AIG") do 42 | -- lazyPure to prevent compiler lifting 43 | IO.lazyPure (fun _ => bvExpr.bitblast) 44 | let aigSize := entry.aig.decls.size 45 | trace[Meta.Tactic.bv] s!"AIG has {aigSize} nodes." 46 | 47 | let (cnf, _) ← 48 | withTraceNode `sat (fun _ => return "Converting AIG to CNF") do 49 | -- lazyPure to prevent compiler lifting 50 | IO.lazyPure (fun _ => 51 | let (entry, map) := entry.relabelNat' 52 | let cnf := Std.Sat.AIG.toCNF entry 53 | (cnf, map) 54 | ) 55 | 56 | let res ← 57 | withTraceNode `sat (fun _ => return "Obtaining external proof certificate") do 58 | runSolver cnf solver lratPath cfg.trimProofs cfg.timeout cfg.binaryProofs 59 | 60 | match res with 61 | | .ok cert => 62 | let certFine ← 63 | withTraceNode `sat (fun _ => return "Verifying LRAT certificate") do 64 | -- lazyPure to prevent compiler lifting 65 | IO.lazyPure (fun _ => LRAT.check cert cnf) 66 | if certFine then 67 | logInfo "unsat" 68 | return (0 : UInt8) 69 | else 70 | logInfo "Error: Failed to check LRAT cert" 71 | return (1 : UInt8) 72 | | .error .. => 73 | logInfo "sat" 74 | return (0 : UInt8) 75 | | none => 76 | logInfo "unsat" 77 | return (0 : UInt8) 78 | catch e => 79 | -- TODO: improve handling of sat cases. This is a temporary workaround. 80 | let message ← e.toMessageData.toString 81 | if message.startsWith "None of the hypotheses are in the supported BitVec fragment" then 82 | -- We fully support SMT-LIB v2.6. Getting the above error message means 83 | -- the goal was reduced to `False` with only `True` as an assumption. 84 | logInfo "sat" 85 | return (0 : UInt8) 86 | else 87 | logError m!"Error: {e.toMessageData}" 88 | return (1 : UInt8) 89 | -------------------------------------------------------------------------------- /Leanwuzla/Sexp/Dsl.lean: -------------------------------------------------------------------------------- 1 | /- 2 | Copyright (c) 2021-2022 by the authors listed in the file AUTHORS and their 3 | institutional affiliations. All rights reserved. 4 | Released under Apache 2.0 license as described in the file LICENSE. 5 | Authors: Wojciech Nawrocki 6 | Source: https://github.com/ufmg-smite/lean-smt/blob/main/Smt/Dsl/Sexp.lean 7 | -/ 8 | import Lean.Parser 9 | 10 | import Leanwuzla.Sexp.Basic 11 | 12 | section 13 | open Lean.Parser 14 | open Lean.PrettyPrinter 15 | 16 | /-- Like `ident` but with no splitting on dots and accepts anything that's not whitespace 17 | or parentheses. So e.g. `<=` works. -/ 18 | def generalIdent : Parser := 19 | withAntiquot (mkAntiquot "generalIdent" `generalIdent) { 20 | fn := fun c s => 21 | let startPos := s.pos 22 | let s := takeWhile1Fn (fun c => !("(){}[].".contains c) ∧ !c.isWhitespace) "expected generalized identifier" c s 23 | mkNodeToken `generalIdent startPos c s } 24 | 25 | def Lean.TSyntax.getGeneralId : TSyntax `generalIdent → String 26 | | ⟨.node _ `generalIdent args⟩ => args[0]!.getAtomVal 27 | | _ => unreachable! 28 | 29 | @[combinator_formatter generalIdent] def generalIdent.formatter : Formatter := pure () 30 | @[combinator_parenthesizer generalIdent] def generalIdent.parenthesizer : Parenthesizer := pure () 31 | end 32 | 33 | instance : ToSexp String := ⟨Sexp.atom⟩ 34 | 35 | instance [ToSexp α] : CoeOut α Sexp := ⟨ToSexp.toSexp⟩ 36 | 37 | declare_syntax_cat sexp 38 | declare_syntax_cat slist 39 | 40 | syntax generalIdent : sexp 41 | syntax "(" ")" : sexp 42 | syntax "(" slist ")" : sexp 43 | syntax "{" term "}" : sexp 44 | 45 | syntax sexp : slist 46 | syntax "...{" term "}" : slist 47 | syntax sexp slist : slist 48 | syntax "...{" term "}" slist : slist 49 | 50 | syntax "sexp!{" sexp "}" : term 51 | syntax "slist!{" slist "}" : term 52 | 53 | macro_rules 54 | | `(slist| $s:sexp) => `([sexp!{$s}]) 55 | | `(slist| ...{ $t:term }) => `(($t : List Sexp)) 56 | | `(slist| $s:sexp $ss:slist) => `(sexp!{$s} :: slist!{$ss}) 57 | | `(slist| ...{ $t:term } $ss:slist) => `(($t : List Sexp) ++ slist!{$ss}) 58 | 59 | macro_rules 60 | | `(sexp| $a:generalIdent) => `(Sexp.atom $(Lean.quote a.getGeneralId)) 61 | | `(sexp| ()) => `(Sexp.expr []) 62 | | `(sexp| ( $ss:slist )) => `(Sexp.expr slist!{$ss}) 63 | | `(sexp| { $t:term }) => `(($t : Sexp)) 64 | 65 | /-- S-expressions can be written using `sexp!{...}` syntax. -/ 66 | macro_rules 67 | | `(sexp!{ $s:sexp }) => `(sexp|$s) 68 | | `(slist!{ $ss:slist }) => `(slist|$ss) 69 | 70 | syntax "sexps!{" "}" : term 71 | syntax "sexps!{" slist "}" : term 72 | macro_rules 73 | | `(sexps!{ }) => `(([] : List Sexp)) 74 | | `(sexps!{ $ss:slist }) => `(slist!{$ss}) 75 | 76 | instance : Repr Sexp where 77 | reprPrec s _ := s!"sexp!\{{toString s}}" 78 | 79 | /-- info: sexp!{()} -/ 80 | #guard_msgs in #eval sexp!{()} 81 | /-- info: sexp!{foo} -/ 82 | #guard_msgs in #eval sexp!{foo} 83 | /-- info: sexp!{foo} -/ 84 | #guard_msgs in #eval sexp!{{Sexp.atom "foo"}} 85 | /-- info: sexp!{(foo)} -/ 86 | #guard_msgs in #eval sexp!{(foo)} 87 | /-- info: sexp!{(foo bar)} -/ 88 | #guard_msgs in #eval sexp!{(foo bar)} 89 | /-- info: sexp!{(foo bar)} -/ 90 | #guard_msgs in #eval sexp!{(foo {Sexp.atom "bar"})} 91 | /-- info: sexp!{(foo bar)} -/ 92 | #guard_msgs in #eval sexp!{({Sexp.atom "foo"} bar)} 93 | /-- info: sexp!{(foo bar)} -/ 94 | #guard_msgs in #eval sexp!{(foo ...{[Sexp.atom "bar"]})} 95 | /-- info: sexp!{(foo bar)} -/ 96 | #guard_msgs in #eval sexp!{(...{[Sexp.atom "foo"]} bar)} 97 | /-- info: sexp!{(foo bar)} -/ 98 | #guard_msgs in #eval sexp!{(...{[Sexp.atom "foo"]} ...{[Sexp.atom "bar"]})} 99 | /-- info: sexp!{(foo bar)} -/ 100 | #guard_msgs in #eval sexp!{(...{[Sexp.atom "foo", Sexp.atom "bar"]})} 101 | /-- info: sexp!{(foo bar)} -/ 102 | #guard_msgs in #eval sexp!{{Sexp.expr [Sexp.atom "foo", Sexp.atom "bar"]}} 103 | /-- info: [] -/ 104 | #guard_msgs in #eval sexps!{} 105 | /-- info: [sexp!{()}] -/ 106 | #guard_msgs in #eval sexps!{()} 107 | /-- info: [sexp!{()}, sexp!{()}] -/ 108 | #guard_msgs in #eval sexps!{() ()} 109 | -------------------------------------------------------------------------------- /Leanwuzla/Sexp/Basic.lean: -------------------------------------------------------------------------------- 1 | /- 2 | Copyright (c) 2021-2022 by the authors listed in the file AUTHORS and their 3 | institutional affiliations. All rights reserved. 4 | Released under Apache 2.0 license as described in the file LICENSE. 5 | Authors: Wojciech Nawrocki 6 | Source: https://github.com/ufmg-smite/lean-smt/blob/main/Smt/Data/Sexp.lean 7 | -/ 8 | 9 | import Std.Internal.Parsec.String 10 | 11 | /-- The type of S-expressions. -/ 12 | inductive Sexp where 13 | | atom : String → Sexp 14 | | expr : List Sexp → Sexp 15 | deriving Repr, BEq, Inhabited 16 | 17 | class ToSexp (α : Type u) where 18 | toSexp : α → Sexp 19 | 20 | namespace Sexp 21 | 22 | def isAtom : Sexp → Bool 23 | | atom _ => true 24 | | _ => false 25 | 26 | def isExpr : Sexp → Bool 27 | | expr _ => true 28 | | _ => false 29 | 30 | partial def serialize : Sexp → String 31 | | atom s => s 32 | | expr ss => "(" ++ (" ".intercalate <| ss.map serialize) ++ ")" 33 | 34 | def serializeMany (ss : List Sexp) : String := 35 | ss.map serialize |> "\n".intercalate 36 | 37 | instance : ToString Sexp := 38 | ⟨serialize⟩ 39 | 40 | namespace Parser 41 | 42 | open Std.Internal.Parsec String 43 | 44 | /-- Parse the s-expression grammar. Supported token kinds are more or less as in 45 | https://smtlib.cs.uiowa.edu/papers/smt-lib-reference-v2.6-r2021-05-12.pdf: 46 | - parentheses `(`/`)` 47 | - symbols `abc` 48 | - quoted symbols `|abc|` 49 | - string literals `"abc"` 50 | - comments `; abc` 51 | -/ 52 | 53 | def comment : Parser Unit := do 54 | skipChar ';' 55 | discard <| many (satisfy (· != '\n')) 56 | skipChar '\n' <|> eof 57 | 58 | def misc : Parser Unit := do 59 | ws 60 | discard <| many (comment *> ws) 61 | 62 | def strLit : Parser String := do 63 | let cstart ← pchar '"' 64 | let s ← manyCharsCore (satisfy (· ≠ '"')) cstart.toString 65 | let cend ← pchar '"' 66 | return s.push cend 67 | 68 | def quotedSym : Parser String := do 69 | let cstart ← pchar '|' 70 | let s ← manyCharsCore (satisfy (· ≠ '|')) cstart.toString 71 | let cend ← pchar '|' 72 | return s.push cend 73 | 74 | def sym : Parser String := 75 | let filter c := !c.isWhitespace && c != '(' && c != ')' && c != '|' && c != '"' && c != ';' 76 | many1Chars (satisfy filter) 77 | 78 | def atom : Parser Sexp := do 79 | let a ← strLit <|> quotedSym <|> sym 80 | return Sexp.atom a 81 | 82 | /-- 83 | Parse all the s-expressions in the given string. For example, `"(abc) (def)"` 84 | contains two. Note that the string may contain extra data, but parsing will 85 | always succeed. 86 | -/ 87 | def manySexps : Parser (List Sexp) := do 88 | let mut stack : List (List Sexp) := [] 89 | let mut curr := [] 90 | let mut next ← misc *> peek? 91 | while h : next.isSome do 92 | match next.get h with 93 | | '(' => 94 | skipChar '(' 95 | stack := curr :: stack 96 | curr := [] 97 | | ')' => 98 | match stack with 99 | | [] => 100 | return curr.reverse 101 | | sexp :: sexps => 102 | skipChar ')' 103 | stack := sexps 104 | curr := .expr curr.reverse :: sexp 105 | | _ => 106 | curr := (← atom) :: curr 107 | next ← misc *> peek? 108 | if !stack.isEmpty then 109 | fail "expected ')'" 110 | return curr.reverse 111 | 112 | def expr : Parser Sexp := do 113 | skipChar '(' 114 | let sexps ← manySexps 115 | skipChar ')' 116 | return Sexp.expr sexps 117 | 118 | /-- 119 | Parse a single s-expression. Note that the string may contain extra data, but 120 | parsing will succeed as soon as a single s-expr is complete. 121 | -/ 122 | def sexp : Parser Sexp := 123 | atom <|> expr 124 | 125 | /-- 126 | Parse all the s-expressions in the given string. For example, `"(abc) (def)"` 127 | contains two. Parsing fails if there is any extra data after the last s-expr. 128 | -/ 129 | def manySexps! : Parser (List Sexp) := do 130 | let sexps ← manySexps 131 | eof 132 | return sexps 133 | 134 | /-- 135 | Parse a single s-expression. Parsing fails if there is any extra data after the 136 | s-expr. 137 | -/ 138 | def sexp! : Parser Sexp := do 139 | let s ← sexp 140 | eof 141 | return s 142 | 143 | end Sexp.Parser 144 | -------------------------------------------------------------------------------- /smtcomp/2025/main.typ: -------------------------------------------------------------------------------- 1 | #import "@preview/charged-ieee:0.1.3": ieee 2 | 3 | #show: ieee.with( 4 | title: [`bv_decide` at the SMT-COMP 2025], 5 | abstract: [ 6 | This paper presents our SMT solver `bv_decide` as submitted to SMT-COMP 2025. 7 | `bv_decide` is a solver for the quantifier-free theory of bit-vectors. 8 | ], 9 | authors: ( 10 | ( 11 | name: "Henrik Böving", 12 | organization: [Lean FRO], 13 | ), 14 | ( 15 | name: "Siddharth Bhat", 16 | organization: [University of Cambridge], 17 | ), 18 | ( 19 | name: "Alex Keizer", 20 | organization: [University of Cambridge], 21 | ), 22 | ( 23 | name: "Luisa Cicolini", 24 | organization: [University of Cambridge], 25 | ), 26 | ( 27 | name: "Leon Frenot", 28 | organization: [ENS Lyon], 29 | ), 30 | ( 31 | name: "Abdalrhman Mohamed", 32 | organization: [Stanford University], 33 | ), 34 | ( 35 | name: "Léo Stefanesco", 36 | organization: [University of Cambridge], 37 | ), 38 | ( 39 | name: "Harun Khan", 40 | organization: [Stanford University], 41 | ), 42 | ( 43 | name: "Josh Clune", 44 | organization: [Carnegie Mellon University], 45 | ), 46 | ( 47 | name: "Clark Barrett", 48 | organization: [Stanford University], 49 | ), 50 | ( 51 | name: "Tobias Grosser", 52 | organization: [University of Cambridge], 53 | ), 54 | ), 55 | bibliography: bibliography("refs.bib"), 56 | ) 57 | 58 | = Introduction 59 | 60 | `bv_decide` is a Lean 4 @lean4 tactic for deciding quantifier-free bit-vector `QF_BV` formulas. 61 | It eagerly determines satisfiability by rewriting and simplifying bit-vector formulas, 62 | bit-blasting them into propositional logic, and invoking a SAT solver with certificate checking. 63 | 64 | Algorithmically, `bv_decide` follows the approach of the Bitwuzla SMT solver @bitwuzla, with the 65 | crucial twist that the solver is implemented and verified as part of the Lean 4 theorem prover 66 | and programming language. Furthermore, to ensure the end-to-end correctness of results, 67 | `bv_decide` also checks the LRAT @lrat certificate produced by the SAT backend. 68 | 69 | = Core (Verified) Components 70 | 71 | 1. *Bitvector Rewrites:* Inspired by Bitwuzla's equational rewrite rules to simplify `QF_BV` formulas. 72 | All rewrite rules are formally verified in Lean for arbitrary bit-widths, ensuring semantic equivalence. 73 | 74 | 2. *Bit-Blasting Engine:* Transforms simplified bit-vector expressions into an And-Inverter Graph (AIG) 75 | and then into CNF. Built atop a verified AIG framework that uses Lean's imperative arrays and 76 | hash maps for efficient in-place operations. 77 | 78 | 3. *LRAT Checker:* After encoding the AIG to CNF, `bv_decide` invokes the SAT backend on the CNF problem, 79 | which emits an LRAT certificate if the problem is UNSAT. The verified LRAT checker then validates 80 | the UNSAT proof and trims unused proof steps, following the `lrat-trim` strategy @lrat-trim. 81 | 82 | = Certifying Components 83 | 84 | The following components are not verified, and instead output certificates that are checked by other 85 | components of `bv_decide`: 86 | 87 | 1. *SAT Backend:* The SAT backend does not need to be verified and is treated as a black box. However, 88 | we require it to produce a valid LRAT certificate, which is then checked by the verified LRAT checker. 89 | 90 | 2. *Rewrite Engine:* The proof generated by this engine during the transformation of the original 91 | bit-vector formula is checked by the Lean kernel. 92 | 93 | = Trusted Code Base 94 | 95 | Despite end-to-end verification, three low-probability sources of error remain: 96 | 97 | 1. *Parsing Errors*, which may cause `bv_decide` to misinterpret SMT-LIB input. `bv_decide` is 98 | developed as an integrated tactic for Lean and thus does not natively support the SMT-LIB format. 99 | To use it in SMT-COMP, we developed a wrapper called `leanwuzla` @leanwuzla that translates 100 | SMT-LIB queries into Lean propositions for `bv_decide` to decide. This wrapper is not formally verified. 101 | However, extensive testing on SMT-LIB 2024 provides high confidence in its correctness. 102 | 103 | 2. *Compiler Bugs* in Lean 4, which might produce a binary misaligned with source semantics. 104 | Again, successful runs on SMT-LIB 2024 benchmarks offer indirect validation. 105 | 106 | 3. *Kernel Unsoundness*, which could undermine proof checking. The Lean kernel follows the De Bruijn criterion 107 | @barendregt2005challenge and has a small, audited codebase, making serious soundness failures extremely unlikely. 108 | Most soundness issues in the past have been due to proofs specifically crafted to exploit particular 109 | implementation vulnerabilities. Since we perform fairly standard reasoning, we do not believe this 110 | to be a serious concern. 111 | 112 | = Competition Entry 113 | 114 | `bv_decide` competes in the single-query track of the `QF_BV` logic at SMT-COMP 2025, 115 | using CaDiCaL 2.1.2 @cadical as its SAT solver. We submit `bv_decide` in two configurations: 116 | - *With kernel checking enabled:* fully verified, with end-to-end correctness guarantees. 117 | - *With kernel checking disabled:* significantly faster, but only partially verified. 118 | 119 | Disabling kernel checking improves performance substantially, but at the cost of expanding the 120 | trusted code base to include Lean's rewrite engine and sacrificing full end-to-end certification. 121 | Nevertheless, the configuration remains far more trustworthy than most solvers: the rewriting rules, 122 | bit-blasting engine, and LRAT certificate validation are still formally verified. 123 | 124 | = License 125 | 126 | Both `bv_decide` and its SMT-LIB wrapper `leanwuzla` are released under the Apache 2.0 license. 127 | -------------------------------------------------------------------------------- /Main.lean: -------------------------------------------------------------------------------- 1 | import Cli 2 | 3 | import Leanwuzla.Parser 4 | import Leanwuzla.Basic 5 | import Leanwuzla.NoKernel 6 | 7 | open Lean 8 | 9 | def parseSmt2File (path : System.FilePath) : MetaM Expr := do 10 | let query ← IO.FS.readFile path 11 | ofExcept (Parser.parseSmt2Query query) 12 | 13 | 14 | open Elab in 15 | def decideSmt (type : Expr) : SolverM UInt8 := do 16 | let mv ← Meta.mkFreshExprMVar type 17 | let (_, mv') ← mv.mvarId!.introsP 18 | trace[Meta.Tactic.bv] m!"Working on goal: {mv'}" 19 | try 20 | mv'.withContext $ IO.FS.withTempFile fun _ lratFile => do 21 | let cfg ← SolverM.getBVDecideConfig 22 | let ctx ← (Tactic.BVDecide.Frontend.TacticContext.new lratFile cfg).run' { declName? := `lrat } 23 | discard <| Tactic.BVDecide.Frontend.bvDecide mv' ctx 24 | catch e => 25 | -- TODO: improve handling of sat cases. This is a temporary workaround. 26 | let message ← e.toMessageData.toString 27 | if message.startsWith "The prover found a counterexample" || 28 | message.startsWith "None of the hypotheses are in the supported BitVec fragment" then 29 | -- We fully support SMT-LIB v2.6. Getting the above error message means 30 | -- the goal was reduced to `False` with only `True` as an assumption. 31 | logInfo "sat" 32 | return (0 : UInt8) 33 | else 34 | logError m!"Error: {e.toMessageData}" 35 | return (1 : UInt8) 36 | let value ← instantiateExprMVars mv 37 | try 38 | Lean.addDecl (.thmDecl { name := ← Lean.mkAuxDeclName, levelParams := [], type, value }) 39 | logInfo "unsat" 40 | return 0 41 | catch e => 42 | logError m!"Error: {e.toMessageData}" 43 | return 1 44 | 45 | def typeCheck (e : Expr) : SolverM UInt8 := do 46 | try 47 | let defn := .defnDecl { 48 | name := ← Lean.mkAuxDeclName 49 | levelParams := [] 50 | type := .sort .zero 51 | value := e 52 | hints := .regular 0 53 | safety := .safe 54 | } 55 | Lean.addDecl defn 56 | return 0 57 | catch e => 58 | logError m!"Error: {e.toMessageData}" 59 | return 1 60 | 61 | def parseAndDecideSmt2File : SolverM UInt8 := do 62 | try 63 | let goalType ← parseSmt2File (← SolverM.getInput) 64 | if ← SolverM.getParseOnly then 65 | logInfo m!"Goal:\n{goalType}" 66 | typeCheck goalType 67 | else 68 | if ← SolverM.getKernelDisabled then 69 | decideSmtNoKernel goalType 70 | else 71 | decideSmt goalType 72 | finally 73 | printTraces 74 | Lean.Language.reportMessages (← Core.getMessageLog) (← getOptions) 75 | 76 | section Cli 77 | 78 | open Cli 79 | 80 | unsafe def runLeanwuzlaCmd (p : Parsed) : IO UInt32 := do 81 | let options := argsToOpts p 82 | let context := argsToContext p 83 | Lean.initSearchPath (← Lean.findSysroot) 84 | enableInitializersExecution 85 | let env ← importModules #[`Std.Tactic.BVDecide, `Leanwuzla.Aux] {} 0 (loadExts := true) 86 | let coreContext := { fileName := "leanwuzla", fileMap := default, options } 87 | let coreState := { env } 88 | let code ← SolverM.run parseAndDecideSmt2File context coreContext coreState 89 | IO.Process.exit code 90 | where 91 | argsToOpts (p : Parsed) : Options := Id.run do 92 | let mut opts := Options.empty 93 | 94 | if p.hasFlag "verbose" then 95 | opts := 96 | opts 97 | |>.setBool `trace.Meta.Tactic.bv true 98 | |>.setBool `trace.Meta.Tactic.sat true 99 | --|>.setBool `trace.profiler true 100 | 101 | if p.hasFlag "vsimp" then 102 | opts := 103 | opts 104 | |>.setBool `trace.Meta.Tactic.simp true 105 | 106 | if p.hasFlag "disableKernel" then 107 | opts := 108 | opts 109 | |>.setBool `debug.skipKernelTC true 110 | 111 | if p.hasFlag "parseOnly" then 112 | opts := 113 | opts 114 | |>.setNat `pp.maxSteps 100000000000000000 115 | |>.setNat `pp.deepTerms.threshold 100000000 116 | 117 | opts := opts.setNat `maxHeartbeats <| p.flag! "maxHeartbeats" |>.as! Nat 118 | opts := opts.setNat `maxRecDepth <| p.flag! "maxRecDepth" |>.as! Nat 119 | opts := opts.setNat `trace.profiler.threshold <| p.flag! "pthreshold" |>.as! Nat 120 | opts := opts.setNat `exponentiation.threshold <| p.flag! "expthreshold" |>.as! Nat 121 | 122 | return opts 123 | 124 | argsToContext (p : Parsed) : Context := 125 | { 126 | acNf := p.hasFlag "acnf" 127 | parseOnly := p.hasFlag "parseOnly" 128 | timeout := p.flag! "timeout" |>.as! Nat 129 | input := p.positionalArg! "input" |>.as! String 130 | maxSteps := p.flag! "maxSteps" |>.as! Nat 131 | disableAndFlatten := p.hasFlag "disableAndFlatten" 132 | disableEmbeddedConstraintSubst := p.hasFlag "disableEmbeddedConstraintSubst" 133 | disableKernel := p.hasFlag "disableKernel" 134 | } 135 | 136 | unsafe def leanwuzlaCmd : Cmd := `[Cli| 137 | leanwuzla VIA runLeanwuzlaCmd; ["0.1.0"] 138 | "Run LeanSAT as an SMT solver on an SMTLIB2 file." 139 | 140 | FLAGS: 141 | v, verbose; "Print profiler trace output from LeanSAT." 142 | acnf; "Activate the normalisation pass up to commutatitvity." 143 | parseOnly; "Only parse and exit right away." 144 | timeout : Nat; "Set the SAT solver timeout in seconds." 145 | 146 | maxHeartbeats : Nat; "Set the maxHeartbeats." 147 | maxRecDepth : Nat; "Set the maxRecDepth." 148 | expthreshold : Nat; "The threshold for maximum exponents. Useful to limit runaway computation." 149 | maxSteps : Nat; "Set the maximum number of simplification steps." 150 | pthreshold : Nat; "The timing threshold for profiler output." 151 | vsimp; "Print the profiler trace output from simp." 152 | disableAndFlatten; "Disable the and flattening pass." 153 | disableEmbeddedConstraintSubst; "Disable the embedded constraints substitution pass." 154 | disableKernel; "Disable the Lean kernel, that is only verify the LRAT cert, no reflection proof" 155 | 156 | ARGS: 157 | input : String; "Path to the smt2 file to work on" 158 | 159 | EXTENSIONS: 160 | defaultValues! #[ 161 | ("timeout", "10"), 162 | ("maxHeartbeats", toString maxHeartbeats.defValue), 163 | ("maxRecDepth", toString maxRecDepth.defValue), 164 | ("pthreshold", toString trace.profiler.threshold.defValue), 165 | ("maxSteps", toString Lean.Meta.Simp.defaultMaxSteps), 166 | ("expthreshold", toString exponentiation.threshold.defValue) 167 | ] 168 | ] 169 | 170 | end Cli 171 | 172 | unsafe def main (args : List String) : IO UInt32 := do 173 | leanwuzlaCmd.validate args 174 | -------------------------------------------------------------------------------- /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 | -------------------------------------------------------------------------------- /Leanwuzla.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 Lean.Elab.Tactic.BVDecide.Frontend.BVDecide 7 | 8 | open Lean.Parser.Tactic 9 | 10 | /-- 11 | Invoke Bitwuzla on an SMT-fied version of a bitvector goal to see if it holds or not. 12 | Does not generate a proof term. 13 | -/ 14 | syntax (name := bvBitwuzla) "bv_bitwuzla" optConfig (str)? : tactic 15 | 16 | /-- 17 | Compare the performance of `bv_decide` and `bv_bitwuzla`. 18 | -/ 19 | syntax (name := bvCompare) "bv_compare" optConfig (str)? : tactic 20 | 21 | namespace Lean.Elab.Tactic.BVDecide 22 | namespace Frontend 23 | 24 | open Std.Sat 25 | open Std.Tactic.BVDecide 26 | open Std.Tactic.BVDecide.Reflect 27 | open Lean.Meta 28 | 29 | 30 | partial def toSMT (expr : BVLogicalExpr) (atomsAssignment : Std.HashMap Nat (Nat × Expr × Bool)) : String := 31 | let (_, buffer) := StateT.run (go expr atomsAssignment) "" 32 | buffer 33 | where 34 | @[inline] 35 | push (str : String) : StateM String Unit := modify fun buf => buf ++ str 36 | 37 | @[specialize] 38 | withParens (arg : StateM String Unit) : StateM String Unit := do 39 | push "(" 40 | arg 41 | push ")" 42 | 43 | @[inline] 44 | pushUnaryOp (opName : String) (arg : StateM String Unit) : StateM String Unit := do 45 | withParens do 46 | push s!"{opName} " 47 | arg 48 | 49 | @[inline] 50 | pushBinaryOp (opName : String) (lhs rhs : StateM String Unit) : StateM String Unit := do 51 | withParens do 52 | push s!"{opName} " 53 | lhs 54 | push " " 55 | rhs 56 | 57 | @[inline] 58 | pushIteOp (c e1 e2 : StateM String Unit) : StateM String Unit := do 59 | withParens do 60 | push s!"ite " 61 | c 62 | push " " 63 | e1 64 | push " " 65 | e2 66 | 67 | go (expr : BVLogicalExpr) (atomsAssignment : Std.HashMap Nat (Nat × Expr × Bool)) : StateM String Unit := do 68 | push "(set-logic QF_BV)\n" 69 | declareConsts atomsAssignment 70 | pushUnaryOp "assert" (goBVLogical expr) 71 | push "\n" 72 | push "(check-sat)\n" 73 | push "(exit)\n" 74 | 75 | declareConsts (atomsAssignment : Std.HashMap Nat (Nat × Expr × Bool)) : StateM String Unit := do 76 | for (atom, (width, _, _)) in atomsAssignment do 77 | push s!"(declare-const x_{atom} (_ BitVec {width}))\n" 78 | 79 | goBVLogical (expr : BVLogicalExpr) : StateM String Unit := do 80 | match expr with 81 | | .literal pred => goBVPred pred 82 | | .const b => 83 | match b with 84 | | true => push "true" 85 | | false => push "false" 86 | | .not expr => pushUnaryOp "not" (goBVLogical expr) 87 | | .ite c e1 e2 => pushIteOp (goBVLogical c) (goBVLogical e1) (goBVLogical e2) 88 | | .gate gate lhs rhs => 89 | let gateStr := 90 | match gate with 91 | | .and => "and" 92 | | .xor => "xor" 93 | | .beq => "=" 94 | | .or => "or" 95 | pushBinaryOp gateStr (goBVLogical lhs) (goBVLogical rhs) 96 | 97 | goBVPred (pred : BVPred) : StateM String Unit := do 98 | match pred with 99 | | .bin lhs op rhs => 100 | let opStr := 101 | match op with 102 | | .eq => "=" 103 | | .ult => "bvult" 104 | pushBinaryOp opStr (goBVExpr lhs) (goBVExpr rhs) 105 | | .getLsbD expr idx => 106 | push s!"(= #b1 ((_ extract {idx} {idx}) " 107 | goBVExpr expr 108 | push "))" 109 | 110 | goBVExpr {w : Nat} (expr : BVExpr w) : StateM String Unit := do 111 | match expr with 112 | | .var idx => push s!"x_{idx}" 113 | | .const val => 114 | let s := (Nat.toDigits 2 val.toNat).asString 115 | let t := (List.replicate (w - s.length) '0').asString 116 | let binStr := t ++ s 117 | push "#b" 118 | push binStr 119 | | .extract start len expr => pushUnaryOp s!"(_ extract {len - 1 + start} {start})" (goBVExpr expr) 120 | | .bin lhs op rhs => 121 | let lhs := goBVExpr lhs 122 | let rhs := goBVExpr rhs 123 | match op with 124 | | .and => pushBinaryOp "bvand" lhs rhs 125 | | .or => pushBinaryOp "bvor" lhs rhs 126 | | .xor => pushBinaryOp "bvxor" lhs rhs 127 | | .add => pushBinaryOp "bvadd" lhs rhs 128 | | .mul => pushBinaryOp "bvmul" lhs rhs 129 | | .udiv => 130 | let zero := goBVExpr <| .const (w := w) 0 131 | withParens do 132 | push "ite " 133 | pushBinaryOp "=" zero rhs 134 | push " " 135 | zero 136 | push " " 137 | pushBinaryOp "bvudiv" lhs rhs 138 | | .umod => 139 | let zero := goBVExpr <| .const (w := w) 0 140 | withParens do 141 | push "ite " 142 | pushBinaryOp "=" zero rhs 143 | push " " 144 | zero 145 | push " " 146 | pushBinaryOp "bvurem" lhs rhs 147 | | .un op operand => 148 | match op with 149 | | .not => pushUnaryOp "bvnot" (goBVExpr operand) 150 | | .rotateLeft n => pushUnaryOp s!"(_ rotate_left {n})" (goBVExpr operand) 151 | | .rotateRight n => pushUnaryOp s!"(_ rotate_right {n})" (goBVExpr operand) 152 | | .arithShiftRightConst n => pushBinaryOp s!"bvashr" (goBVExpr operand) (goBVExpr (bvConst w n)) 153 | | .reverse => panic! "Reverse currently not supported in SMTLIB production" 154 | | .clz => panic! "Reverse currently not supported in SMTLIB production" 155 | | .append lhs rhs _ => pushBinaryOp "concat" (goBVExpr lhs) (goBVExpr rhs) 156 | | .replicate n expr _ => pushUnaryOp s!"(_ repeat {n})" (goBVExpr expr) 157 | | .shiftLeft lhs rhs => pushBinaryOp "bvshl" (goBVExpr lhs) (goBVExpr rhs) 158 | | .shiftRight lhs rhs => pushBinaryOp "bvlshr" (goBVExpr lhs) (goBVExpr rhs) 159 | | .arithShiftRight lhs rhs => pushBinaryOp "bvashr" (goBVExpr lhs) (goBVExpr rhs) 160 | 161 | emitTruncate {w : Nat} (expr : BVExpr w) (targetWidth : Nat) : StateM String Unit := do 162 | pushUnaryOp s!"(_ extract {targetWidth - 1} 0)" (goBVExpr expr) 163 | 164 | @[inline] 165 | bvConst (w : Nat) (n : Nat) : BVExpr w := .const (BitVec.ofNat w n) 166 | 167 | def smtQuery (solverPath : System.FilePath) (problemPath : System.FilePath) (timeout : Nat) : 168 | CoreM External.SolverResult := do 169 | let cmd := solverPath.toString 170 | let opts ← getOptions 171 | let verbose := diagnostics.get opts 172 | let mut args := #[ 173 | problemPath.toString, 174 | "-v=1" 175 | ] 176 | 177 | if verbose then 178 | args := args.push "-p" 179 | 180 | let out? ← External.runInterruptible timeout { cmd, args, stdin := .piped, stdout := .piped, stderr := .null } 181 | match out? with 182 | | .timeout => 183 | throwError "The SMT solver timed out while solving the problem." 184 | | .success { exitCode := exitCode, stdout := stdout, stderr := stderr} => 185 | if exitCode == 255 then 186 | throwError s!"Failed to execute external prover:\n{stderr}" 187 | else 188 | let stdoutLines := stdout.splitOn "\n" 189 | let solvingLine := stdoutLines[stdoutLines.length - 2]! 190 | assert! solvingLine.startsWith "solving_context::time_solve:" 191 | trace[Meta.Tactic.bv] solvingLine 192 | if stdoutLines.contains "sat" then 193 | return .sat #[] 194 | else if stdoutLines.contains "unsat" then 195 | return .unsat 196 | else 197 | throwError s!"The external prover produced unexpected output, stdout:\n{stdout}stderr:\n{stderr}" 198 | 199 | 200 | 201 | private axiom bitwuzlaCorrect (expr : BVLogicalExpr) : expr.Unsat 202 | 203 | def bitwuzlaCounterExample : String := "Bitwuzla found a counter example" 204 | def bitwuzlaSuccess : String := "Bitwuzla thinks it's right but can't trust the wuzla!" 205 | 206 | def bitwuzla (g : MVarId) (reflectionResult : ReflectionResult) (atomsAssignment : Std.HashMap Nat (Nat × Expr × Bool)) 207 | (solverPath : System.FilePath) (cfg : BVDecideConfig) : 208 | MetaM (Except CounterExample UnsatProver.Result) := do 209 | let smt := toSMT reflectionResult.bvExpr atomsAssignment 210 | trace[Meta.Tactic.bv] s!"Encoded as SMT: {smt}" 211 | let res ← 212 | IO.FS.withTempFile fun handle path => do 213 | handle.putStr smt 214 | handle.flush 215 | withTraceNode `bv (fun _ => return "Proving with bitwuzla") do 216 | smtQuery solverPath path cfg.timeout 217 | match res with 218 | | .sat .. => return .error ⟨g, {}, #[]⟩ 219 | | .unsat => return .ok ⟨mkApp (mkConst ``bitwuzlaCorrect) (toExpr reflectionResult.bvExpr), ""⟩ 220 | 221 | def bvBitwuzla (g : MVarId) (solverPath : System.FilePath) (cfg : BVDecideConfig) : 222 | MetaM (Except CounterExample Unit) := do 223 | let some g ← Normalize.bvNormalize g cfg | return .ok () 224 | let unsatProver : UnsatProver := fun g reflectionResult atomsAssignment => do 225 | withTraceNode `bv (fun _ => return "Preparing LRAT reflection term") do 226 | bitwuzla g reflectionResult atomsAssignment solverPath cfg 227 | match ← closeWithBVReflection g unsatProver with 228 | | .ok .. => return .ok () 229 | | .error err => return .error err 230 | 231 | 232 | @[tactic bvBitwuzla] 233 | def evalBvBitwuzla : Tactic := fun 234 | | `(tactic| bv_bitwuzla $cfg:optConfig $solverPath:str) => do 235 | let cfg ← elabBVDecideConfig cfg 236 | liftMetaFinishingTactic fun g => do 237 | match ← bvBitwuzla g solverPath.getString cfg with 238 | | .ok .. => throwError bitwuzlaSuccess 239 | | .error .. => throwError bitwuzlaCounterExample 240 | | `(tactic| bv_bitwuzla $cfg:optConfig) => do 241 | let cfg ← elabBVDecideConfig cfg 242 | liftMetaFinishingTactic fun g => do 243 | match ← bvBitwuzla g "bitwuzla" cfg with 244 | | .ok .. => throwError bitwuzlaSuccess 245 | | .error .. => throwError bitwuzlaCounterExample 246 | | _ => throwUnsupportedSyntax 247 | 248 | structure BitwuzlaPerf where 249 | success : Bool 250 | overallTime : Float 251 | solvingContextTime : Float 252 | 253 | structure LeansatSuccessTimings where 254 | timeRewrite: Float 255 | timeBitBlasting : Float 256 | timeSatSolving : Float 257 | timeLratTrimming : Float 258 | timeLratChecking : Float 259 | 260 | structure LeansatFailureTimings where 261 | timeRewrite : Float 262 | timeSatSolving : Float 263 | 264 | inductive LeansatPerf where 265 | | success (overallTime : Float) (timings : LeansatSuccessTimings) 266 | | failure (overallTime : Float) (timings : LeansatFailureTimings) 267 | 268 | structure Comparision where 269 | bitwuzlaPerf : Option BitwuzlaPerf 270 | leansatPerf : Option LeansatPerf 271 | 272 | def BitwuzlaPerf.toString (perf : BitwuzlaPerf) : String := 273 | if perf.success then 274 | s!"Bitwuzla proved the goal after {perf.overallTime}ms, solving context: {perf.solvingContextTime}ms" 275 | else 276 | s!"Bitwuzla provided a counter example after {perf.overallTime}ms, solving context: {perf.solvingContextTime}ms" 277 | 278 | instance : ToString BitwuzlaPerf where 279 | toString := BitwuzlaPerf.toString 280 | 281 | def LeansatSuccessTimings.toString (timings : LeansatSuccessTimings) : String := 282 | let { timeRewrite, timeBitBlasting, timeSatSolving, timeLratTrimming, timeLratChecking } := timings 283 | s!"rewriting {timeRewrite}ms, bitblasting {timeBitBlasting}ms, SAT solving {timeSatSolving}ms, LRAT trimming {timeLratTrimming}ms, LRAT checking {timeLratChecking}ms" 284 | 285 | instance : ToString LeansatSuccessTimings where 286 | toString := LeansatSuccessTimings.toString 287 | 288 | def LeansatFailureTimings.toString (timings : LeansatFailureTimings) : String := 289 | let { timeRewrite, timeSatSolving } := timings 290 | s!"rewriting {timeRewrite} SAT solving {timeSatSolving}ms" 291 | 292 | instance : ToString LeansatFailureTimings where 293 | toString := LeansatFailureTimings.toString 294 | 295 | def LeansatPerf.toString (perf : LeansatPerf) : String := 296 | match perf with 297 | | .success overallTime timings => 298 | s!"LeanSAT proved the goal after {overallTime}ms: {timings}" 299 | | .failure overallTime timings => 300 | s!"LeanSAT provided a counter example after {overallTime}ms: {timings}" 301 | 302 | instance : ToString LeansatPerf where 303 | toString := LeansatPerf.toString 304 | 305 | def Comparision.toString (comp : Comparision) : String := 306 | match comp.bitwuzlaPerf, comp.leansatPerf with 307 | | none, none => "Bitwuzla failed.\nLeanSAT failed." 308 | | some bitwuzla, none => bitwuzla.toString ++ "\nLeanSAT failed." 309 | | none, some leansat => "Bitwuzla failed.\n" ++ leansat.toString 310 | | some bitwuzla, some leansat => bitwuzla.toString ++ "\n" ++ leansat.toString 311 | 312 | instance : ToString Comparision where 313 | toString := Comparision.toString 314 | 315 | def TraceData.durationMs (data : TraceData) : Float := (data.stopTime - data.startTime) * 1000.0 316 | 317 | partial def parseSuccessTrace (traces : PersistentArray TraceElem) : IO LeansatSuccessTimings := do 318 | let traces := traces.toArray.map TraceElem.msg 319 | let (_, time) ← go traces |>.run { 320 | timeRewrite := 0, 321 | timeBitBlasting := 0, 322 | timeSatSolving := 0, 323 | timeLratTrimming := 0, 324 | timeLratChecking := 0, 325 | } 326 | return time 327 | where 328 | go (msgs : Array MessageData) : StateT LeansatSuccessTimings IO Unit := do 329 | for msg in msgs do 330 | match msg with 331 | | .trace data msg children => 332 | let msg ← msg.toString 333 | match msg with 334 | | "Preprocessing goal" => 335 | modify fun s => { s with timeRewrite := TraceData.durationMs data } 336 | | "Bitblasting BVLogicalExpr to AIG" => 337 | modify fun s => { s with timeBitBlasting := TraceData.durationMs data } 338 | | "Running SAT solver" => 339 | modify fun s => { s with timeSatSolving := TraceData.durationMs data } 340 | | "Obtaining LRAT certificate" => 341 | modify fun s => { s with timeLratTrimming := TraceData.durationMs data } 342 | | "Verifying LRAT certificate" | "Compiling expr term" | "Compiling proof certificate term" 343 | | "Compiling reflection proof term" => 344 | modify fun s => { s with timeLratChecking := s.timeLratChecking + TraceData.durationMs data } 345 | | _ => pure () 346 | go children 347 | | .withContext _ msg => go #[msg] 348 | | _ => continue 349 | 350 | partial def parseFailureTrace (traces : PersistentArray TraceElem) : IO LeansatFailureTimings := do 351 | let traces := traces.toArray.map TraceElem.msg 352 | let (_, time) ← go traces |>.run { timeRewrite := 0.0, timeSatSolving := 0.0 } 353 | return time 354 | where 355 | go (msgs : Array MessageData) : StateT LeansatFailureTimings IO Unit := do 356 | for msg in msgs do 357 | match msg with 358 | | .trace data msg children => 359 | let msg ← msg.toString 360 | match msg with 361 | | "Preprocessing goal" => 362 | modify fun s => { s with timeRewrite := TraceData.durationMs data } 363 | | "Running SAT solver" => 364 | modify fun s => { s with timeSatSolving := TraceData.durationMs data } 365 | | _ => pure () 366 | go children 367 | | .withContext _ msg => go #[msg] 368 | | _ => continue 369 | 370 | partial def evalBitwuzla (g : MVarId) (arg : System.FilePath × BVDecideConfig) : MetaM BitwuzlaPerf := do 371 | let (solverPath, cfg) := arg 372 | g.withContext do 373 | let t1 ← IO.monoNanosNow 374 | let res ← bvBitwuzla g solverPath cfg 375 | let t2 ← IO.monoNanosNow 376 | let overallTime := (Float.ofNat <| t2 - t1) / 1000000.0 377 | let (_, solvingContextTime) ← parseBitwuzlaTrace ((← getTraces).toArray.map TraceElem.msg) |>.run 0 378 | return { success := res.isOk, overallTime, solvingContextTime } 379 | where 380 | parseBitwuzlaTrace (msgs : Array MessageData) : StateT Float IO Unit := do 381 | for msg in msgs do 382 | match msg with 383 | | .trace _ msg children => 384 | let msg ← msg.toString 385 | let pref := "solving_context::time_solve: " 386 | if msg.startsWith pref then 387 | let msg := (msg.splitOn " ")[1]!.dropRight 2 388 | set <| Float.ofInt msg.toNat! 389 | parseBitwuzlaTrace children 390 | | .withContext _ msg => parseBitwuzlaTrace #[msg] 391 | | _ => continue 392 | 393 | def evalLeanSat (g : MVarId) (cfg : TacticContext) : MetaM LeansatPerf := do 394 | g.withContext do 395 | let t1 ← IO.monoNanosNow 396 | let result ← bvDecide' g cfg 397 | let t2 ← IO.monoNanosNow 398 | let overallTime := (Float.ofNat <| t2 - t1) / 1000000.0 399 | 400 | let traces ← getTraces 401 | match result with 402 | | .ok _ => 403 | return .success overallTime (← parseSuccessTrace traces) 404 | | .error _ => 405 | return .failure overallTime (← parseFailureTrace traces) 406 | 407 | def bvCompare (g : MVarId) (solverPath : System.FilePath) (ctx : TacticContext) : MetaM Comparision := do 408 | let bitwuzlaPerf ← measure evalBitwuzla g (solverPath, ctx.config) 409 | let leansatPerf ← measure evalLeanSat g ctx 410 | 411 | return { bitwuzlaPerf , leansatPerf } 412 | where 413 | withFreshTraceState {α : Type} (x : MetaM α) : MetaM α := do 414 | let traces ← getTraceState 415 | resetTraceState 416 | try x finally setTraceState traces 417 | 418 | measure {α : Type _} {β : Type _} (f : MVarId → α → MetaM β) (g : MVarId) (arg : α) : MetaM (Option β) := do 419 | let setTraceOptions (opt : Options) : Options := 420 | opt 421 | |>.setBool `trace.profiler true 422 | |>.setBool `trace.Meta.Tactic.bv true 423 | |>.setBool `trace.Meta.Tactic.sat true 424 | |>.setNat `trace.profiler.threshold 1 425 | 426 | try 427 | withOptions setTraceOptions <| withoutModifyingEnv <| withoutModifyingState <| withFreshTraceState do 428 | f g arg 429 | catch e => 430 | logError e.toMessageData 431 | return none 432 | 433 | @[tactic bvCompare] 434 | def evalBvCompare : Tactic := fun 435 | | `(tactic| bv_compare $cfg:optConfig $solverPath:str) => do 436 | let cfg ← elabBVDecideConfig cfg 437 | IO.FS.withTempFile fun _ lratFile => do 438 | let ctx ← TacticContext.new lratFile cfg 439 | let g ← getMainGoal 440 | let res ← bvCompare g solverPath.getString ctx 441 | logInfo <| toString res 442 | | `(tactic| bv_compare $cfg:optConfig) => do 443 | let cfg ← elabBVDecideConfig cfg 444 | IO.FS.withTempFile fun _ lratFile => do 445 | let ctx ← TacticContext.new lratFile cfg 446 | let g ← getMainGoal 447 | let res ← bvCompare g "bitwuzla" ctx 448 | logInfo <| toString res 449 | | _ => throwUnsupportedSyntax 450 | 451 | end Frontend 452 | end Lean.Elab.Tactic.BVDecide 453 | -------------------------------------------------------------------------------- /Leanwuzla/Parser.lean: -------------------------------------------------------------------------------- 1 | import Leanwuzla.Aux 2 | import Leanwuzla.Sexp 3 | 4 | open Lean 5 | 6 | structure Parser.State where 7 | /-- Current de Bruijn level. -/ 8 | level : Nat := 0 9 | /-- A mapping from variable names to their corresponding type and de Bruijn 10 | level (not index). So, the variables are indexed from the bottom of the 11 | stack rather than from the top (i.e., the order in which the symbols are 12 | introduced in the SMT-LIB file). To compute the de Bruijn index, we 13 | subtract the variable's level from the current level. -/ 14 | bvars : Std.HashMap Name (Expr × Nat) := {} 15 | /-- A mapping from user-defined types (i.e., aliases) to their corresponding 16 | canonical types. We need this mapping to provide correct type-class 17 | instances for user-defined types. -/ 18 | uvars : Std.HashMap Name Expr := {} 19 | deriving Repr 20 | 21 | abbrev ParserM := StateT Parser.State (Except MessageData) 22 | 23 | namespace Parser 24 | 25 | private def mkBool : Expr := 26 | .const ``Bool [] 27 | 28 | private def mkBitVec (w : Nat) : Expr := 29 | .app (.const ``BitVec []) (mkNatLit w) 30 | 31 | private def getBitVecWidth (α : Expr) : ParserM Nat := do 32 | match α with 33 | | .app (.const ``BitVec []) w => return w.nat?.get! 34 | | _ => throw m!"Error: expected BitVec type, got {α}" 35 | 36 | private def mkInstBEqBool : Expr := 37 | mkApp2 (.const ``instBEqOfDecidableEq [0]) mkBool 38 | (.const ``instDecidableEqBool []) 39 | 40 | private def mkInstBEqBitVec (w : Nat) : Expr := 41 | mkApp2 (.const ``instBEqOfDecidableEq [0]) (mkBitVec w) 42 | (.app (.const ``instDecidableEqBitVec []) (mkNatLit w)) 43 | 44 | private def mkBitVecAppend (w v : Nat) : Expr := 45 | mkApp4 (.const ``HAppend.hAppend [0, 0, 0]) 46 | (mkBitVec w) (mkBitVec v) (mkBitVec (w + v)) 47 | (mkApp2 (.const ``BitVec.instHAppendHAddNat []) (mkNatLit w) (mkNatLit v)) 48 | 49 | private def mkBitVecAnd (w : Nat) : Expr := 50 | mkApp4 (.const ``HAnd.hAnd [0, 0, 0]) 51 | (mkBitVec w) (mkBitVec w) (mkBitVec w) 52 | (mkApp2 (.const ``instHAndOfAndOp [0]) (mkBitVec w) 53 | (.app (.const ``BitVec.instAndOp []) (mkNatLit w))) 54 | 55 | private def mkBitVecOr (w : Nat) : Expr := 56 | mkApp4 (.const ``HOr.hOr [0, 0, 0]) 57 | (mkBitVec w) (mkBitVec w) (mkBitVec w) 58 | (mkApp2 (.const ``instHOrOfOrOp [0]) (mkBitVec w) 59 | (.app (.const ``BitVec.instOrOp []) (mkNatLit w))) 60 | 61 | private def mkBitVecXor (w : Nat) : Expr := 62 | mkApp4 (.const ``HXor.hXor [0, 0, 0]) 63 | (mkBitVec w) (mkBitVec w) (mkBitVec w) 64 | (mkApp2 (.const ``instHXorOfXor [0]) (mkBitVec w) 65 | (.app (.const ``BitVec.instXor []) (mkNatLit w))) 66 | 67 | private def mkBitVecNot (w : Nat) : Expr := 68 | mkApp2 (.const ``Complement.complement [0]) 69 | (mkBitVec w) 70 | (.app (.const ``BitVec.instComplement []) (mkNatLit w)) 71 | 72 | private def mkBitVecNeg (w : Nat) : Expr := 73 | mkApp2 (.const ``Neg.neg [0]) 74 | (mkBitVec w) 75 | (.app (.const ``BitVec.instNeg []) (mkNatLit w)) 76 | 77 | private def mkBitVecAdd (w : Nat) : Expr := 78 | mkApp4 (.const ``HAdd.hAdd [0, 0, 0]) 79 | (mkBitVec w) (mkBitVec w) (mkBitVec w) 80 | (mkApp2 (.const ``instHAdd [0]) (mkBitVec w) 81 | (.app (.const ``BitVec.instAdd []) (mkNatLit w))) 82 | 83 | private def mkBitVecSub (w : Nat) : Expr := 84 | mkApp4 (.const ``HSub.hSub [0, 0, 0]) 85 | (mkBitVec w) (mkBitVec w) (mkBitVec w) 86 | (mkApp2 (.const ``instHSub [0]) (mkBitVec w) 87 | (.app (.const ``BitVec.instSub []) (mkNatLit w))) 88 | 89 | private def mkBitVecMul (w : Nat) : Expr := 90 | mkApp4 (.const ``HMul.hMul [0, 0, 0]) 91 | (mkBitVec w) (mkBitVec w) (mkBitVec w) 92 | (mkApp2 (.const ``instHMul [0]) (mkBitVec w) 93 | (.app (.const ``BitVec.instMul []) (mkNatLit w))) 94 | 95 | private def mkBitVecMod (w : Nat) : Expr := 96 | mkApp4 (.const ``HMod.hMod [0, 0, 0]) 97 | (mkBitVec w) (mkBitVec w) (mkBitVec w) 98 | (mkApp2 (.const ``instHMod [0]) (mkBitVec w) 99 | (.app (.const ``BitVec.instMod []) (mkNatLit w))) 100 | 101 | private def mkBitVecShiftLeft (w : Nat) : Expr := 102 | mkApp4 (.const ``HShiftLeft.hShiftLeft [0, 0, 0]) 103 | (mkBitVec w) (mkBitVec w) (mkBitVec w) 104 | (mkApp2 (.const ``BitVec.instHShiftLeft []) (mkNatLit w) (mkNatLit w)) 105 | 106 | private def mkBitVecShiftRight (w : Nat) : Expr := 107 | mkApp4 (.const ``HShiftRight.hShiftRight [0, 0, 0]) 108 | (mkBitVec w) (mkBitVec w) (mkBitVec w) 109 | (mkApp2 (.const ``BitVec.instHShiftRight []) (mkNatLit w) (mkNatLit w)) 110 | 111 | def smtSymbolToName (s : String) : Name := 112 | let s := if s.startsWith "|" && s.endsWith "|" then s.extract ⟨1⟩ (s.endPos - ⟨1⟩) else s 113 | -- Quote the string if a natural translation to Name fails 114 | if s.toName == .anonymous then 115 | Name.mkSimple s 116 | else 117 | s.toName 118 | 119 | /-- Returns two types: the first is the canonical type and the second is the 120 | user-provided one (mainly for pretty-printing). -/ 121 | def parseSort (s : Sexp) : ParserM (Expr × Expr) := do 122 | match s with 123 | | sexp!{Bool} => 124 | return (mkBool, mkBool) 125 | | sexp!{(_ BitVec {w})} => 126 | let w := w.serialize.toNat! 127 | return (mkBitVec w, mkBitVec w) 128 | | sexp!{({sc} ...{as})} => 129 | let (bsc, sc) ← parseSort sc 130 | let as ← as.mapM parseSort 131 | let (bas, as) := as.unzip 132 | return (mkAppN bsc bas.toArray, mkAppN sc as.toArray) 133 | | .atom n => 134 | let state ← get 135 | let n := smtSymbolToName n 136 | let some (.sort 1, i) := state.bvars[n]? | throw m!"Error: unexpected sort {s}" 137 | let ut := .bvar (state.level - i - 1) 138 | let ct := match state.uvars[n]? with 139 | | some ct => ct 140 | | none => ut 141 | return (ct, ut) 142 | | _ => throw m!"Error: unsupported sort {s}" 143 | 144 | partial def parseTerm (s : Sexp) : ParserM (Expr × Expr) := do 145 | try 146 | go s 147 | catch e => 148 | throw m!"Error: {e}\nfailed to parse term {s}" 149 | where 150 | go (e : Sexp) : ParserM (Expr × Expr) := do 151 | if let sexp!{(let (...{_}) {_})} := e then 152 | -- SMT-LIB supports nesting of parallel let expressions. Both can be 153 | -- very long. So, we use tail-call recursion to avoid stack overflows. 154 | let state ← get 155 | let (bindings, b) := getNestedLetBindingsAndBody [] e 156 | let bindings ← parseNestedBindings bindings 157 | let (tb, b) ← parseTerm b 158 | set state 159 | let e := bindings.foldr (fun (n, t, v) b => .letE n t v b true) b 160 | return (tb, e) 161 | if let sexp!{true} := e then 162 | return (mkBool, .const ``true []) 163 | if let sexp!{false} := e then 164 | return (mkBool, .const ``false []) 165 | if let sexp!{(not {p})} := e then 166 | let (_, p) ← parseTerm p 167 | return (mkBool, .app (.const ``not []) p) 168 | if let sexp!{(=> ...{ps})} := e then 169 | let ps ← ps.mapM (fun p => return (← parseTerm p).snd) 170 | let p := ps.dropLast.foldr (mkApp2 (.const ``implies [])) ps.getLast! 171 | return (mkBool, p) 172 | if let sexp!{(and ...{ps})} := e then 173 | return ← leftAssocOpBool (.const ``and []) ps 174 | if let sexp!{(or ...{ps})} := e then 175 | return ← leftAssocOpBool (.const ``or []) ps 176 | if let sexp!{(xor ...{ps})} := e then 177 | return ← leftAssocOpBool (.const ``xor []) ps 178 | if let sexp!{(= {x} {y})} := e then 179 | let (α, x) ← parseTerm x 180 | let (_, y) ← parseTerm y 181 | let hα ← if α == mkBool 182 | then pure mkInstBEqBool 183 | else pure (mkInstBEqBitVec (← getBitVecWidth α)) 184 | return (mkBool, mkApp4 (.const ``BEq.beq [0]) α hα x y) 185 | if let sexp!{(distinct ...{xs})} := e then 186 | return ← pairwiseDistinct xs 187 | if let sexp!{(ite {c} {t} {e})} := e then 188 | let (_, c) ← parseTerm c 189 | let (α, t) ← parseTerm t 190 | let (_, e) ← parseTerm e 191 | return (α, mkApp4 (mkConst ``cond [1]) α c t e) 192 | if let sexp!{(concat ...{xs})} := e then 193 | let (α, acc) ← parseTerm xs.head! 194 | let w ← getBitVecWidth α 195 | let f := fun (w, acc) x => do 196 | let (β, x) ← parseTerm x 197 | let v ← getBitVecWidth β 198 | return (w + v, mkApp2 (mkBitVecAppend w v) acc x) 199 | let (w, acc) ← xs.tail.foldlM f (w, acc) 200 | return (mkBitVec w, acc) 201 | if let sexp!{(bvand ...{xs})} := e then 202 | return ← leftAssocOpBitVec mkBitVecAnd xs 203 | if let sexp!{(bvor ...{xs})} := e then 204 | return ← leftAssocOpBitVec mkBitVecOr xs 205 | if let sexp!{(bvxor ...{xs})} := e then 206 | return ← leftAssocOpBitVec mkBitVecXor xs 207 | if let sexp!{(bvnot {x})} := e then 208 | let (α, x) ← parseTerm x 209 | let w ← getBitVecWidth α 210 | return (α, .app (mkBitVecNot w) x) 211 | if let sexp!{(bvnand {x} {y})} := e then 212 | let (α, x) ← parseTerm x 213 | let (_, y) ← parseTerm y 214 | let w ← getBitVecWidth α 215 | return (α, mkApp3 (.const ``BitVec.nand []) (mkNatLit w) x y) 216 | if let sexp!{(bvnor {x} {y})} := e then 217 | let (α, x) ← parseTerm x 218 | let (_, y) ← parseTerm y 219 | let w ← getBitVecWidth α 220 | return (α, mkApp3 (.const ``BitVec.nor []) (mkNatLit w) x y) 221 | if let sexp!{(bvxnor {x} {y})} := e then 222 | let (α, x) ← parseTerm x 223 | let (_, y) ← parseTerm y 224 | let w ← getBitVecWidth α 225 | return (α, mkApp3 (.const ``BitVec.xnor []) (mkNatLit w) x y) 226 | if let sexp!{(bvcomp {x} {y})} := e then 227 | let (α, x) ← parseTerm x 228 | let (_, y) ← parseTerm y 229 | let w ← getBitVecWidth α 230 | return (mkBitVec 1, mkApp3 (.const ``BitVec.compare []) (mkNatLit w) x y) 231 | if let sexp!{(bvmul ...{xs})} := e then 232 | return ← leftAssocOpBitVec mkBitVecMul xs 233 | if let sexp!{(bvadd ...{xs})} := e then 234 | return ← leftAssocOpBitVec mkBitVecAdd xs 235 | if let sexp!{(bvsub ...{xs})} := e then 236 | return ← leftAssocOpBitVec mkBitVecSub xs 237 | if let sexp!{(bvneg {x})} := e then 238 | let (α, x) ← parseTerm x 239 | let w ← getBitVecWidth α 240 | return (α, .app (mkBitVecNeg w) x) 241 | if let sexp!{(bvudiv {x} {y})} := e then 242 | let (α, x) ← parseTerm x 243 | let (_, y) ← parseTerm y 244 | let w ← getBitVecWidth α 245 | return (α, mkApp3 (.const ``BitVec.smtUDiv []) (mkNatLit w) x y) 246 | if let sexp!{(bvurem {x} {y})} := e then 247 | let (α, x) ← parseTerm x 248 | let (_, y) ← parseTerm y 249 | let w ← getBitVecWidth α 250 | return (α, mkApp2 (mkBitVecMod w) x y) 251 | if let sexp!{(bvsdiv {x} {y})} := e then 252 | let (α, x) ← parseTerm x 253 | let (_, y) ← parseTerm y 254 | let w ← getBitVecWidth α 255 | return (α, mkApp3 (.const ``BitVec.smtSDiv []) (mkNatLit w) x y) 256 | if let sexp!{(bvsrem {x} {y})} := e then 257 | let (α, x) ← parseTerm x 258 | let (_, y) ← parseTerm y 259 | let w ← getBitVecWidth α 260 | return (α, mkApp3 (.const ``BitVec.srem []) (mkNatLit w) x y) 261 | if let sexp!{(bvsmod {x} {y})} := e then 262 | let (α, x) ← parseTerm x 263 | let (_, y) ← parseTerm y 264 | let w ← getBitVecWidth α 265 | return (α, mkApp3 (.const ``BitVec.smod []) (mkNatLit w) x y) 266 | if let sexp!{(bvshl {x} {y})} := e then 267 | let (α, x) ← parseTerm x 268 | let (_, y) ← parseTerm y 269 | let w ← getBitVecWidth α 270 | return (α, mkApp2 (mkBitVecShiftLeft w) x y) 271 | if let sexp!{(bvlshr {x} {y})} := e then 272 | let (α, x) ← parseTerm x 273 | let (_, y) ← parseTerm y 274 | let w ← getBitVecWidth α 275 | return (α, mkApp2 (mkBitVecShiftRight w) x y) 276 | if let sexp!{(bvashr {x} {y})} := e then 277 | let (α, x) ← parseTerm x 278 | let (_, y) ← parseTerm y 279 | let w ← getBitVecWidth α 280 | return (α, mkApp4 (.const ``BitVec.sshiftRight' []) (mkNatLit w) (mkNatLit w) x y) 281 | if let sexp!{(bvult {x} {y})} := e then 282 | let (α, x) ← parseTerm x 283 | let (_, y) ← parseTerm y 284 | let w ← getBitVecWidth α 285 | return (mkBool, mkApp3 (.const ``BitVec.ult []) (mkNatLit w) x y) 286 | if let sexp!{(bvule {x} {y})} := e then 287 | let (α, x) ← parseTerm x 288 | let (_, y) ← parseTerm y 289 | let w ← getBitVecWidth α 290 | return (mkBool, mkApp3 (.const ``BitVec.ule []) (mkNatLit w) x y) 291 | if let sexp!{(bvugt {x} {y})} := e then 292 | let (α, x) ← parseTerm x 293 | let (_, y) ← parseTerm y 294 | let w ← getBitVecWidth α 295 | return (mkBool, mkApp3 (.const ``BitVec.ugt []) (mkNatLit w) x y) 296 | if let sexp!{(bvuge {x} {y})} := e then 297 | let (α, x) ← parseTerm x 298 | let (_, y) ← parseTerm y 299 | let w ← getBitVecWidth α 300 | return (mkBool, mkApp3 (.const ``BitVec.uge []) (mkNatLit w) x y) 301 | if let sexp!{(bvslt {x} {y})} := e then 302 | let (α, x) ← parseTerm x 303 | let (_, y) ← parseTerm y 304 | let w ← getBitVecWidth α 305 | return (mkBool, mkApp3 (.const ``BitVec.slt []) (mkNatLit w) x y) 306 | if let sexp!{(bvsle {x} {y})} := e then 307 | let (α, x) ← parseTerm x 308 | let (_, y) ← parseTerm y 309 | let w ← getBitVecWidth α 310 | return (mkBool, mkApp3 (.const ``BitVec.sle []) (mkNatLit w) x y) 311 | if let sexp!{(bvsgt {x} {y})} := e then 312 | let (α, x) ← parseTerm x 313 | let (_, y) ← parseTerm y 314 | let w ← getBitVecWidth α 315 | return (mkBool, mkApp3 (.const ``BitVec.sgt []) (mkNatLit w) x y) 316 | if let sexp!{(bvsge {x} {y})} := e then 317 | let (α, x) ← parseTerm x 318 | let (_, y) ← parseTerm y 319 | let w ← getBitVecWidth α 320 | return (mkBool, mkApp3 (.const ``BitVec.sge []) (mkNatLit w) x y) 321 | if let sexp!{((_ extract {i} {j}) {x})} := e then 322 | let i := i.serialize.toNat! 323 | let j := j.serialize.toNat! 324 | let (α, x) ← parseTerm x 325 | let w ← getBitVecWidth α 326 | let start := j 327 | let len := i - j + 1 328 | return (mkBitVec (i - j + 1), mkApp4 (.const ``BitVec.extractLsb' []) (mkNatLit w) (mkNatLit start) (mkNatLit len) x) 329 | if let sexp!{((_ repeat {i}) {x})} := e then 330 | let i := i.serialize.toNat! 331 | let (α, x) ← parseTerm x 332 | let w ← getBitVecWidth α 333 | return (mkBitVec (w * i), mkApp3 (.const ``BitVec.replicate []) (mkNatLit w) (mkNatLit i) x) 334 | if let sexp!{((_ zero_extend {i}) {x})} := e then 335 | let i := i.serialize.toNat! 336 | let (α, x) ← parseTerm x 337 | let w ← getBitVecWidth α 338 | return (mkBitVec (w + i), mkApp3 (.const ``BitVec.setWidth []) (mkNatLit w) (mkNatLit (w + i)) x) 339 | if let sexp!{((_ sign_extend {i}) {x})} := e then 340 | let i := i.serialize.toNat! 341 | let (α, x) ← parseTerm x 342 | let w ← getBitVecWidth α 343 | return (mkBitVec (w + i), mkApp3 (.const ``BitVec.signExtend []) (mkNatLit w) (mkNatLit (w + i)) x) 344 | if let sexp!{((_ rotate_left {i}) {x})} := e then 345 | let i := i.serialize.toNat! 346 | let (α, x) ← parseTerm x 347 | let w ← getBitVecWidth α 348 | return (α, mkApp3 (.const ``BitVec.rotateLeft []) (mkNatLit w) x (mkNatLit i)) 349 | if let sexp!{((_ rotate_right {i}) {x})} := e then 350 | let i := i.serialize.toNat! 351 | let (α, x) ← parseTerm x 352 | let w ← getBitVecWidth α 353 | return (α, mkApp3 (.const ``BitVec.rotateRight []) (mkNatLit w) x (mkNatLit i)) 354 | if let some r ← parseVar? e then 355 | return r 356 | if let some r := parseBVLiteral? s then 357 | return r 358 | if let sexp!{({f} ...{as})} := s then 359 | let (α, f) ← parseTerm f 360 | let as ← as.mapM (fun a => return (← parseTerm a).snd) 361 | return (retType α, mkAppN f as.toArray) 362 | throw m!"Error: unsupported term {e}" 363 | retType : Expr → Expr 364 | | .forallE _ _ b _ => retType b 365 | | e => e 366 | parseVar? (s : Sexp) : ParserM (Option (Expr × Expr)) := do 367 | let .atom n := s | return none 368 | let state ← get 369 | let n := smtSymbolToName n 370 | let some (t, i) := state.bvars[n]? | return none 371 | return some (t, .bvar (state.level - i - 1)) 372 | parseBVLiteral? (s : Sexp) : Option (Expr × Expr) := 373 | match s with 374 | | sexp!{(_ {.atom v} {.atom w})} => 375 | if v.startsWith "bv" then 376 | let v := v.drop 2 377 | let w := w.toNat! 378 | let v := v.toNat! 379 | some (mkBitVec w, mkApp2 (.const ``BitVec.ofNat []) (mkNatLit w) (mkNatLit v)) 380 | else 381 | none 382 | | sexp!{{.atom s}} => 383 | if s.startsWith "#b" then 384 | let s := s.drop 2 385 | let w := s.length 386 | let v := s.foldl (fun v c => v <<< 1 + (if c == '1' then 1 else 0)) 0 387 | some (mkBitVec w, mkApp2 (.const ``BitVec.ofNat []) (mkNatLit w) (mkNatLit v)) 388 | else if s.startsWith "#x" then 389 | let s := (s.drop 2).toUpper 390 | let w := 4 * s.length 391 | let f v c := 392 | let d := if c.isDigit then c.toNat - '0'.toNat else c.toNat - 'A'.toNat + 10 393 | v <<< 4 + d 394 | let v := s.foldl f 0 395 | some (mkBitVec w, mkApp2 (.const ``BitVec.ofNat []) (mkNatLit w) (mkNatLit v)) 396 | else 397 | none 398 | | _ => 399 | none 400 | leftAssocOpBool (op : Expr) (as : List Sexp) : ParserM (Expr × Expr) := do 401 | let as ← as.mapM (fun a => return (← parseTerm a).snd) 402 | return (mkBool, as.tail.foldl (mkApp2 op) as.head!) 403 | leftAssocOpBitVec (op : Nat → Expr) (as : List Sexp) : ParserM (Expr × Expr) := do 404 | let (α, a) ← parseTerm as.head! 405 | let op := op (← getBitVecWidth α) 406 | -- Do not reparse a! 407 | let as ← as.tail.mapM (fun a => return (← parseTerm a).snd) 408 | return (α, as.foldl (mkApp2 op) a) 409 | pairwiseDistinct (as : List Sexp) : ParserM (Expr × Expr) := do 410 | if h : as.length < 2 then 411 | throw m!"Error: expected at least two arguments for `distinct`" 412 | else 413 | let (α, as0) ← parseTerm as[0] 414 | let (_, as1) ← parseTerm as[1] 415 | let hα ← if α == mkBool 416 | then pure mkInstBEqBool 417 | else pure (mkInstBEqBitVec (← getBitVecWidth α)) 418 | let mut acc : Expr := mkApp4 (.const ``bne [0]) α hα as0 as1 419 | for hi : i in [2:as.length] do 420 | let (_, asi) ← parseTerm as[i] 421 | acc := mkApp2 (.const ``and []) acc (mkApp4 (.const ``bne [0]) α hα as0 asi) 422 | for hi : i in [1:as.length] do 423 | for hj : j in [i + 1:as.length] do 424 | let (_, asi) ← parseTerm as[i] 425 | let (_, asj) ← parseTerm as[j] 426 | acc := mkApp2 (.const ``and []) acc (mkApp4 (.const ``bne [0]) α hα asi asj) 427 | return (mkBool, acc) 428 | parseNestedBindings (bindings : List (List Sexp)) : ParserM (List (Name × Expr × Expr)) := do 429 | let bindings ← bindings.mapM parseParallelBindings 430 | return bindings.flatten 431 | parseParallelBindings (bindings : List Sexp) : ParserM (List (Name × Expr × Expr)) := do 432 | -- Note: bindings in a parallel let expression use the same context. In 433 | -- particular, variable shadowing only applies after all the bindings are 434 | -- parsed. 435 | let state ← get 436 | let bindings ← bindings.mapM parseBinding 437 | let (level, bvars) := bindings.foldl (fun (lvl, bvs) (n, t, _) => (lvl + 1, bvs.insert n (t, lvl))) (state.level, state.bvars) 438 | set { bvars, level : Parser.State } 439 | return bindings 440 | parseBinding (binding : Sexp) : ParserM (Name × Expr × Expr) := do 441 | match binding with 442 | | sexp!{({n} {v})} => 443 | let n := smtSymbolToName n.serialize 444 | let (t, v) ← parseTerm v 445 | modify fun state => { state with level := state.level + 1 } 446 | return (n, t, v) 447 | | _ => 448 | throw m!"Error: unsupported binding {binding}" 449 | getNestedLetBindingsAndBody (bindings : List (List Sexp)) : Sexp → (List (List Sexp) × Sexp) 450 | | sexp!{(let (...{bs}) {b})} => getNestedLetBindingsAndBody (bs :: bindings) b 451 | | b => (bindings.reverse, b) 452 | 453 | def withTypeDefs (defs : List Sexp) (k : ParserM Expr) : ParserM Expr := do 454 | let state ← get 455 | let defs ← defs.mapM parseTypeDef 456 | let b ← k 457 | set state 458 | -- Note: We set `nonDep` to `false` for user-defined types to ensure 459 | -- type-checking works. Although this could be inefficient, it should be 460 | -- acceptable since user-defined types are rare. 461 | return defs.foldr (fun (n, t, v) b => .letE n t v b false) b 462 | where 463 | parseTypeDef (defn : Sexp) : ParserM (Name × Expr × Expr) := do 464 | match defn with 465 | | sexp!{(define-sort {n} (...{ps}) {b})} => 466 | let state ← get 467 | let ps ← ps.mapM parseParam 468 | let rt := .sort 1 469 | let (cb, ub) ← parseSort b 470 | let n := smtSymbolToName n.serialize 471 | let t := ps.foldr (fun (n, t) b => .forallE n t b .default) rt 472 | let bv := ps.foldr (fun (n, t) b => .lam n t b .default) cb 473 | let nv := ps.foldr (fun (n, t) b => .lam n t b .default) ub 474 | let bvars := state.bvars.insert n (t, state.level) 475 | let uvars := state.uvars.insert n bv 476 | let level := state.level + 1 477 | set { level, bvars, uvars : Parser.State } 478 | return (n, t, nv) 479 | | _ => 480 | throw m!"Error: unsupported type def {defn}" 481 | parseParam (p : Sexp) : ParserM (Name × Expr) := do 482 | match p with 483 | | .atom n => 484 | let n := smtSymbolToName n 485 | let t := .sort 1 486 | modify fun state => { level := state.level + 1, bvars := state.bvars.insert n (t, state.level) } 487 | return (n, t) 488 | | _ => 489 | throw m!"Error: unsupported type param {p}" 490 | 491 | def withFunDecls (decls : List Sexp) (k : ParserM Expr) : ParserM Expr := do 492 | let state ← get 493 | let decls ← decls.mapM parseFunDecl 494 | let b ← k 495 | set state 496 | return decls.foldr (fun (n, t) b => .forallE n t b .default) b 497 | where 498 | parseFunDecl (decl : Sexp) : ParserM (Name × Expr) := do 499 | match decl with 500 | | sexp!{(declare-fun {n} (...{ps}) {s})} => 501 | let state ← get 502 | let ps ← ps.mapM parseParamSort 503 | let (crt, urt) ← parseSort s 504 | let n := smtSymbolToName n.serialize 505 | let (ct, ut) := ps.foldr (fun (ct, ut) (cb, ub) => (.forallE n ct cb .default, .forallE n ut ub .default)) (crt, urt) 506 | let bvars := state.bvars.insert n (ct, state.level) 507 | let level := state.level + 1 508 | set { bvars, level : Parser.State } 509 | return (n, ut) 510 | | _ => 511 | throw m!"Error: unsupported fun decl {decl}" 512 | parseParamSort (s : Sexp) : ParserM (Expr × Expr) := do 513 | let tp ← parseSort s 514 | modify fun state => { state with level := state.level + 1 } 515 | return tp 516 | 517 | def withFunDefs (defs : List Sexp) (k : ParserM Expr) : ParserM Expr := do 518 | let state ← get 519 | -- it's common for SMT-LIB queries to be "letified" using define-fun to 520 | -- minimize their size. We don't recurse on each definition to avoid stack 521 | -- overflows. 522 | let defs ← defs.mapM parseFunDef 523 | let b ← k 524 | set state 525 | return defs.foldr (fun (n, t, v) b => .letE n t v b true) b 526 | where 527 | parseFunDef (defn : Sexp) : ParserM (Name × Expr × Expr) := do 528 | match defn with 529 | | sexp!{(define-fun {n} (...{ps}) {s} {b})} => 530 | let state ← get 531 | let ps ← ps.mapM parseParam 532 | let (crt, urt) ← parseSort s 533 | let (_, b) ← parseTerm b 534 | let n := smtSymbolToName n.serialize 535 | let (ct, ut) := ps.foldr (fun (n, ct, ut) (cb, ub) => (.forallE n ct cb .default, .forallE n ut ub .default)) (crt, urt) 536 | let v := ps.foldr (fun (n, _, t) b => .lam n t b .default) b 537 | let bvars := state.bvars.insert n (ct, state.level) 538 | let level := state.level + 1 539 | set { bvars, level : Parser.State } 540 | return (n, ut, v) 541 | | _ => 542 | throw m!"Error: unsupported fun def {defn}" 543 | parseParam (p : Sexp) : ParserM (Name × Expr × Expr) := do 544 | match p with 545 | | sexp!{({n} {s})} => 546 | let n := smtSymbolToName n.serialize 547 | let (ct, ut) ← parseSort s 548 | modify fun state => { level := state.level + 1, bvars := state.bvars.insert n (ct, state.level) } 549 | return (n, ct, ut) 550 | | _ => 551 | throw m!"Error: unsupported fun param {p}" 552 | 553 | def parseAssert : Sexp → ParserM Expr 554 | | sexp!{(assert {p})} => 555 | return (← parseTerm p).snd 556 | | s => 557 | throw m!"Error: unsupported assert {s}" 558 | 559 | structure Query where 560 | typeDefs : List Sexp := [] 561 | funDecls : List Sexp := [] 562 | funDefs : List Sexp := [] 563 | asserts : List Sexp := [] 564 | 565 | def parseQuery (query : Query) : ParserM Expr := do 566 | withTypeDefs query.typeDefs <| withFunDecls query.funDecls <| withFunDefs query.funDefs do 567 | let conjs ← query.asserts.mapM parseAssert 568 | let p := if h : 0 < conjs.length 569 | then conjs.tail.foldl (mkApp2 (.const ``and [])) conjs[0] 570 | else .const ``true [] 571 | return mkApp3 (.const ``Eq [1]) (.const ``Bool []) (.app (.const ``not []) p) (.const ``true []) 572 | 573 | def filterCmds (sexps : List Sexp) : Query := 574 | go {} sexps 575 | where 576 | go (query : Query) : List Sexp → Query 577 | | sexp!{(define-sort {n} {ps} {b})} :: cmds => 578 | go { query with typeDefs := sexp!{(define-sort {n} {ps} {b})} :: query.typeDefs } cmds 579 | | sexp!{(declare-const {n} {s})} :: cmds => 580 | go { query with funDecls := sexp!{(declare-const {n} {s})} :: query.funDecls } cmds 581 | | sexp!{(declare-fun {n} {ps} {s})} :: cmds => 582 | go { query with funDecls := sexp!{(declare-fun {n} {ps} {s})} :: query.funDecls } cmds 583 | | sexp!{(define-fun {n} {ps} {s} {b})} :: cmds => 584 | go { query with funDefs := sexp!{(define-fun {n} {ps} {s} {b})} :: query.funDefs } cmds 585 | | sexp!{(assert {p})} :: cmds => 586 | go { query with asserts := sexp!{(assert {p})} :: query.asserts } cmds 587 | -- TODO: We should parse `(check-sat)` command. We currently return `sat` if 588 | -- `(check-sat)` command is missing. 589 | | _ :: cmds => 590 | go query cmds 591 | | [] => 592 | { typeDefs := query.typeDefs.reverse 593 | funDecls := query.funDecls.reverse 594 | funDefs := query.funDefs.reverse 595 | asserts := query.asserts.reverse } 596 | 597 | def parseSmt2Query (query : String) : Except MessageData Expr := 598 | match Sexp.Parser.manySexps!.run query with 599 | | Except.error e => 600 | .error s!"{e}" 601 | | Except.ok cmds => 602 | let query := filterCmds cmds 603 | (parseQuery query).run' {} 604 | 605 | end Parser 606 | --------------------------------------------------------------------------------