├── .github └── workflows │ └── ci.yml ├── .gitignore ├── AUTHORS ├── LICENSE ├── README.md ├── Smt.lean ├── Smt ├── Arith.lean ├── Attribute.lean ├── Auto.lean ├── BitVec.lean ├── Bool.lean ├── Builtin.lean ├── Data │ ├── Graph.lean │ └── Sexp.lean ├── Dsl │ └── Sexp.lean ├── Int.lean ├── Nat.lean ├── Options.lean ├── Preprocess.lean ├── Preprocess │ └── Iff.lean ├── Prop.lean ├── Quant.lean ├── Rat.lean ├── Real.lean ├── Recognizers.lean ├── Reconstruct.lean ├── Reconstruct │ ├── Arith.lean │ ├── Arith │ │ ├── ArithMulSign.lean │ │ ├── ArithMulSign │ │ │ ├── Lemmas.lean │ │ │ └── Tactic.lean │ │ ├── ArithPolyNorm.lean │ │ ├── MulPosNeg.lean │ │ ├── MulPosNeg │ │ │ ├── Instances.lean │ │ │ ├── Lemmas.lean │ │ │ └── Tactic.lean │ │ ├── Rewrites.lean │ │ ├── SumBounds.lean │ │ ├── SumBounds │ │ │ ├── Instances.lean │ │ │ ├── Lemmas.lean │ │ │ └── Tactic.lean │ │ ├── TangentPlane.lean │ │ ├── TangentPlane │ │ │ └── Theorems.lean │ │ ├── TightBounds.lean │ │ ├── TightBounds │ │ │ ├── Lemmas.lean │ │ │ └── Tactic.lean │ │ ├── TransFns.lean │ │ ├── TransFns │ │ │ ├── ArithTransExpApproxAboveNeg.lean │ │ │ ├── ArithTransExpApproxAbovePos.lean │ │ │ ├── ArithTransExpApproxBelow.lean │ │ │ ├── ArithTransExpNeg.lean │ │ │ ├── ArithTransExpPositivity.lean │ │ │ ├── ArithTransExpSuperLin.lean │ │ │ ├── ArithTransExpZero.lean │ │ │ ├── ArithTransPi.lean │ │ │ ├── ArithTransSineApproxAboveNeg.lean │ │ │ ├── ArithTransSineApproxAbovePos.lean │ │ │ ├── ArithTransSineApproxBelowNeg.lean │ │ │ ├── ArithTransSineApproxBelowPos.lean │ │ │ ├── ArithTransSineBounds.lean │ │ │ ├── ArithTransSineShift.lean │ │ │ ├── ArithTransSineSymmetry.lean │ │ │ ├── ArithTransSineTangentPi.lean │ │ │ ├── ArithTransSineTangentZero.lean │ │ │ └── Utils.lean │ │ ├── Trichotomy.lean │ │ └── Trichotomy │ │ │ ├── Lemmas.lean │ │ │ └── Tactic.lean │ ├── BitVec.lean │ ├── BitVec │ │ └── Bitblast.lean │ ├── Bool │ │ ├── Lemmas.lean │ │ └── Tactic.lean │ ├── Builtin.lean │ ├── Builtin │ │ ├── AC.lean │ │ ├── Absorb.lean │ │ ├── Lemmas.lean │ │ ├── Rewrites.lean │ │ └── Tactic.lean │ ├── Certified │ │ ├── Coe.lean │ │ ├── Defs.lean │ │ ├── Env.lean │ │ ├── EufRules.lean │ │ ├── ModusPonens.lean │ │ ├── Rules.lean │ │ ├── Term.lean │ │ ├── Types.lean │ │ └── Utils.lean │ ├── Int.lean │ ├── Int │ │ ├── Core.lean │ │ ├── Lemmas.lean │ │ ├── Polynorm.lean │ │ └── Rewrites.lean │ ├── Prop.lean │ ├── Prop │ │ ├── Core.lean │ │ ├── Factor.lean │ │ ├── Lemmas.lean │ │ ├── LiftOrNToImp.lean │ │ ├── LiftOrNToNeg.lean │ │ ├── PermutateOr.lean │ │ ├── Pull.lean │ │ ├── Resolution.lean │ │ └── Rewrites.lean │ ├── Quant.lean │ ├── Quant │ │ └── Lemmas.lean │ ├── Rat.lean │ ├── Rat │ │ ├── Core.lean │ │ ├── Lemmas.lean │ │ ├── Polynorm.lean │ │ └── Rewrites.lean │ ├── Real.lean │ ├── Real │ │ ├── Core.lean │ │ ├── Lemmas.lean │ │ ├── Polynorm.lean │ │ └── Rewrites.lean │ ├── Rewrite.lean │ ├── Timed.lean │ ├── UF.lean │ ├── UF │ │ ├── Congruence.lean │ │ └── Rewrites.lean │ └── Util.lean ├── String.lean ├── Tactic.lean ├── Tactic │ ├── Concretize.lean │ ├── EqnDef.lean │ ├── Smt.lean │ ├── WHNFConfigurable.lean │ ├── WHNFConfigurableRef.lean │ └── WHNFSmt.lean ├── Translate.lean ├── Translate │ ├── BitVec.lean │ ├── Bool.lean │ ├── Commands.lean │ ├── Int.lean │ ├── Nat.lean │ ├── Prop.lean │ ├── Query.lean │ ├── Real.lean │ ├── Solver.lean │ ├── String.lean │ └── Term.lean ├── UF.lean └── Util.lean ├── Test ├── Auto.expected ├── Auto.lean ├── BitVec │ ├── Add.expected │ ├── Add.lean │ ├── Shift.expected │ ├── Shift.lean │ ├── XorComm.expected │ └── XorComm.lean ├── Bool │ ├── Addition.expected │ ├── Addition.lean │ ├── Assoc.expected │ ├── Assoc.lean │ ├── Comm.expected │ ├── Comm.lean │ ├── Cong.expected │ ├── Cong.lean │ ├── Conjunction.expected │ ├── Conjunction.lean │ ├── DisjunctiveSyllogism.expected │ ├── DisjunctiveSyllogism.lean │ ├── Exists'.expected │ ├── Exists'.lean │ ├── Falsum.expected │ ├── Falsum.lean │ ├── HypotheticalSyllogism.expected │ ├── HypotheticalSyllogism.lean │ ├── ModusPonens'.expected │ ├── ModusPonens'.lean │ ├── ModusPonens.expected │ ├── ModusPonens.lean │ ├── ModusTollens.expected │ ├── ModusTollens.lean │ ├── PropExt.expected │ ├── PropExt.lean │ ├── Refl.expected │ ├── Refl.lean │ ├── Resolution.expected │ ├── Resolution.lean │ ├── Simplification.expected │ ├── Simplification.lean │ ├── Symm.expected │ ├── Symm.lean │ ├── Trans.expected │ ├── Trans.lean │ ├── Triv'''.expected │ ├── Triv''.expected │ ├── Triv''.lean │ ├── Triv'.expected │ ├── Triv'.lean │ ├── Triv.expected │ ├── Triv.lean │ ├── Verum.expected │ └── Verum.lean ├── Concretize │ ├── Add.expected │ ├── Add.lean │ ├── BitVec.expected │ └── BitVec.lean ├── EUF │ ├── Example.expected │ ├── Example.lean │ ├── Issue100.expected │ ├── Issue100.lean │ ├── Issue126.expected │ ├── Issue126.lean │ ├── Issue97.expected │ └── Issue97.lean ├── EqnDef │ ├── Extract.expected │ └── Extract.lean ├── Examples │ ├── Demo.expected │ └── Demo.lean ├── Int │ ├── Arith.expected │ ├── Arith.lean │ ├── Binders.expected │ ├── Binders.lean │ ├── Cong.expected │ ├── Cong.lean │ ├── DefineSort.expected │ ├── DefineSort.lean │ ├── Dvd.expected │ ├── Dvd.lean │ ├── DvdTimeout.expected │ ├── DvdTimeout.lean │ ├── ForallExists.expected │ ├── ForallExists.lean │ ├── Let.expected │ ├── Let.lean │ ├── Neg.expected │ ├── Neg.lean │ ├── linarith.expected │ └── linarith.lean ├── Nat │ ├── Cong.expected │ ├── Cong.lean │ ├── Exists''.expected │ ├── Exists''.lean │ ├── Exists'.expected │ ├── Exists'.lean │ ├── Forall'.expected │ ├── Forall'.lean │ ├── Max.expected │ ├── Max.lean │ ├── NeqZero.expected │ ├── NeqZero.lean │ ├── Sum'.expected │ ├── Sum'.lean │ ├── Triv'.expected │ ├── Triv'.lean │ ├── Triv.expected │ ├── Triv.lean │ ├── ZeroSub'.expected │ ├── ZeroSub'.lean │ ├── ZeroSub.expected │ └── ZeroSub.lean ├── Playground │ └── WHNFExamples.lean ├── Prop │ ├── Addition.expected │ ├── Addition.lean │ ├── Assoc.expected │ ├── Assoc.lean │ ├── Comm.expected │ ├── Comm.lean │ ├── Cong.expected │ ├── Cong.lean │ ├── Conjunction.expected │ ├── Conjunction.lean │ ├── DisjunctiveSyllogism.expected │ ├── DisjunctiveSyllogism.lean │ ├── Exists'.expected │ ├── Exists'.lean │ ├── Falsum.expected │ ├── Falsum.lean │ ├── HypotheticalSyllogism.expected │ ├── HypotheticalSyllogism.lean │ ├── Iff.expected │ ├── Iff.lean │ ├── ModusPonens'.expected │ ├── ModusPonens'.lean │ ├── ModusPonens.expected │ ├── ModusPonens.lean │ ├── ModusTollens.expected │ ├── ModusTollens.lean │ ├── PropExt.expected │ ├── PropExt.lean │ ├── Refl.expected │ ├── Refl.lean │ ├── Resolution.expected │ ├── Resolution.lean │ ├── Simplification.expected │ ├── Simplification.lean │ ├── Symm.expected │ ├── Symm.lean │ ├── Trans.expected │ ├── Trans.lean │ ├── Triv''.expected │ ├── Triv''.lean │ ├── Triv'.expected │ ├── Triv'.lean │ ├── Triv.expected │ ├── Triv.lean │ ├── Verum.expected │ └── Verum.lean ├── Quant.expected ├── Quant.lean ├── Rat │ ├── linarith.expected │ └── linarith.lean ├── Real │ ├── Density.expected │ ├── Density.lean │ ├── Inverse.expected │ └── Inverse.lean ├── Reconstruction │ ├── AndElim.expected │ ├── AndElim.lean │ ├── Arith │ │ ├── ArithMulSign.expected │ │ ├── ArithMulSign.lean │ │ ├── MulPosNeg.expected │ │ ├── MulPosNeg.lean │ │ ├── SumBounds.expected │ │ ├── SumBounds.lean │ │ ├── TightBounds.expected │ │ ├── TightBounds.lean │ │ ├── Trichotomy.expected │ │ └── Trichotomy.lean │ ├── Factor.expected │ ├── Factor.lean │ ├── LiftOrNToImp.expected │ ├── LiftOrNToImp.lean │ ├── LiftOrNToNeg.expected │ ├── LiftOrNToNeg.lean │ ├── NotOrElim.expected │ ├── NotOrElim.lean │ ├── PermutateOr.expected │ ├── PermutateOr.lean │ ├── Pull.expected │ ├── Pull.lean │ ├── Resolution.expected │ └── Resolution.lean ├── Solver │ ├── Error.expected │ ├── Error.lean │ ├── Interactive.expected │ ├── Interactive.lean │ ├── Model.expected │ ├── Model.lean │ ├── Proof.expected │ ├── Proof.lean │ ├── Triv.expected │ └── Triv.lean ├── String │ ├── Append.expected │ ├── Append.lean │ ├── Contains.expected │ ├── Contains.lean │ ├── Empty.expected │ ├── Empty.lean │ ├── GetOp.expected │ ├── GetOp.lean │ ├── Length.expected │ ├── Length.lean │ ├── Lt.expected │ ├── Lt.lean │ ├── Replace.expected │ └── Replace.lean ├── Unit │ ├── Int.expected │ ├── Int.lean │ ├── Quant.expected │ ├── Quant.lean │ ├── Rat.expected │ ├── Rat.lean │ ├── Real.expected │ └── Real.lean ├── linarith.expected └── linarith.lean ├── lake-manifest.json ├── lakefile.lean └── lean-toolchain /.github/workflows/ci.yml: -------------------------------------------------------------------------------- 1 | name: CI 2 | 3 | on: [push, pull_request, workflow_dispatch] 4 | 5 | jobs: 6 | build: 7 | name: ${{ matrix.name }} 8 | runs-on: ${{ matrix.os }} 9 | strategy: 10 | matrix: 11 | include: 12 | - name: Linux-aarch64 13 | os: ubuntu-22.04-arm 14 | - name: Linux-x86_64 15 | os: ubuntu-22.04 16 | - name: macOS-aarch64 17 | os: macos-14 18 | - name: macOS-x86_64 19 | os: macos-13 20 | - name: Windows-x86_64 21 | os: windows-latest 22 | steps: 23 | - name: Checkout 24 | uses: actions/checkout@v4 25 | - name: Set up Lean 26 | uses: leanprover/lean-action@v1 27 | with: 28 | auto-config: false 29 | use-mathlib-cache: false 30 | - name: Build and Test Lean-SMT 31 | run: | 32 | lake update 33 | lake build Smt Smt.Auto Smt.Rat Smt.Real 34 | lake run test 35 | shell: bash 36 | -------------------------------------------------------------------------------- /.gitignore: -------------------------------------------------------------------------------- 1 | /.lake 2 | -------------------------------------------------------------------------------- /AUTHORS: -------------------------------------------------------------------------------- 1 | The authors, designers, and main contributors to lean-smt are listed below. 2 | lean-smt's copyright is held by these individuals and the affiliated 3 | institutions at the time of their contributions (note that some authors have 4 | had more than one affiliated institution). See the file LICENSE for details on 5 | the copyright and licensing of lean-smt. 6 | 7 | The developers and authors of lean-smt are: 8 | 9 | Current: 10 | Haniel Barbosa, Universidade Federal de Minas Gerais 11 | Tomaz Mascarenhas, Universidade Federal de Minas Gerais 12 | Abdalrhman Mohamed, The University of Iowa 13 | Cesare Tinelli, The University of Iowa 14 | 15 | The developers who contributed to lean-smt are: 16 | Wojciech Nawrocki, SRI International, Carnegie Mellon University 17 | -------------------------------------------------------------------------------- /README.md: -------------------------------------------------------------------------------- 1 | # Lean-SMT 2 | 3 | This project provides Lean tactics to discharge goals into SMT solvers. 4 | It is under active development and is currently in a beta phase. While it is 5 | usable, it is important to note that there are still some rough edges and 6 | ongoing improvements being made. 7 | 8 | ## Supported Theories 9 | `lean-smt` currently supports the theories of Uninterpreted Functions and Linear 10 | Integer/Real Arithmetic with quantifiers. Mathlib is required for Real 11 | Arithmetic. Support for the theory of Bitvectors is at an experimental stage. 12 | Support for additional theories is in progress. 13 | 14 | ## Requirements 15 | `lean-smt` depends on [`lean-cvc5`](https://github.com/abdoo8080/lean-cvc5), 16 | which currently only supports Linux and macOS. 17 | 18 | ## Setup 19 | To use `lean-smt` in your project, add the following lines to your list of 20 | dependencies in `lakefile.toml`: 21 | ```toml 22 | [[require]] 23 | name = "smt" 24 | scope = "ufmg-smite" 25 | rev = "main" 26 | ``` 27 | If your build configuration is in `lakefile.lean`, add the following line to 28 | your dependencies: 29 | ```lean 30 | require smt from git "https://github.com/ufmg-smite/lean-smt.git" @ "main" 31 | ``` 32 | 33 | ## Usage 34 | `lean-smt` comes with one main tactic, `smt`, that translates the current goal 35 | into an SMT query, sends the query to cvc5, and (if the solver returns `unsat`) 36 | replays cvc5's proof in Lean. cvc5's proofs may contain holes, returned as Lean 37 | goals. You can fill these holes manually or with other tactics. To use the `smt` 38 | tactic, you just need to import the `Smt` library: 39 | ```lean 40 | import Smt 41 | 42 | example [Nonempty U] {f : U → U → U} {a b c d : U} 43 | (h0 : a = b) (h1 : c = d) (h2 : p1 ∧ True) (h3 : (¬ p1) ∨ (p2 ∧ p3)) 44 | (h4 : (¬ p3) ∨ (¬ (f a c = f b d))) : False := by 45 | smt [h0, h1, h2, h3, h4] 46 | ``` 47 | To use the `smt` tactic on Real arithmetic goals, import `Smt.Real`: 48 | ```lean 49 | import Smt 50 | import Smt.Real 51 | 52 | example (ε : Real) (h1 : ε > 0) : ε / 2 + ε / 3 + ε / 7 < ε := by 53 | smt [h1] 54 | ``` 55 | `lean-smt` integrates with 56 | [`lean-auto`](https://github.com/leanprover-community/lean-auto) to provide 57 | basic hammer-like capabilities. To set the `smt` tactic as a backend for `auto`, 58 | import `Smt.Auto` and set `auto.native` to `true`: 59 | ```lean 60 | import Mathlib.Algebra.Group.Defs 61 | import Smt 62 | import Smt.Auto 63 | 64 | set_option auto.native true 65 | 66 | variable [Group G] 67 | 68 | theorem inverse : ∀ (a : G), a * a⁻¹ = 1 := by 69 | auto [mul_assoc, one_mul, inv_mul_cancel] 70 | 71 | theorem identity : ∀ (a : G), a * 1 = a := by 72 | auto [mul_assoc, one_mul, inv_mul_cancel, inverse] 73 | 74 | theorem unique_identity : ∀ (e : G), (∀ a, e * a = a) ↔ e = 1 := by 75 | auto [mul_assoc, one_mul, inv_mul_cancel] 76 | ``` 77 | -------------------------------------------------------------------------------- /Smt.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: Abdalrhman Mohamed 6 | -/ 7 | 8 | import Smt.BitVec 9 | import Smt.Bool 10 | import Smt.Builtin 11 | import Smt.Int 12 | import Smt.Nat 13 | import Smt.Options 14 | import Smt.Prop 15 | import Smt.Quant 16 | import Smt.String 17 | import Smt.Tactic 18 | import Smt.UF 19 | 20 | import Smt.Translate.Solver 21 | -------------------------------------------------------------------------------- /Smt/Arith.lean: -------------------------------------------------------------------------------- 1 | /- 2 | Copyright (c) 2021-2023 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: Tomaz Gomes Mascarenhas 6 | -/ 7 | 8 | import Smt.Reconstruct.Arith 9 | import Smt.Translate.Int 10 | import Smt.Translate.Real 11 | -------------------------------------------------------------------------------- /Smt/Attribute.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: Abdalrhman Mohamed 6 | -/ 7 | 8 | import Lean 9 | 10 | namespace Smt.Attribute 11 | 12 | open Lean 13 | 14 | /-- An extension to Lean's runtime environment to support SMT attributes. 15 | Maintains a set of function declarations for the `smt` tactic to utilize 16 | while generating the SMT query. -/ 17 | abbrev SmtExtension := SimpleScopedEnvExtension (Name × Name) (Std.HashMap Name (Std.HashSet Name)) 18 | 19 | /-- Adds a declaration to the set of function declarations maintained by the SMT 20 | environment extension. -/ 21 | def addSmtEntry (d : Std.HashMap Name (Std.HashSet Name)) (e : (Name × Name)) := 22 | d.insert e.fst ((d.getD e.fst {}).insert e.snd) 23 | 24 | initialize smtExt : SmtExtension ← registerSimpleScopedEnvExtension { 25 | name := `SmtExt 26 | initial := {} 27 | addEntry := addSmtEntry 28 | } 29 | 30 | /-- Throws unexpected type error. -/ 31 | def throwUnexpectedType (t : Name) (n : Name) : AttrM Unit := 32 | throwError s!"unexpected type at '{n}', `{t}` expected" 33 | 34 | /-- Validates the tagged declaration. -/ 35 | def validate (n : Name) (t : Name) : AttrM Unit := do 36 | match (← getEnv).find? n with 37 | | none => throwError s!"unknown constant '{n}'" 38 | | some info => 39 | match info.type with 40 | | Expr.const c .. => if c != t then throwUnexpectedType t n 41 | | _ => throwUnexpectedType t n 42 | 43 | /-- Registers an SMT attribute with the provided name and description and links 44 | it against `ext`. -/ 45 | def registerSmtAttr (attrName : Name) (typeName : Name) (attrDescr : String) 46 | : IO Unit := 47 | registerBuiltinAttribute { 48 | name := attrName 49 | descr := attrDescr 50 | applicationTime := AttributeApplicationTime.afterTypeChecking 51 | add := fun decl stx attrKind => do 52 | trace[smt.attr] s!"attrName: {attrName}, attrDescr: {attrDescr}" 53 | trace[smt.attr] s!"decl: {decl}, stx: {stx}, attrKind: {attrKind}" 54 | Attribute.Builtin.ensureNoArgs stx 55 | validate decl typeName 56 | setEnv (smtExt.addEntry (← getEnv) (typeName, decl)) 57 | erase := fun declName => do 58 | let s := smtExt.getState (← getEnv) 59 | let s := s.erase declName 60 | modifyEnv fun env => smtExt.modifyState env fun _ => s 61 | } 62 | 63 | initialize registerSmtAttr `smt_translate `Smt.Translator 64 | "Utilize this function to translate Lean expressions to SMT terms." 65 | 66 | initialize registerSmtAttr `smt_sort_reconstruct `Smt.SortReconstructor 67 | "Utilize this function to translate cvc5 sorts to Lean expressions." 68 | 69 | initialize registerSmtAttr `smt_term_reconstruct `Smt.TermReconstructor 70 | "Utilize this function to translate cvc5 terms to Lean expressions." 71 | 72 | initialize registerSmtAttr `smt_proof_reconstruct `Smt.ProofReconstructor 73 | "Utilize this function to translate cvc5 proofs to Lean expressions." 74 | 75 | end Smt.Attribute 76 | -------------------------------------------------------------------------------- /Smt/Auto.lean: -------------------------------------------------------------------------------- 1 | /- 2 | Copyright (c) 2021-2024 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: Abdalrhman Mohamed 6 | -/ 7 | 8 | import Auto.Tactic 9 | import Smt.Tactic.Smt 10 | 11 | open Lean in 12 | def Smt.smtSolverFunc (ls : Array Auto.Lemma) (is : Array Auto.Lemma) : MetaM Expr := do 13 | let fi l := do 14 | let userName ← mkFreshUserName `inst 15 | let type ← Meta.mkAppM ``Inhabited #[l.type] 16 | let value ← Meta.mkAppOptM ``Inhabited.mk #[l.type, l.proof] 17 | return { userName, type, value } 18 | let is ← is.mapM fi 19 | let fl l := do 20 | let userName ← mkFreshUserName `h 21 | let type ← Auto.Lam2D.approxSimpNF l.type 22 | let value := l.proof 23 | return { userName, type, value } 24 | let ls ← ls.mapM fl 25 | let h ← Meta.mkFreshExprMVar (Expr.const ``False []) 26 | let (_, mv) ← h.mvarId!.assertHypotheses is 27 | let (fvs, mv) ← mv.assertHypotheses ls 28 | mv.withContext do 29 | let hs := fvs.map (.fvar ·) 30 | let timeout := (← getOptions).get? ``auto.smt.timeout 31 | _ ← smt { timeout } mv hs.toList 32 | -- Note: auto should allow solvers to export new goals to users 33 | -- for mv in mvs do 34 | -- logInfo m!"new : {mv}" 35 | instantiateMVars h 36 | -------------------------------------------------------------------------------- /Smt/BitVec.lean: -------------------------------------------------------------------------------- 1 | /- 2 | Copyright (c) 2021-2023 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: Abdalrhman Mohamed 6 | -/ 7 | 8 | import Smt.Reconstruct.BitVec 9 | import Smt.Translate.BitVec 10 | -------------------------------------------------------------------------------- /Smt/Bool.lean: -------------------------------------------------------------------------------- 1 | /- 2 | Copyright (c) 2021-2023 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: Abdalrhman Mohamed 6 | -/ 7 | 8 | import Smt.Translate.Bool 9 | -------------------------------------------------------------------------------- /Smt/Builtin.lean: -------------------------------------------------------------------------------- 1 | /- 2 | Copyright (c) 2021-2023 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: Abdalrhman Mohamed 6 | -/ 7 | 8 | import Smt.Reconstruct.Builtin 9 | -------------------------------------------------------------------------------- /Smt/Data/Graph.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: Abdalrhman Mohamed, Wojciech Nawrocki 6 | -/ 7 | 8 | import Std.Data.HashMap 9 | import Std.Data.HashSet 10 | import Lean.Message 11 | 12 | open Lean 13 | 14 | abbrev Graph (α) (β) [BEq α] [Hashable α] := Std.HashMap α (Std.HashMap α β) 15 | 16 | namespace Graph 17 | 18 | variable {α} {β} [BEq α] [Hashable α] (g : Graph α β) (v u : α) (e : β) 19 | 20 | def empty : Graph α β := {} 21 | 22 | def vertices : List α := g.fold (fun a v _ => v :: a) [] 23 | 24 | def neighbors? : Option (List α) := 25 | g[v]? >>= fun es => some (es.fold (fun a v _ => v :: a) []) 26 | 27 | def neighbors! : List α := match (g.neighbors? v) with 28 | | some ns => ns 29 | | none => panic! "vertex is not in the graph" 30 | 31 | def addVertex : Graph α β := g.insert v {} 32 | 33 | def addEdge : Graph α β := g.insert v ((g[v]!).insert u e) 34 | 35 | def weight? : Option β := g[v]? >>= fun es => es[u]? 36 | 37 | partial def dfs [Monad m] (f : α → m Unit) : m Unit := 38 | StateT.run' (s := {}) do 39 | for v in g.vertices do 40 | visitVertex v 41 | where 42 | visitVertex (v : α) : StateT (Std.HashSet α) m Unit := do 43 | let vs ← get 44 | if vs.contains v then 45 | return 46 | set (vs.insert v) 47 | for u in g.neighbors! v do 48 | visitVertex u 49 | f v 50 | 51 | partial def orderedDfs [Monad m] (vs : List α) (f : α → m Unit) : m Unit := 52 | StateT.run' (s := {}) do 53 | for v in vs do 54 | visitVertex v 55 | where 56 | visitVertex (v : α) : StateT (Std.HashSet α) m Unit := do 57 | let vs ← get 58 | if vs.contains v then 59 | return 60 | set (vs.insert v) 61 | for u in g.neighbors! v do 62 | visitVertex u 63 | f v 64 | 65 | open Std.Format in 66 | protected def format [ToFormat α] [ToFormat β] : Format := 67 | bracket "{" (joinSep (g.vertices.map formatVertex) ("," ++ line)) "}" 68 | where 69 | formatVertex (v : α) : Format := 70 | f!"{v} ↦ {formatNeighbors (g.neighbors! v)}" 71 | formatNeighbors (ns : List α) : Format := 72 | bracket "{" (joinSep ns ("," ++ line)) "}" 73 | 74 | instance [ToFormat α] [ToFormat β] : ToFormat (Graph α β) where 75 | format g := Graph.format g 76 | 77 | open Lean MessageData in 78 | protected def toMessageData [ToMessageData α] [ToMessageData β] : MessageData := 79 | bracket "{" (joinSep (g.vertices.map formatVertex) ("," ++ Format.line)) "}" 80 | where 81 | formatVertex (v : α) : MessageData := 82 | m!"{v} ↦ {formatNeighbors (g.neighbors! v)}" 83 | formatNeighbors (ns : List α) : MessageData := 84 | bracket "{" (joinSep (ns.map toMessageData) ("," ++ Format.line)) "}" 85 | 86 | open Lean in 87 | instance [ToMessageData α] [ToMessageData β] : ToMessageData (Graph α β) where 88 | toMessageData g := Graph.toMessageData g 89 | 90 | end Graph 91 | -------------------------------------------------------------------------------- /Smt/Int.lean: -------------------------------------------------------------------------------- 1 | /- 2 | Copyright (c) 2021-2023 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: Abdalrhman Mohamed 6 | -/ 7 | 8 | import Smt.Reconstruct.Int 9 | import Smt.Translate.Int 10 | -------------------------------------------------------------------------------- /Smt/Nat.lean: -------------------------------------------------------------------------------- 1 | /- 2 | Copyright (c) 2021-2023 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: Abdalrhman Mohamed 6 | -/ 7 | 8 | import Smt.Translate.Nat 9 | -------------------------------------------------------------------------------- /Smt/Options.lean: -------------------------------------------------------------------------------- 1 | /- 2 | Copyright (c) 2021-2024 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: Abdalrhman Mohamed, Tomaz Gomes Mascarenhas 6 | -/ 7 | 8 | import Lean 9 | 10 | namespace Smt 11 | 12 | open Lean 13 | 14 | initialize 15 | registerTraceClass `smt 16 | registerTraceClass `smt.attr 17 | registerTraceClass `smt.translate 18 | registerTraceClass `smt.translate.expr 19 | registerTraceClass `smt.translate.query 20 | registerTraceClass `smt.solve 21 | registerTraceClass `smt.reconstruct 22 | registerTraceClass `smt.reconstruct.sort 23 | registerTraceClass `smt.reconstruct.term 24 | registerTraceClass `smt.reconstruct.proof 25 | 26 | end Smt 27 | -------------------------------------------------------------------------------- /Smt/Preprocess.lean: -------------------------------------------------------------------------------- 1 | /- 2 | Copyright (c) 2021-2024 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: Abdalrhman Mohamed, Tomaz Gomes Mascarenhas 6 | -/ 7 | 8 | import Smt.Preprocess.Iff 9 | -------------------------------------------------------------------------------- /Smt/Preprocess/Iff.lean: -------------------------------------------------------------------------------- 1 | /- 2 | Copyright (c) 2021-2024 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: Abdalrhman Mohamed, Tomaz Gomes Mascarenhas 6 | -/ 7 | 8 | import Lean 9 | import Qq 10 | 11 | namespace Smt.Preprocess 12 | 13 | open Lean Qq 14 | 15 | theorem iff_eq_eq : (p ↔ q) = (p = q) := propext ⟨propext, (· ▸ ⟨(·), (·)⟩)⟩ 16 | 17 | theorem eq_resolve {p q : Prop} (hp : p) (hpq : p = q) : q := hpq ▸ hp 18 | 19 | def replaceIff (e : Expr) : MetaM Expr := 20 | let f e := 21 | if let some ((l : Q(Prop)), (r : Q(Prop))) := e.app2? ``Iff then 22 | q($l = $r) 23 | else 24 | none 25 | Meta.mkAppM ``Eq #[e, e.replace f] 26 | 27 | def containsIff (e : Expr) : Bool := 28 | (Expr.const ``Iff []).occurs e 29 | 30 | def elimIff (mv : MVarId) (hs : List Expr) : MetaM (List Expr × MVarId) := mv.withContext do 31 | let t ← instantiateMVars (← mv.getType) 32 | let ts ← hs.mapM (Meta.inferType · >>= instantiateMVars) 33 | if !(containsIff t || ts.any containsIff) then 34 | return (hs, mv) 35 | let simpTheorems ← #[``eq_self, ``iff_eq_eq].foldlM (·.addConst ·) ({} : Meta.SimpTheorems) 36 | let simpTheorems := #[simpTheorems] 37 | let congrTheorems ← Meta.getSimpCongrTheorems 38 | let ctx ← Meta.Simp.mkContext {} simpTheorems congrTheorems 39 | let (hs, mv) ← elimIffLocalDecls mv hs ctx 40 | let mv ← elimIffTarget mv ctx 41 | return (hs, mv) 42 | where 43 | elimIffLocalDecls mv hs ctx := mv.withContext do 44 | let mut newHs := [] 45 | let mut toAssert := #[] 46 | for h in hs do 47 | let type ← Meta.inferType h 48 | let eq ← replaceIff (← instantiateMVars type) 49 | let (_, l, r) := eq.eq?.get! 50 | if l == r then 51 | newHs := h :: newHs 52 | else 53 | let userName ← if h.isFVar then h.fvarId!.getUserName else Lean.mkFreshId 54 | let type := r 55 | let (r, _) ← Meta.simp eq ctx 56 | let value ← Meta.mkAppM ``eq_resolve #[h, ← Meta.mkOfEqTrue (← r.getProof)] 57 | toAssert := toAssert.push { userName, type, value } 58 | let (fvs, mv) ← mv.assertHypotheses toAssert 59 | newHs := newHs.reverse ++ (fvs.map (.fvar ·)).toList 60 | return (newHs, mv) 61 | elimIffTarget mv ctx := mv.withContext do 62 | let eq ← replaceIff (← instantiateMVars (← mv.getType)) 63 | let (r, _) ← Meta.simp eq ctx 64 | if r.expr.isTrue then 65 | mv.replaceTargetEq eq.appArg! (← Meta.mkOfEqTrue (← r.getProof)) 66 | else 67 | return mv 68 | 69 | end Smt.Preprocess 70 | -------------------------------------------------------------------------------- /Smt/Prop.lean: -------------------------------------------------------------------------------- 1 | /- 2 | Copyright (c) 2021-2023 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: Abdalrhman Mohamed 6 | -/ 7 | 8 | import Smt.Reconstruct.Prop 9 | import Smt.Translate.Prop 10 | -------------------------------------------------------------------------------- /Smt/Quant.lean: -------------------------------------------------------------------------------- 1 | /- 2 | Copyright (c) 2021-2023 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: Abdalrhman Mohamed 6 | -/ 7 | 8 | import Smt.Reconstruct.Quant 9 | -------------------------------------------------------------------------------- /Smt/Rat.lean: -------------------------------------------------------------------------------- 1 | /- 2 | Copyright (c) 2021-2023 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: Abdalrhman Mohamed 6 | -/ 7 | 8 | import Smt.Reconstruct.Rat 9 | -------------------------------------------------------------------------------- /Smt/Real.lean: -------------------------------------------------------------------------------- 1 | /- 2 | Copyright (c) 2021-2023 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: Abdalrhman Mohamed 6 | -/ 7 | 8 | import Smt.Reconstruct.Real 9 | import Smt.Translate.Real 10 | -------------------------------------------------------------------------------- /Smt/Reconstruct/Arith/ArithMulSign.lean: -------------------------------------------------------------------------------- 1 | /- 2 | Copyright (c) 2021-2023 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: Tomaz Gomes Mascarenhas 6 | -/ 7 | 8 | -- implementation of: 9 | -- https://cvc5.github.io/docs/cvc5-1.0.2/proofs/proof_rules.html#_CPPv4N4cvc58internal6PfRule15ARITH_MULT_SIGNE 10 | 11 | import Smt.Reconstruct.Arith.ArithMulSign.Lemmas 12 | import Smt.Reconstruct.Arith.ArithMulSign.Tactic 13 | -------------------------------------------------------------------------------- /Smt/Reconstruct/Arith/ArithMulSign/Lemmas.lean: -------------------------------------------------------------------------------- 1 | /- 2 | Copyright (c) 2021-2023 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: Tomaz Gomes Mascarenhas 6 | -/ 7 | 8 | import Mathlib.Algebra.Order.Ring.Defs 9 | import Mathlib.Data.Real.Basic 10 | 11 | import Smt.Reconstruct.Arith.MulPosNeg.Lemmas 12 | 13 | namespace Smt.Reconstruct.Arith 14 | 15 | variable {α : Type} 16 | 17 | variable [LinearOrderedCommRing α] 18 | 19 | variable {a b : α} 20 | 21 | mutual 22 | 23 | theorem powNegEven : ∀ {k : Nat} {z : α}, z < 0 → Even k → z ^ k > 0 := by 24 | intros k z h₁ h₂ 25 | cases k with 26 | | zero => simp 27 | | succ k' => 28 | have k'NotEven := Nat.even_add_one.mp h₂ 29 | have k'Odd := Nat.not_even_iff_odd.mp k'NotEven 30 | have rc := powNegOdd h₁ k'Odd 31 | simp [Pow.pow] 32 | have mulZ := arith_mul_neg_lt ⟨h₁, rc⟩ 33 | simp at mulZ 34 | rw [pow_succ, mul_comm] 35 | exact mulZ 36 | 37 | theorem powNegOdd : ∀ {k : Nat} {z : α}, z < 0 → Odd k → z ^ k < 0 := by 38 | intros k z h₁ h₂ 39 | cases k with 40 | | zero => simp at h₂ 41 | | succ k' => 42 | simp only [Odd, add_left_inj] at h₂ 43 | have k'Even : Even k' := (even_iff_exists_two_nsmul k').mpr h₂ 44 | have rc := powNegEven h₁ k'Even 45 | simp only [gt_iff_lt] 46 | have mulZ := arith_mul_neg_lt ⟨h₁, rc⟩ 47 | simp only [mul_zero, gt_iff_lt] at mulZ 48 | rw [pow_succ, mul_comm] 49 | exact mul_neg_of_neg_of_pos h₁ rc 50 | 51 | end 52 | 53 | theorem powPos : ∀ {k : Nat} {z : α}, z > 0 → z ^ k > 0 := by 54 | intro k z h 55 | cases k with 56 | | zero => simp 57 | | succ k' => 58 | have ih := @powPos k' z h 59 | rw [pow_succ] 60 | have h₂ := arith_mul_pos_lt ⟨h, ih⟩ 61 | simp at h₂ 62 | rw [mul_comm] 63 | exact h₂ 64 | 65 | theorem nonZeroEvenPow : ∀ {k : Nat} {z : α}, 66 | z ≠ 0 → Even k → z ^ k > 0 := by 67 | intros k z h₁ h₂ 68 | match lt_trichotomy z 0 with 69 | | Or.inl zNeg => exact powNegEven zNeg h₂ 70 | | Or.inr (Or.inl zZero) => rw [zZero] at h₁; simp at h₁ 71 | | Or.inr (Or.inr zPos) => exact powPos zPos 72 | 73 | theorem combineSigns₁ : a > 0 → b > 0 → b * a > 0 := by 74 | intros h₁ h₂ 75 | have h := mul_lt_mul_of_pos_left h₁ h₂ 76 | simp at h 77 | exact h 78 | 79 | theorem combineSigns₂ : a > 0 → b < 0 → b * a < 0 := by 80 | intros h₁ h₂ 81 | have h := mul_lt_mul_of_pos_right h₂ h₁ 82 | simp at h 83 | exact h 84 | 85 | theorem combineSigns₃ : a < 0 → b > 0 → b * a < 0 := by 86 | intros h₁ h₂ 87 | have h := mul_lt_mul_of_pos_left h₁ h₂ 88 | simp at h 89 | exact h 90 | 91 | theorem combineSigns₄ : a < 0 → b < 0 → b * a > 0 := by 92 | intros h₁ h₂ 93 | have h := mul_lt_mul_of_neg_left h₁ h₂ 94 | simp at h 95 | exact h 96 | 97 | theorem castPos : ∀ (a : Int), a > 0 → ↑a > 0 := by 98 | intros a h 99 | simp [h] 100 | 101 | theorem castNeg : ∀ (a : Int), a < 0 → ↑a < 0 := by 102 | intros a h 103 | simp [h] 104 | 105 | instance : HMul ℤ ℝ ℝ where 106 | hMul z r := ↑z * r 107 | 108 | instance : HMul ℝ ℤ ℝ where 109 | hMul r z := r * ↑z 110 | 111 | end Smt.Reconstruct.Arith 112 | -------------------------------------------------------------------------------- /Smt/Reconstruct/Arith/MulPosNeg.lean: -------------------------------------------------------------------------------- 1 | /- 2 | Copyright (c) 2021-2023 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: Tomaz Gomes Mascarenhas 6 | -/ 7 | 8 | -- implementation of: 9 | -- https://cvc5.github.io/docs/cvc5-1.0.2/proofs/proof_rules.html#_CPPv4N4cvc58internal6PfRule14ARITH_MULT_POSE 10 | -- and 11 | -- https://cvc5.github.io/docs/cvc5-1.0.2/proofs/proof_rules.html#_CPPv4N4cvc58internal6PfRule14ARITH_MULT_NEGE 12 | 13 | import Smt.Reconstruct.Arith.MulPosNeg.Lemmas 14 | import Smt.Reconstruct.Arith.MulPosNeg.Tactic 15 | -------------------------------------------------------------------------------- /Smt/Reconstruct/Arith/MulPosNeg/Instances.lean: -------------------------------------------------------------------------------- 1 | /- 2 | Copyright (c) 2021-2023 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: Tomaz Gomes Mascarenhas 6 | -/ 7 | 8 | import Mathlib.Algebra.Order.Ring.Lemmas 9 | import Mathlib.Data.Real.Basic 10 | 11 | namespace Smt.Reconstruct.Arith 12 | 13 | instance lorInt : LinearOrderedRing Int := inferInstance 14 | noncomputable instance lorReal : LinearOrderedRing Real := inferInstance 15 | 16 | end Smt.Reconstruct.Arith 17 | -------------------------------------------------------------------------------- /Smt/Reconstruct/Arith/SumBounds.lean: -------------------------------------------------------------------------------- 1 | /- 2 | Copyright (c) 2021-2023 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: Tomaz Gomes Mascarenhas 6 | -/ 7 | 8 | import Smt.Reconstruct.Arith.SumBounds.Lemmas 9 | import Smt.Reconstruct.Arith.SumBounds.Tactic 10 | -------------------------------------------------------------------------------- /Smt/Reconstruct/Arith/SumBounds/Instances.lean: -------------------------------------------------------------------------------- 1 | /- 2 | Copyright (c) 2021-2023 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: Tomaz Gomes Mascarenhas 6 | -/ 7 | 8 | import Mathlib.Data.Real.Basic 9 | 10 | namespace Smt.Reconstruct.Arith 11 | 12 | open Function 13 | 14 | instance : CovariantClass Nat Nat (· + ·) (· < ·) where 15 | elim := by 16 | intros a b c h 17 | exact Nat.add_lt_add_left h a 18 | 19 | instance : CovariantClass Nat Nat (swap (· + ·)) (· < ·) where 20 | elim := by 21 | intros a b c h 22 | exact Nat.add_lt_add_right h a 23 | 24 | instance : CovariantClass Nat Nat (· + ·) (· ≤ ·) where 25 | elim := by 26 | intros a b c h 27 | exact Nat.add_le_add_left h a 28 | 29 | instance : CovariantClass Nat Nat (swap (· + ·)) (· ≤ ·) where 30 | elim := by 31 | intros a b c h 32 | exact Nat.add_le_add_right h a 33 | 34 | instance : CovariantClass Int Int (· + ·) (· < ·) where 35 | elim := by 36 | intros a b c h 37 | exact Int.add_lt_add_left h a 38 | 39 | instance : CovariantClass Int Int (swap (· + ·)) (· < ·) where 40 | elim := by 41 | intros a b c h 42 | exact Int.add_lt_add_right h a 43 | 44 | instance : CovariantClass Int Int (· + ·) (· ≤ ·) where 45 | elim := by 46 | intros a b c h 47 | exact Int.add_le_add_left h a 48 | 49 | instance : CovariantClass Int Int (swap (· + ·)) (· ≤ ·) where 50 | elim := by 51 | intros a b c h 52 | exact Int.add_le_add_right h a 53 | 54 | instance : CovariantClass Real Real (· + ·) (· < ·) where 55 | elim := by 56 | intros a b c h 57 | exact add_lt_add_left h a 58 | 59 | instance : CovariantClass Real Real (swap (· + ·)) (· < ·) where 60 | elim := by 61 | intros a b c h 62 | exact add_lt_add_right h a 63 | 64 | instance : CovariantClass Real Real (· + ·) (· ≤ ·) where 65 | elim := by 66 | intros a b c h 67 | exact add_le_add_left h a 68 | 69 | instance : CovariantClass Real Real (swap (· + ·)) (· ≤ ·) where 70 | elim := by 71 | intros a b c h 72 | exact add_le_add_right h a 73 | 74 | end Smt.Reconstruct.Arith 75 | -------------------------------------------------------------------------------- /Smt/Reconstruct/Arith/SumBounds/Lemmas.lean: -------------------------------------------------------------------------------- 1 | /- 2 | Copyright (c) 2021-2023 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: Tomaz Gomes Mascarenhas 6 | -/ 7 | 8 | import Mathlib.Algebra.Order.Ring.Defs 9 | 10 | namespace Smt.Reconstruct.Arith 11 | 12 | open Function 13 | 14 | variable {α : Type} [LinearOrderedRing α] 15 | 16 | variable {a b c d : α} 17 | 18 | theorem sumBounds₁ : a < b → c < d → a + c < b + d := by 19 | intros h₁ h₂ 20 | have r₁: a + c < a + d := add_lt_add_left h₂ a 21 | have r₂: a + d < b + d := add_lt_add_right h₁ d 22 | exact lt_trans r₁ r₂ 23 | 24 | theorem sumBounds₂ : a < b → c ≤ d → a + c < b + d := by 25 | intros h₁ h₂ 26 | have r₁: a + c ≤ a + d := add_le_add_left h₂ a 27 | have r₂: a + d < b + d := add_lt_add_right h₁ d 28 | exact lt_of_le_of_lt r₁ r₂ 29 | 30 | theorem sumBounds₃ : a < b → c = d → a + c < b + d := by 31 | intros h₁ h₂ 32 | rewrite [h₂] 33 | exact add_lt_add_right h₁ d 34 | 35 | theorem sumBounds₄ : a ≤ b → c < d → a + c < b + d := by 36 | intros h₁ h₂ 37 | have r₁ : a + c < a + d := add_lt_add_left h₂ a 38 | have r₂ : a + d ≤ b + d := add_le_add_right h₁ d 39 | exact lt_of_lt_of_le r₁ r₂ 40 | 41 | theorem sumBounds₅ : a ≤ b → c ≤ d → a + c ≤ b + d := by 42 | intros h₁ h₂ 43 | have r₁ : a + c ≤ a + d := add_le_add_left h₂ a 44 | have r₂ : a + d ≤ b + d := add_le_add_right h₁ d 45 | exact le_trans r₁ r₂ 46 | 47 | theorem sumBounds₆ : a ≤ b → c = d → a + c ≤ b + d := by 48 | intros h₁ h₂ 49 | rewrite [h₂] 50 | exact add_le_add_right h₁ d 51 | 52 | theorem sumBounds₇ : a = b → c < d → a + c < b + d := by 53 | intros h₁ h₂ 54 | rewrite [h₁] 55 | exact add_lt_add_left h₂ b 56 | 57 | theorem sumBounds₈ : a = b → c ≤ d → a + c ≤ b + d := by 58 | intros h₁ h₂ 59 | rewrite [h₁] 60 | exact add_le_add_left h₂ b 61 | 62 | theorem sumBounds₉ : a = b → c = d → a + c ≤ b + d := by 63 | intros h₁ h₂ 64 | rewrite [h₁, h₂] 65 | exact le_refl (b + d) 66 | 67 | end Smt.Reconstruct.Arith 68 | -------------------------------------------------------------------------------- /Smt/Reconstruct/Arith/TangentPlane.lean: -------------------------------------------------------------------------------- 1 | /- 2 | Copyright (c) 2021-2023 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: Tomaz Gomes Mascarenhas 6 | -/ 7 | 8 | -- implementation of: 9 | -- https://cvc5.github.io/docs/cvc5-1.0.2/proofs/proof_rules.html#_CPPv4N4cvc58internal6PfRule18ARITH_MULT_TANGENTE 10 | 11 | import Smt.Reconstruct.Arith.TangentPlane.Theorems 12 | -------------------------------------------------------------------------------- /Smt/Reconstruct/Arith/TightBounds.lean: -------------------------------------------------------------------------------- 1 | /- 2 | Copyright (c) 2021-2023 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: Tomaz Gomes Mascarenhas 6 | -/ 7 | 8 | import Smt.Reconstruct.Arith.TightBounds.Lemmas 9 | import Smt.Reconstruct.Arith.TightBounds.Tactic 10 | -------------------------------------------------------------------------------- /Smt/Reconstruct/Arith/TightBounds/Lemmas.lean: -------------------------------------------------------------------------------- 1 | /- 2 | Copyright (c) 2021-2023 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: Tomaz Gomes Mascarenhas 6 | -/ 7 | 8 | import Mathlib.Data.Real.Archimedean 9 | 10 | import Smt.Reconstruct.Arith.MulPosNeg.Lemmas 11 | 12 | namespace Smt.Reconstruct.Arith 13 | 14 | variable {α : Type} [LinearOrderedRing α] [FloorRing α] 15 | 16 | theorem Real.neg_lt_neg {a b : α} (h : a < b) : -a > -b := by 17 | simp 18 | exact h 19 | 20 | theorem Real.neg_le_neg {a b : α} (h : a ≤ b) : -a ≥ -b := by 21 | simp 22 | exact h 23 | 24 | theorem castLE' : ∀ {a b : Int}, a ≤ b → (a : α) ≤ b := by simp 25 | 26 | theorem intTightLb' : ∀ {i : Int} {c : α}, i > c → i ≥ ⌊c⌋ + 1 := by 27 | intros i c h 28 | cases lt_trichotomy i (⌊c⌋ + 1) with 29 | | inl iltc => 30 | have ilec := (@Int.lt_iff_add_one_le i (⌊c⌋ + 1)).mp iltc 31 | simp at ilec 32 | have c_le_floor := Int.floor_le c 33 | have cast_ilec := le_trans (castLE' ilec) c_le_floor 34 | have abs := lt_of_le_of_lt cast_ilec h 35 | simp at abs 36 | | inr h' => cases h' with 37 | | inl ieqc => exact le_of_eq (Eq.symm ieqc) 38 | | inr igtc => exact le_of_lt igtc 39 | 40 | theorem intTightUb' : ∀ {i : Int} {c : α}, i < c → i ≤ ⌈c⌉ - 1 := by 41 | intros i c h 42 | have neg_c_lt_neg_i := Real.neg_lt_neg h 43 | have i_le_floor_neg_c: -i ≥ ⌊-c⌋ + 1 := by 44 | apply intTightLb' 45 | norm_cast at neg_c_lt_neg_i 46 | rw [Int.floor_neg] at i_le_floor_neg_c 47 | have i_plus_one_le_c := Int.neg_le_neg i_le_floor_neg_c 48 | simp at i_plus_one_le_c 49 | rw [add_comm] at i_plus_one_le_c 50 | have pf: i + 1 - 1 ≤ ⌈c⌉ - 1 := Int.add_le_add i_plus_one_le_c le_rfl 51 | simp at pf 52 | exact pf 53 | 54 | end Smt.Reconstruct.Arith 55 | -------------------------------------------------------------------------------- /Smt/Reconstruct/Arith/TightBounds/Tactic.lean: -------------------------------------------------------------------------------- 1 | /- 2 | Copyright (c) 2021-2023 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: Tomaz Gomes Mascarenhas 6 | -/ 7 | 8 | import Smt.Reconstruct.Arith.TightBounds.Lemmas 9 | import Smt.Reconstruct.Util 10 | 11 | import Lean 12 | import Mathlib.Algebra.Order.Floor 13 | 14 | namespace Smt.Reconstruct.Arith 15 | 16 | open Lean 17 | open Meta Elab.Tactic Expr 18 | 19 | def isIntLt : Expr → Bool 20 | | app (app (app (app _ (const ``Int ..)) ..) ..) .. => true 21 | | _ => false 22 | 23 | def intTightMeta (mvar : MVarId) (h : Expr) (thmName outName : Name) 24 | : MetaM MVarId := 25 | mvar.withContext do 26 | let t ← inferType h 27 | let arg ← 28 | if isIntLt t then 29 | mkAppM ``castLT #[h] 30 | else pure h 31 | let answer ← mkAppM thmName #[arg] 32 | let answerType ← inferType answer 33 | let (_, mvar') ← MVarId.intro1P $ ← mvar.assert outName answerType answer 34 | return mvar' 35 | 36 | syntax (name := intTightUb) "intTightUb" term : tactic 37 | @[tactic intTightUb] def evalIntTightUb : Tactic := fun stx => 38 | withMainContext do 39 | trace[smt.debug] m!"[intTightUb] start time: {← IO.monoNanosNow}ns" 40 | let h ← elabTerm stx[1] none 41 | let fname ← mkFreshId 42 | let mvar ← getMainGoal 43 | let mvar' ← intTightMeta mvar h ``intTightUb' fname 44 | replaceMainGoal [mvar'] 45 | evalTactic (← `(tactic| exact $(mkIdent fname))) 46 | trace[smt.debug] m!"[intTightUb] end time: {← IO.monoNanosNow}ns" 47 | 48 | syntax (name := intTightLb) "intTightLb" term : tactic 49 | @[tactic intTightLb] def evalIntTightLb : Tactic := fun stx => 50 | withMainContext do 51 | trace[smt.debug] m!"[intTightLb] start time: {← IO.monoNanosNow}ns" 52 | let h ← elabTerm stx[1] none 53 | let fname ← mkFreshId 54 | let mvar ← getMainGoal 55 | let mvar' ← intTightMeta mvar h ``intTightLb' fname 56 | replaceMainGoal [mvar'] 57 | evalTactic (← `(tactic| exact $(mkIdent fname))) 58 | trace[smt.debug] m!"[intTightLb] end time: {← IO.monoNanosNow}ns" 59 | 60 | end Smt.Reconstruct.Arith 61 | -------------------------------------------------------------------------------- /Smt/Reconstruct/Arith/TransFns.lean: -------------------------------------------------------------------------------- 1 | /- 2 | Copyright (c) 2021-2023 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: Tomaz Gomes Mascarenhas 6 | -/ 7 | 8 | -- implementation of rules about transcendental functions from cvc5 9 | 10 | import Smt.Reconstruct.Arith.TransFns.ArithTransExpApproxAboveNeg 11 | import Smt.Reconstruct.Arith.TransFns.ArithTransExpApproxAbovePos 12 | import Smt.Reconstruct.Arith.TransFns.ArithTransExpApproxBelow 13 | import Smt.Reconstruct.Arith.TransFns.ArithTransExpNeg 14 | import Smt.Reconstruct.Arith.TransFns.ArithTransExpPositivity 15 | import Smt.Reconstruct.Arith.TransFns.ArithTransExpSuperLin 16 | import Smt.Reconstruct.Arith.TransFns.ArithTransExpZero 17 | import Smt.Reconstruct.Arith.TransFns.ArithTransPi 18 | import Smt.Reconstruct.Arith.TransFns.ArithTransSineApproxAboveNeg 19 | import Smt.Reconstruct.Arith.TransFns.ArithTransSineApproxAbovePos 20 | import Smt.Reconstruct.Arith.TransFns.ArithTransSineApproxBelowNeg 21 | import Smt.Reconstruct.Arith.TransFns.ArithTransSineApproxBelowPos 22 | import Smt.Reconstruct.Arith.TransFns.ArithTransSineBounds 23 | import Smt.Reconstruct.Arith.TransFns.ArithTransSineShift 24 | import Smt.Reconstruct.Arith.TransFns.ArithTransSineSymmetry 25 | import Smt.Reconstruct.Arith.TransFns.ArithTransSineTangentPi 26 | import Smt.Reconstruct.Arith.TransFns.ArithTransSineTangentZero 27 | -------------------------------------------------------------------------------- /Smt/Reconstruct/Arith/TransFns/ArithTransExpApproxBelow.lean: -------------------------------------------------------------------------------- 1 | /- 2 | Copyright (c) 2021-2023 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: Harun Khan 6 | -/ 7 | 8 | /- 9 | Implementation of: 10 | https://cvc5.github.io/docs/cvc5-1.0.2/proofs/proof_rules.html#_CPPv4N4cvc58internal6PfRule28ARITH_TRANS_EXP_APPROX_BELOWE 11 | -/ 12 | 13 | import Smt.Reconstruct.Arith.TransFns.Utils 14 | 15 | namespace Smt.Reconstruct.Arith 16 | 17 | open Set Real 18 | 19 | theorem iteratedDeriv_exp (n : Nat) : iteratedDeriv n exp = exp := by 20 | induction' n with n hn 21 | · simp 22 | · simp [iteratedDeriv_succ, hn] 23 | 24 | theorem DifferentiableOn_iteratedDerivWithin {f : ℝ → ℝ} (hf : ContDiff ℝ ⊤ f) (hx : a < b) : 25 | DifferentiableOn ℝ (iteratedDerivWithin d f (Icc a b)) (Ioo a b) := by 26 | apply DifferentiableOn.mono _ Set.Ioo_subset_Icc_self 27 | apply ContDiffOn.differentiableOn_iteratedDerivWithin (n := d + 1) _ (by norm_cast; simp) (uniqueDiffOn_Icc hx) 28 | apply ContDiff.contDiffOn (by apply ContDiff.of_le hf (by norm_cast; simp)) 29 | 30 | -- can definitely be shortened. same proof below 31 | theorem arithTransExpApproxBelow₁ (d n : ℕ) (_ : d = 2*n + 1) (hx : 0 < x) : 32 | Real.exp x ≥ taylorWithinEval Real.exp d Set.univ 0 x := by 33 | have h2 : DifferentiableOn ℝ (iteratedDerivWithin d rexp (Icc 0 x)) (Ioo 0 x) := by 34 | apply DifferentiableOn.mono _ Set.Ioo_subset_Icc_self 35 | apply ContDiffOn.differentiableOn_iteratedDerivWithin (n := d + 1) _ (by norm_cast; simp) (uniqueDiffOn_Icc hx) 36 | apply ContDiff.contDiffOn ((contDiff_infty.mp contDiff_exp) _) 37 | have ⟨x', hx', H⟩ := taylor_mean_remainder_lagrange hx (ContDiff.contDiffOn (s := Icc 0 x) (n := d) contDiff_exp) h2 38 | rw [taylorWithinEval_eq _ (left_mem_Icc.mpr (le_of_lt hx)) (uniqueDiffOn_Icc hx) contDiff_exp] at H 39 | rw [ge_iff_le, ←sub_nonneg, H] 40 | rw [iteratedDerivWithin_eq_iteratedDeriv contDiff_exp (uniqueDiffOn_Icc hx) _ (Ioo_subset_Icc_self hx'), iteratedDeriv_exp] 41 | apply mul_nonneg _ (by apply inv_nonneg.mpr; simp) 42 | apply mul_nonneg (le_of_lt (Real.exp_pos x')) (by simp [le_of_lt hx]) 43 | 44 | 45 | -- see the last line. this probably holds for any function. 46 | theorem arithTransExpApproxBelow₂ (d n : ℕ) (h : d = 2*n + 1) (hx : x < 0) : 47 | Real.exp x ≥ taylorWithinEval Real.exp d Set.univ 0 x := by 48 | have ⟨x', hx', H⟩ := taylor_mean_remainder_lagrange₁ hx contDiff_exp (n := d) 49 | rw [taylorWithinEval_eq _ (right_mem_Icc.mpr (le_of_lt hx)) (uniqueDiffOn_Icc hx) contDiff_exp] at H 50 | rw [ge_iff_le, ←sub_nonneg, H] 51 | rw [iteratedDerivWithin_eq_iteratedDeriv contDiff_exp (uniqueDiffOn_Icc hx) _ (Ioo_subset_Icc_self hx'), iteratedDeriv_exp] 52 | apply mul_nonneg _ (by apply inv_nonneg.mpr; simp) 53 | apply mul_nonneg (le_of_lt (Real.exp_pos x')) 54 | apply Even.pow_nonneg; rw [h, show 2*n + 1 + 1 = 2*(n+1) by ring]; simp 55 | 56 | end Smt.Reconstruct.Arith 57 | 58 | -------------------------------------------------------------------------------- /Smt/Reconstruct/Arith/TransFns/ArithTransExpNeg.lean: -------------------------------------------------------------------------------- 1 | /- 2 | Copyright (c) 2021-2023 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: Tomaz Mascarenhas 6 | -/ 7 | 8 | /- 9 | Implementation of: 10 | https://cvc5.github.io/docs/cvc5-1.0.2/proofs/proof_rules.html#_CPPv4N4cvc58internal6PfRule19ARITH_TRANS_EXP_NEGE 11 | -/ 12 | 13 | import Mathlib.Data.Complex.Exponential 14 | 15 | namespace Smt.Reconstruct.Arith 16 | 17 | theorem arithTransExpNeg (t : ℝ) : t < 0 ↔ Real.exp t < 1 := 18 | Iff.comm.mp Real.exp_lt_one_iff 19 | 20 | end Smt.Reconstruct.Arith 21 | -------------------------------------------------------------------------------- /Smt/Reconstruct/Arith/TransFns/ArithTransExpPositivity.lean: -------------------------------------------------------------------------------- 1 | /- 2 | Copyright (c) 2021-2023 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: Tomaz Mascarenhas 6 | -/ 7 | 8 | /- 9 | Implementation of: 10 | https://cvc5.github.io/docs/cvc5-1.0.2/proofs/proof_rules.html#_CPPv4N4cvc58internal6PfRule26ARITH_TRANS_EXP_POSITIVITYE 11 | -/ 12 | 13 | import Mathlib.Data.Complex.Exponential 14 | 15 | namespace Smt.Reconstruct.Arith 16 | 17 | theorem arithTransExpPositivity (t : Real) : Real.exp t > 0 := Real.exp_pos t 18 | 19 | end Smt.Reconstruct.Arith 20 | -------------------------------------------------------------------------------- /Smt/Reconstruct/Arith/TransFns/ArithTransExpSuperLin.lean: -------------------------------------------------------------------------------- 1 | /- 2 | Copyright (c) 2021-2023 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: Tomaz Mascarenhas 6 | -/ 7 | 8 | /- 9 | Implementation of: 10 | https://cvc5.github.io/docs/cvc5-1.0.2/proofs/proof_rules.html#_CPPv4N4cvc58internal6PfRule25ARITH_TRANS_EXP_SUPER_LINE 11 | -/ 12 | 13 | import Mathlib.Data.Complex.Exponential 14 | 15 | import Smt.Reconstruct.Prop 16 | 17 | namespace Smt.Reconstruct.Arith 18 | 19 | open Smt.Reconstruct.Prop 20 | 21 | theorem arithTransExpSuperLin' (t : ℝ) (ht : t > 0) : 22 | Real.exp t > t + 1 := 23 | Real.add_one_lt_exp (ne_of_gt ht) 24 | 25 | theorem arithTransExpSuperLin (t : ℝ) : 26 | t ≤ 0 ∨ Real.exp t > t + 1 := by 27 | apply orImplies 28 | intro h 29 | simp at h 30 | exact arithTransExpSuperLin' t h 31 | 32 | end Smt.Reconstruct.Arith 33 | -------------------------------------------------------------------------------- /Smt/Reconstruct/Arith/TransFns/ArithTransExpZero.lean: -------------------------------------------------------------------------------- 1 | /- 2 | Copyright (c) 2021-2023 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: Tomaz Mascarenhas 6 | -/ 7 | 8 | /- 9 | Implementation of: 10 | https://cvc5.github.io/docs/cvc5-1.0.2/proofs/proof_rules.html#_CPPv4N4cvc58internal6PfRule20ARITH_TRANS_EXP_ZEROE 11 | -/ 12 | 13 | import Mathlib.Data.Complex.Exponential 14 | 15 | namespace Smt.Reconstruct.Arith 16 | 17 | theorem arithTransExpZero (t : ℝ) : 18 | t = 0 ↔ Real.exp t = 1 := by 19 | apply Iff.intro 20 | · intro h; rw [h]; simp 21 | · intro h; simp at h; assumption 22 | 23 | end Smt.Reconstruct.Arith 24 | -------------------------------------------------------------------------------- /Smt/Reconstruct/Arith/TransFns/ArithTransSineApproxAboveNeg.lean: -------------------------------------------------------------------------------- 1 | /- 2 | Copyright (c) 2021-2023 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: Harun Khan, Tomaz Mascarenhas 6 | -/ 7 | 8 | /- 9 | Implementation of: 10 | https://cvc5.github.io/docs/cvc5-1.0.2/proofs/proof_rules.html#_CPPv4N4cvc58internal6PfRule33ARITH_TRANS_SINE_APPROX_ABOVE_NEGE 11 | -/ 12 | 13 | import Mathlib.Analysis.SpecialFunctions.Trigonometric.Deriv 14 | import Mathlib.Analysis.Convex.SpecificFunctions.Deriv 15 | 16 | import Smt.Reconstruct.Arith.TransFns.ArithTransExpApproxAbovePos 17 | import Smt.Reconstruct.Arith.TransFns.ArithTransExpApproxAboveNeg 18 | import Smt.Reconstruct.Arith.TransFns.Utils 19 | 20 | open Set Real 21 | 22 | namespace Smt.Reconstruct.Arith 23 | 24 | theorem sineApproxAboveNeg (d k : Nat) (hd : d = 4*k + 3) (hx : x < 0) (hx2 : -π ≤ x): 25 | let p : ℕ → ℝ → ℝ := fun d => taylorWithinEval Real.sin d Set.univ 0 26 | sin x ≤ p d x := by 27 | intro p 28 | have ⟨x', hx', H⟩ := taylor_mean_remainder_lagrange₁ (n := d) hx contDiff_sin 29 | rw [taylorWithinEval_eq _ (right_mem_Icc.mpr (le_of_lt hx)) (uniqueDiffOn_Icc hx) (contDiff_sin)] at H 30 | 31 | rw [←sub_nonpos, H] 32 | rw [iteratedDerivWithin_eq_iteratedDeriv contDiff_sin (uniqueDiffOn_Icc hx) _ (Ioo_subset_Icc_self hx')] 33 | have : (d+1)%4 = 0 := by simp [hd, Nat.add_mod] 34 | simp only [this, iteratedDeriv_sin_cos, reduceIte, three_ne_zero, sub_zero, show 3 ≠ 1 by decide, show 3 ≠ 0 by decide, show 3 ≠ 2 by decide] 35 | apply mul_nonpos_of_nonpos_of_nonneg _ (by apply inv_nonneg.mpr; simp) 36 | apply mul_nonpos_of_nonpos_of_nonneg (Real.sin_nonpos_of_nonnpos_of_neg_pi_le (le_of_lt (mem_Ioo.mp hx').2) (le_trans hx2 (le_of_lt (mem_Ioo.mp hx').1))) 37 | apply Even.pow_nonneg (by rw [even_iff_two_dvd]; omega) 38 | 39 | theorem arithTransSineApproxAboveNeg (d k : Nat) (hd : d = 4*k + 3) (l u t : ℝ) 40 | (ht : l ≤ t ∧ t ≤ u) (hu : u < 0) (hl : -π ≤ l) : 41 | let p: ℝ → ℝ := taylorWithinEval Real.sin d Set.univ 0 42 | Real.sin t ≤ ((p l - p u) / (l - u)) * (t - l) + p l := by 43 | intro p 44 | have hp : ∀ x, p x = taylorWithinEval Real.sin d Set.univ 0 x := by simp 45 | apply le_convex_of_le ht 46 | (by rw [hp]; exact sineApproxAboveNeg d k hd (by linarith) hl) 47 | (by rw [hp]; exact sineApproxAboveNeg d k hd hu (by linarith)) 48 | convexOn_sin_Icc (mem_Icc.mpr ⟨hl, by linarith⟩) 49 | (mem_Icc.mpr ⟨by linarith, le_of_lt hu⟩) 50 | -------------------------------------------------------------------------------- /Smt/Reconstruct/Arith/TransFns/ArithTransSineApproxAbovePos.lean: -------------------------------------------------------------------------------- 1 | /- 2 | Copyright (c) 2021-2023 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: Harun Khan 6 | -/ 7 | 8 | /- 9 | Implementation of: 10 | https://cvc5.github.io/docs/cvc5-1.0.2/proofs/proof_rules.html#_CPPv4N4cvc58internal6PfRule33ARITH_TRANS_SINE_APPROX_ABOVE_POSE 11 | -/ 12 | 13 | import Smt.Reconstruct.Arith.TransFns.ArithTransSineApproxBelowNeg 14 | 15 | open Set Real 16 | 17 | namespace Smt.Reconstruct.Arith 18 | 19 | theorem arithTransSineApproxAbovePos (d k : ℕ) (hd : d = 4*k + 1) 20 | (hx : 0 < x) (hx2 : x ≤ π) : 21 | Real.sin x ≤ taylorWithinEval Real.sin d Set.univ 0 x := by 22 | rw [← neg_neg x, sin_neg, taylorSin_neg, neg_le_neg_iff] 23 | apply arithTransSineApproxBelowNeg d k hd (by linarith) (by linarith) 24 | -------------------------------------------------------------------------------- /Smt/Reconstruct/Arith/TransFns/ArithTransSineApproxBelowNeg.lean: -------------------------------------------------------------------------------- 1 | /- 2 | Copyright (c) 2021-2023 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: Harun Khan 6 | -/ 7 | 8 | /- 9 | Implementation of: 10 | https://cvc5.github.io/docs/cvc5-1.0.2/proofs/proof_rules.html#_CPPv4N4cvc58internal6PfRule33ARITH_TRANS_SINE_APPROX_BELOW_NEGE 11 | -/ 12 | 13 | import Smt.Reconstruct.Arith.TransFns.ArithTransSineApproxBelowPos 14 | 15 | open Set Real 16 | 17 | namespace Smt.Reconstruct.Arith 18 | 19 | theorem arithTransSineApproxBelowNeg (d k : Nat) (hd : d = 4*k + 1) (hx : x < 0) (hx2 : -π ≤ x): 20 | let p : ℕ → ℝ → ℝ := fun d => taylorWithinEval Real.sin d Set.univ 0 21 | p d x ≤ sin x := by 22 | intro p 23 | have ⟨x', hx', H⟩ := taylor_mean_remainder_lagrange₁ (n := d) hx contDiff_sin 24 | rw [taylorWithinEval_eq _ (right_mem_Icc.mpr (le_of_lt hx)) (uniqueDiffOn_Icc hx) (contDiff_sin)] at H 25 | rw [←sub_nonneg, H] 26 | rw [iteratedDerivWithin_eq_iteratedDeriv contDiff_sin (uniqueDiffOn_Icc hx) _ (Ioo_subset_Icc_self hx')] 27 | have : (d+1)%4 = 2 := by simp [hd, Nat.add_mod] 28 | simp only [this, iteratedDeriv_sin_cos, reduceIte, one_ne_zero, sub_zero, show 1 ≠ 3 by decide, show 1 ≠ 2 by decide, two_ne_zero, show 2 ≠ 1 by decide] 29 | apply mul_nonneg _ (by apply inv_nonneg.mpr; simp) 30 | simp only [Pi.neg_apply, neg_mul, Left.nonneg_neg_iff] 31 | apply mul_nonpos_of_nonpos_of_nonneg (Real.sin_nonpos_of_nonnpos_of_neg_pi_le (le_of_lt (mem_Ioo.mp hx').2) (le_trans hx2 (le_of_lt (mem_Ioo.mp hx').1))) 32 | apply Even.pow_nonneg (by rw [even_iff_two_dvd]; omega) 33 | -------------------------------------------------------------------------------- /Smt/Reconstruct/Arith/TransFns/ArithTransSineBounds.lean: -------------------------------------------------------------------------------- 1 | /- 2 | Copyright (c) 2021-2023 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: Tomaz Mascarenhas 6 | -/ 7 | 8 | /- 9 | Implementation of: 10 | https://cvc5.github.io/docs/cvc5-1.0.2/proofs/proof_rules.html#_CPPv4N4cvc58internal6PfRule23ARITH_TRANS_SINE_BOUNDSE 11 | -/ 12 | 13 | import Mathlib.Data.Complex.Exponential 14 | 15 | namespace Smt.Reconstruct.Arith 16 | 17 | open Real 18 | 19 | theorem arithTransSineBounds : ∀ (t : ℝ), sin t ≤ 1 ∧ sin t ≥ -1 := by 20 | intro t 21 | apply And.intro 22 | · exact sin_le_one t 23 | · exact neg_one_le_sin t 24 | 25 | end Smt.Reconstruct.Arith 26 | -------------------------------------------------------------------------------- /Smt/Reconstruct/Arith/TransFns/ArithTransSineSymmetry.lean: -------------------------------------------------------------------------------- 1 | /- 2 | Copyright (c) 2021-2023 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: Tomaz Mascarenhas 6 | -/ 7 | 8 | /- 9 | Implementation of: 10 | https://cvc5.github.io/docs/cvc5-1.0.2/proofs/proof_rules.html#_CPPv4N4cvc58internal6PfRule25ARITH_TRANS_SINE_SYMMETRYE 11 | -/ 12 | 13 | import Mathlib.Data.Complex.Exponential 14 | 15 | namespace Smt.Reconstruct.Arith 16 | 17 | open Real 18 | 19 | theorem arithTransSineSymmetry : ∀ (t : ℝ), sin t + sin (- t) = 0 := by 20 | intro t 21 | rw [sin_neg] 22 | simp 23 | 24 | end Smt.Reconstruct.Arith 25 | -------------------------------------------------------------------------------- /Smt/Reconstruct/Arith/TransFns/ArithTransSineTangentPi.lean: -------------------------------------------------------------------------------- 1 | /- 2 | Copyright (c) 2021-2023 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: Tomaz Mascarenhas 6 | -/ 7 | 8 | /- 9 | Implementation of: 10 | https://cvc5.github.io/docs/cvc5-1.0.2/proofs/proof_rules.html#_CPPv4N4cvc58internal6PfRule27ARITH_TRANS_SINE_TANGENT_PIE 11 | -/ 12 | 13 | import Mathlib.Analysis.SpecialFunctions.Trigonometric.Basic 14 | import Mathlib.Analysis.SpecialFunctions.Trigonometric.Bounds 15 | import Mathlib.Data.Complex.Exponential 16 | 17 | namespace Smt.Reconstruct.Arith 18 | 19 | open Real 20 | 21 | theorem arithTransSineTangentPi₁ : ∀ (t : ℝ), 22 | t > - Real.pi → sin t > - Real.pi - t := by 23 | intros t ht 24 | set ε := Real.pi + t 25 | have e_pos : 0 < ε := neg_lt_iff_pos_add'.mp ht 26 | have te : t = -Real.pi + ε := eq_neg_add_of_add_eq rfl 27 | rw [te] 28 | have : sin (-Real.pi + ε) = - Real.sin ε := 29 | by rw [add_comm, <- sub_eq_add_neg]; exact sin_sub_pi ε 30 | rw [this] 31 | simp only [sub_add_cancel_left, gt_iff_lt, neg_lt_neg_iff] 32 | exact sin_lt e_pos 33 | 34 | theorem arithTransSineTangentPi₂ : ∀ (t : ℝ), 35 | t < Real.pi → sin t < Real.pi - t := by 36 | intros t ht 37 | set ε := Real.pi - t 38 | have e_pos : 0 < ε := sub_pos.mpr ht 39 | have te : t = Real.pi - ε := Eq.symm (sub_sub_self π t) 40 | rw [te] 41 | have : sin (Real.pi - ε) = sin ε := sin_pi_sub ε 42 | rw [this] 43 | exact sin_lt e_pos 44 | 45 | theorem arithTransSineTangentPi : ∀ (t : ℝ), 46 | (t > -Real.pi → sin t > -Real.pi - t) ∧ (t < Real.pi → sin t < Real.pi - t) := fun t => 47 | ⟨arithTransSineTangentPi₁ t, arithTransSineTangentPi₂ t⟩ 48 | 49 | end Smt.Reconstruct.Arith 50 | -------------------------------------------------------------------------------- /Smt/Reconstruct/Arith/TransFns/ArithTransSineTangentZero.lean: -------------------------------------------------------------------------------- 1 | /- 2 | Copyright (c) 2021-2023 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: Tomaz Mascarenhas 6 | -/ 7 | 8 | /- 9 | Implementation of: 10 | https://cvc5.github.io/docs/cvc5-1.0.2/proofs/proof_rules.html#_CPPv4N4cvc58internal6PfRule29ARITH_TRANS_SINE_TANGENT_ZEROE 11 | -/ 12 | 13 | import Mathlib.Analysis.SpecialFunctions.Trigonometric.Bounds 14 | import Mathlib.Data.Complex.Exponential 15 | 16 | namespace Smt.Reconstruct.Arith 17 | 18 | open Real 19 | 20 | theorem sin_neg' : ∀ (t : ℝ), sin t = - sin (-t) := by 21 | intro t 22 | rw [sin_neg t] 23 | simp 24 | 25 | theorem arithTransSinTangentZero : ∀ (t : ℝ), 26 | (t > 0 → sin t < t) ∧ (t < 0 → sin t > t) := by 27 | intro t 28 | apply And.intro 29 | · exact sin_lt 30 | · intro ht 31 | rw [sin_neg'] 32 | apply lt_neg.mpr 33 | apply sin_lt 34 | linarith 35 | 36 | end Smt.Reconstruct.Arith 37 | -------------------------------------------------------------------------------- /Smt/Reconstruct/Arith/Trichotomy.lean: -------------------------------------------------------------------------------- 1 | /- 2 | Copyright (c) 2021-2023 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: Tomaz Gomes Mascarenhas 6 | -/ 7 | 8 | -- implementation of: 9 | -- https://cvc5.github.io/docs/cvc5-1.0.2/proofs/proof_rules.html#_CPPv4N4cvc58internal6PfRule16ARITH_TRICHOTOMYE 10 | 11 | import Smt.Reconstruct.Arith.Trichotomy.Lemmas 12 | import Smt.Reconstruct.Arith.Trichotomy.Tactic 13 | -------------------------------------------------------------------------------- /Smt/Reconstruct/Arith/Trichotomy/Lemmas.lean: -------------------------------------------------------------------------------- 1 | /- 2 | Copyright (c) 2021-2023 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: Tomaz Gomes Mascarenhas 6 | -/ 7 | 8 | import Aesop 9 | import Mathlib.Order.Defs.LinearOrder 10 | 11 | import Smt.Reconstruct.Prop.Lemmas 12 | 13 | namespace Smt.Reconstruct.Arith 14 | 15 | open Smt.Reconstruct.Prop 16 | 17 | variable {α : Type} 18 | 19 | variable [LinearOrder α] 20 | 21 | variable {a b : α} 22 | 23 | theorem trichotomy₁ : ¬ a < b → ¬ a = b → a > b := by 24 | intros h₁ h₂ 25 | have tr := lt_trichotomy a b 26 | exact (orImplies₃ ((orImplies₃ tr) h₁)) h₂ 27 | 28 | theorem trichotomy₂ : ¬ a = b → ¬ a < b → a > b := by 29 | intros h₁ h₂ 30 | have tr := lt_trichotomy a b 31 | exact (orImplies₃ ((orImplies₃ tr) h₂)) h₁ 32 | 33 | theorem trichotomy₃ : ¬ a < b → ¬ a > b → a = b := by 34 | intros h₁ h₂ 35 | have tr := lt_trichotomy a b 36 | have tr': a < b ∨ b < a ∨ a = b := by aesop 37 | exact (orImplies₃ ((orImplies₃ tr') h₁)) h₂ 38 | 39 | theorem trichotomy₄ : ¬ a > b → ¬ a < b → a = b := by 40 | intros h₁ h₂ 41 | have tr := lt_trichotomy a b 42 | have tr': a < b ∨ b < a ∨ a = b := by aesop 43 | exact (orImplies₃ ((orImplies₃ tr') h₂)) h₁ 44 | 45 | theorem trichotomy₅ : ¬ a = b → ¬ a > b → a < b := by 46 | intros h₁ h₂ 47 | have tr := lt_trichotomy a b 48 | have tr': a = b ∨ b < a ∨ a < b := by aesop 49 | exact (orImplies₃ ((orImplies₃ tr') h₁)) h₂ 50 | 51 | theorem trichotomy₆ : ¬ a > b → ¬ a = b → a < b := by 52 | intros h₁ h₂ 53 | have tr := lt_trichotomy a b 54 | have tr' : a = b ∨ b < a ∨ a < b := by aesop 55 | exact (orImplies₃ ((orImplies₃ tr') h₂)) h₁ 56 | 57 | theorem not_gt_of_le : a ≤ b → ¬ a > b := 58 | λ h₁ h₂ => absurd h₂ (not_lt_of_ge h₁) 59 | 60 | end Smt.Reconstruct.Arith 61 | -------------------------------------------------------------------------------- /Smt/Reconstruct/Arith/Trichotomy/Tactic.lean: -------------------------------------------------------------------------------- 1 | /- 2 | Copyright (c) 2021-2023 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: Tomaz Gomes Mascarenhas 6 | -/ 7 | 8 | import Lean 9 | 10 | -- we must import those to have visible instances of LinearOrder of 11 | -- Nat, Int and Real 12 | import Mathlib.Data.Nat.Order.Lemmas 13 | import Mathlib.Data.Int.Order.Lemmas 14 | import Mathlib.Data.Real.Basic 15 | 16 | import Smt.Reconstruct.Arith.Trichotomy.Lemmas 17 | 18 | namespace Smt.Reconstruct.Arith 19 | 20 | open Lean 21 | open Meta Elab.Tactic Expr 22 | 23 | def trichotomyMeta (mvar : MVarId) (h₁ h₂ : Expr) (name : Name) : MetaM MVarId := 24 | mvar.withContext do 25 | let t₁ ← inferType h₁ 26 | let t₂ ← inferType h₂ 27 | let thmName : Name ← 28 | match ← getOp t₁, ← getOp t₂ with 29 | | ``LT.lt , ``Eq => pure ``trichotomy₁ 30 | | ``Eq , ``LT.lt => pure ``trichotomy₂ 31 | | ``LT.lt , ``GT.gt => pure ``trichotomy₃ 32 | | ``GT.gt , ``LT.lt => pure ``trichotomy₄ 33 | | ``Eq , ``GT.gt => pure ``trichotomy₅ 34 | | ``GT.gt , ``Eq => pure ``trichotomy₆ 35 | | _ , _ => throwError "[trichotomy] invalid operation" 36 | let answer ← mkAppM thmName #[h₁, h₂] 37 | let answerType ← inferType answer 38 | let (_, mvar') ← MVarId.intro1P $ ← mvar.assert name answerType answer 39 | return mvar' 40 | 41 | syntax (name := trichotomy) "trichotomy" term "," term : tactic 42 | 43 | @[tactic trichotomy] def evalTrichotomy : Tactic := fun stx => 44 | withMainContext do 45 | trace[smt.profile.reconstruct] m!"[trichotomy] start time: {← IO.monoNanosNow}ns" 46 | let h₁ ← notLeToLt (← elabTerm stx[1] none) 47 | let h₂ ← notLeToLt (← elabTerm stx[3] none) 48 | let fname ← mkFreshId 49 | let mvar ← getMainGoal 50 | let mvar' ← trichotomyMeta mvar h₁ h₂ fname 51 | replaceMainGoal [mvar'] 52 | evalTactic (← `(tactic| exact $(mkIdent fname))) 53 | trace[smt.profile.reconstruct] m!"[trichotomy] end time: {← IO.monoNanosNow}ns" 54 | where 55 | notLeToLt : Expr → MetaM Expr := λ e => do 56 | match ← inferType e with 57 | | app (app (app (app (app (const ``LE.le ..) ..) ..) ..) ..) _ => mkAppM ``not_gt_of_le #[e] 58 | | app (app (app (app (const ``LE.le ..) ..) ..) ..) _ => mkAppM ``not_gt_of_le #[e] 59 | | app (app (app (app (app (const ``GE.ge ..) ..) ..) ..) ..) _ => mkAppM ``not_lt_of_ge #[e] 60 | | app (app (app (app (const ``GE.ge ..) ..) ..) ..) _ => mkAppM ``not_lt_of_ge #[e] 61 | | _ => pure e 62 | 63 | end Smt.Reconstruct.Arith 64 | -------------------------------------------------------------------------------- /Smt/Reconstruct/BitVec/Bitblast.lean: -------------------------------------------------------------------------------- 1 | /- 2 | Copyright (c) 2021-2024 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: Abdalrhman Mohamed 6 | -/ 7 | 8 | import Lean 9 | 10 | namespace BitVec 11 | 12 | def bb (x : BitVec w) : BitVec w := 13 | (iunfoldr (fun i _ => ((), x.getLsbD i)) ()).snd 14 | 15 | theorem self_eq_bb (x : BitVec w) : x = x.bb := by 16 | unfold bb 17 | rw [iunfoldr_replace_snd (λ _ => ()) (init:=rfl)] 18 | intro i 19 | rfl 20 | 21 | -- def beq (x : BitVec w) (y : BitVec w) : Bool := 22 | -- go w 23 | -- where 24 | -- go : Nat → Bool 25 | -- | 0 => x.getLsb 0 == y.getLsb 0 26 | -- | i + 1 => go i && x.getLsb i == y.getLsb i 27 | 28 | -- TODO: fix reconstruction associativity of ∧ to use above version 29 | def beq (x : BitVec w) (y : BitVec w) : Bool := 30 | go (w - 1) 31 | where 32 | go : Nat → Bool 33 | | 0 => x.getLsbD (w - 1) == y.getLsbD (w - 1) 34 | | i + 1 => x.getLsbD ((w - 1) - (i + 1)) == y.getLsbD ((w - 1) - (i + 1)) && go i 35 | 36 | def eq_eq_beq (x : BitVec w) (y : BitVec w) : (x = y) = x.beq y := 37 | sorry 38 | 39 | /-- Carry function for bitwise addition. -/ 40 | def adcb' (x y c : Bool) : Bool × Bool := (x && y || Bool.xor x y && c, (Bool.xor (Bool.xor x y) c)) 41 | 42 | theorem adcb_eq_adcb' : adcb = adcb' := by 43 | funext x y c 44 | cases x <;> cases y <;> cases c <;> rfl 45 | 46 | /-- Bitwise addition implemented via a ripple carry adder. -/ 47 | def adc' (x y : BitVec w) : Bool → Bool × BitVec w := 48 | iunfoldr fun (i : Fin w) c => adcb' (x.getLsbD i) (y.getLsbD i) c 49 | 50 | theorem adc_eq_adc' : @adc = @adc' := by 51 | funext x y c 52 | rw [adc, adc', adcb_eq_adcb'] 53 | 54 | theorem add_eq_adc' (x y : BitVec w) : x + y = (adc' x y false).snd := by 55 | rw [add_eq_adc, adc_eq_adc'] 56 | 57 | end BitVec 58 | -------------------------------------------------------------------------------- /Smt/Reconstruct/Bool/Lemmas.lean: -------------------------------------------------------------------------------- 1 | /- 2 | Copyright (c) 2021-2024 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: Abdalrhman Mohamed 6 | -/ 7 | 8 | import Smt.Reconstruct.Prop.Core 9 | 10 | namespace Smt.Reconstruct.Bool 11 | 12 | theorem decide_true_eq : decide True = true := rfl 13 | 14 | theorem decide_false_eq : decide False = false := rfl 15 | 16 | theorem decide_not_eq [hp: Decidable p] : decide (¬p) = !(decide p) := 17 | match hp with 18 | | isFalse _ => rfl 19 | | isTrue _ => rfl 20 | 21 | theorem decide_and_eq [hp: Decidable p] [hq: Decidable q] : decide (p ∧ q) = (decide p && decide q) := 22 | match hp, hq with 23 | | isFalse _, isFalse _ => rfl 24 | | isFalse _, isTrue _ => rfl 25 | | isTrue _, isFalse _ => rfl 26 | | isTrue _, isTrue _ => rfl 27 | 28 | theorem decide_or_eq [hp: Decidable p] [hq: Decidable q] : decide (p ∨ q) = (decide p || decide q) := 29 | match hp, hq with 30 | | isFalse _, isFalse _ => rfl 31 | | isFalse _, isTrue _ => rfl 32 | | isTrue _, isFalse _ => rfl 33 | | isTrue _, isTrue _ => rfl 34 | 35 | theorem decide_eq_eq [hp: Decidable p] [hq: Decidable q] : decide (p = q) = (decide p == decide q) := 36 | match hp, hq with 37 | | isFalse _, isFalse _ => rfl 38 | | isFalse _, isTrue _ => rfl 39 | | isTrue _, isFalse _ => rfl 40 | | isTrue _, isTrue _ => rfl 41 | 42 | theorem decide_xor_eq [hp: Decidable p] [hq: Decidable q] : decide (XOr p q) = Bool.xor (decide p) (decide q) := 43 | match hp, hq with 44 | | isFalse _, isFalse _ => rfl 45 | | isFalse _, isTrue _ => rfl 46 | | isTrue _, isFalse _ => rfl 47 | | isTrue _, isTrue _ => rfl 48 | 49 | theorem eq_of_decide_eq [hp : Decidable p] [hq : Decidable q] (_ : decide p = decide q) : p = q := 50 | match hp, hq with 51 | | .isFalse hnp, .isFalse hnq => propext ⟨fun hp => False.elim (hnp hp), fun hq => False.elim (hnq hq)⟩ 52 | | .isTrue hp, .isTrue hq => propext ⟨fun _ => hq, fun _ => hp⟩ 53 | 54 | end Smt.Reconstruct.Bool 55 | -------------------------------------------------------------------------------- /Smt/Reconstruct/Bool/Tactic.lean: -------------------------------------------------------------------------------- 1 | /- 2 | Copyright (c) 2021-2024 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: Abdalrhman Mohamed 6 | -/ 7 | 8 | import Lean 9 | 10 | import Smt.Reconstruct.Bool.Lemmas 11 | 12 | namespace Smt.Reconstruct.Bool 13 | 14 | open Lean 15 | 16 | def traceBoolify (r : Except Exception MVarId) : MetaM MessageData := 17 | return match r with 18 | | .ok _ => m!"{checkEmoji}" 19 | | _ => m!"{bombEmoji}" 20 | 21 | def boolify (mv : MVarId) : MetaM MVarId := withTraceNode `smt.reconstruct.boolify traceBoolify do 22 | let ns := [``Bool.decide_eq_true, ``decide_true_eq, ``decide_false_eq, ``decide_not_eq, ``decide_and_eq, ``decide_or_eq, ``decide_eq_eq, ``decide_xor_eq] 23 | let simpTheorems ← ns.foldrM (fun n a => a.addTheorem (.decl n) (.const n [])) {} 24 | let ctx ← Meta.Simp.mkContext {} simpTheorems 25 | let (some mv, _) ← Meta.simpTarget mv ctx | throwError "failed to boolify goal:{mv}" 26 | return mv 27 | 28 | namespace Tactic 29 | 30 | syntax (name := boolify) "boolify" : tactic 31 | 32 | open Lean.Elab Tactic in 33 | @[tactic boolify] def evalBoolify : Tactic := fun _ => withMainContext do 34 | let mv ← Tactic.getMainGoal 35 | let mv ← Bool.boolify mv 36 | replaceMainGoal [mv] 37 | 38 | end Smt.Reconstruct.Bool.Tactic 39 | -------------------------------------------------------------------------------- /Smt/Reconstruct/Builtin/AC.lean: -------------------------------------------------------------------------------- 1 | import Lean 2 | 3 | theorem Eq.same_root (hac : a = c) (hbc : b = c) : a = b := hac ▸ hbc ▸ rfl 4 | 5 | namespace Lean.Meta.AC 6 | 7 | open Lean.Elab Tactic 8 | 9 | /-- Similar to `rewriteUnnormalized`, but rewrite is only applied at the top level. -/ 10 | def rewriteUnnormalizedTop (mv : MVarId) : MetaM Unit := do 11 | let some (α, l, r) := (← mv.getType).eq? 12 | | throwError "[ac_rfl_top] expected a top level equality with AC operator on lhs and/or rhs, got {← mv.getType}" 13 | let lvl ← Meta.getLevel α 14 | let (nl, pl) ← normalize l 15 | let (nr, pr) ← normalize r 16 | if nl == r then 17 | let some pl := pl | throwError "[ac_rfl_top] expected {l} to have an AC operator" 18 | mv.assign pl 19 | else if l == nr then 20 | let some pr := pr | throwError "[ac_rfl_top] expected {r} to have an AC operator" 21 | mv.assign (mkApp4 (.const ``Eq.symm [lvl]) α r l pr) 22 | else if nl == nr then 23 | let some pl := pl | throwError "[ac_rfl_top] expected {l} to have an AC operator" 24 | let some pr := pr | throwError "[ac_rfl_top] expected {r} to have an AC operator" 25 | mv.assign (mkApp6 (.const ``Eq.same_root [lvl]) α l nl r pl pr) 26 | else 27 | throwError "[ac_rfl_top] expected {l} and {r} to have the same AC operator" 28 | where 29 | normalize (e : Expr) : MetaM (Expr × Option Expr) := do 30 | let bin op l r := e | return (e, none) 31 | let some pc ← preContext op | return (e, none) 32 | let (p, ne) ← buildNormProof pc l r 33 | return (ne, some p) 34 | 35 | syntax (name := ac_rfl_top) "ac_rfl_top" : tactic 36 | 37 | @[tactic ac_rfl_top] def evalacRflTop : Lean.Elab.Tactic.Tactic := fun _ => do 38 | let goal ← getMainGoal 39 | goal.withContext (rewriteUnnormalizedTop goal) 40 | 41 | local instance : Std.Associative (α := Nat) (· + ·) := ⟨Nat.add_assoc⟩ 42 | local instance : Std.Commutative (α := Nat) (· + ·) := ⟨Nat.add_comm⟩ 43 | 44 | example (a b c d : Nat) : 2 * (a + b + c + d) = 2 * (d + (b + c) + a) := by 45 | try ac_rfl_top 46 | ac_rfl 47 | 48 | example (a b c d : Nat) : a + b + c + d + (2 * (a + b)) = d + (b + c) + a + (2 * (b + a)) := by 49 | try ac_rfl_top 50 | ac_rfl 51 | 52 | example (a b c d : Nat) : a + b + c + d + (a + b) = d + (b + c) + a + (b + a) := by 53 | ac_rfl_top 54 | 55 | end Lean.Meta.AC 56 | -------------------------------------------------------------------------------- /Smt/Reconstruct/Builtin/Absorb.lean: -------------------------------------------------------------------------------- 1 | /- 2 | Copyright (c) 2021-2025 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: Abdalrhman Mohamed 6 | -/ 7 | 8 | namespace Smt.Reconstruct.Builtin 9 | 10 | class Absorb (mul : α → α → α) where 11 | zero : α 12 | /-- Zero is a left absorbing element for multiplication -/ 13 | zero_mul : ∀ (a : α), mul zero a = zero 14 | /-- Zero is a right absorbing element for multiplication -/ 15 | mul_zero : ∀ (a : α), mul a zero = zero 16 | 17 | instance : Absorb And where 18 | zero := False 19 | zero_mul := false_and 20 | mul_zero := and_false 21 | 22 | instance : Absorb Or where 23 | zero := True 24 | zero_mul := true_or 25 | mul_zero := or_true 26 | 27 | instance : @Absorb Int (· * ·) where 28 | zero := 0 29 | zero_mul := Int.zero_mul 30 | mul_zero := Int.mul_zero 31 | 32 | instance : @Absorb (BitVec w) (· &&& ·) where 33 | zero := 0#w 34 | zero_mul := @BitVec.zero_and w 35 | mul_zero := @BitVec.and_zero w 36 | 37 | instance : @Absorb (BitVec w) (· ||| ·) where 38 | zero := BitVec.allOnes w 39 | zero_mul := @BitVec.allOnes_or w 40 | mul_zero := @BitVec.or_allOnes w 41 | 42 | instance : @Absorb (BitVec w) (· * ·) where 43 | zero := 0#w 44 | zero_mul := @BitVec.zero_mul w 45 | mul_zero := @BitVec.mul_zero w 46 | 47 | namespace Absorb 48 | 49 | def Context α := Nat → α 50 | 51 | inductive Expr where 52 | | zero 53 | | atom (v : Nat) 54 | | op (l r : Expr) 55 | deriving Inhabited, Repr 56 | 57 | namespace Expr 58 | 59 | def containsZero : Expr → Bool 60 | | Expr.zero => true 61 | | Expr.atom _ => false 62 | | Expr.op l r => containsZero l || containsZero r 63 | 64 | def eval [hα : @Absorb α mul] (ctx : Context α) : Expr → α 65 | | Expr.zero => hα.zero 66 | | Expr.atom v => ctx v 67 | | Expr.op l r => mul (eval (hα := hα) ctx l) (eval (hα := hα) ctx r) 68 | 69 | theorem eval_eq_zero_from_containsZero [hα : @Absorb α mul] (ctx : Context α) (h : containsZero e) : 70 | eval (hα := hα) ctx e = hα.zero := by 71 | induction e with 72 | | zero => rfl 73 | | atom v => contradiction 74 | | op l r ih₁ ih₂ => 75 | unfold eval 76 | simp only [containsZero, Bool.or_eq_true] at h 77 | cases h <;> rename_i h 78 | · rw [ih₁ h, hα.zero_mul] 79 | · rw [ih₂ h, hα.mul_zero] 80 | 81 | end Expr 82 | 83 | end Smt.Reconstruct.Builtin.Absorb 84 | -------------------------------------------------------------------------------- /Smt/Reconstruct/Builtin/Lemmas.lean: -------------------------------------------------------------------------------- 1 | /- 2 | Copyright (c) 2021-2023 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: Abdalrhman Mohamed 6 | -/ 7 | 8 | import Smt.Reconstruct.Prop.Core 9 | 10 | namespace Smt.Reconstruct.Builtin 11 | 12 | theorem iteElim1 [hc : Decidable c] : ite c a b → ¬ c ∨ a := by 13 | intros h 14 | cases hc with 15 | | isTrue hc => exact Or.inr h 16 | | isFalse hnc => exact Or.inl hnc 17 | 18 | theorem iteElim2 [hc : Decidable c] : ite c a b → c ∨ b := by 19 | intros h 20 | cases hc with 21 | | isTrue hc => exact Or.inl hc 22 | | isFalse hnc => exact Or.inr h 23 | 24 | theorem notIteElim1 [hc : Decidable c] : ¬ ite c a b → ¬ c ∨ ¬ a := by 25 | intros h 26 | cases hc with 27 | | isTrue hc => exact Or.inr h 28 | | isFalse hnc => exact Or.inl hnc 29 | 30 | theorem notIteElim2 [hc : Decidable c] : ¬ ite c a b → c ∨ ¬ b := by 31 | intros h 32 | cases hc with 33 | | isTrue hc => exact Or.inl hc 34 | | isFalse hnc => exact Or.inr h 35 | 36 | theorem orImplies : ∀ {p q : Prop}, (¬ p → q) → p ∨ q := 37 | by intros p q h 38 | exact match Classical.em p with 39 | | Or.inl pp => Or.inl pp 40 | | Or.inr npp => match Classical.em q with 41 | | Or.inl pq => Or.inr pq 42 | | Or.inr npq => False.elim (npq (h npp)) 43 | 44 | theorem notNotElim : ∀ {p : Prop}, ¬ ¬ p → p := 45 | by intros p h 46 | exact match Classical.em p with 47 | | Or.inl pp => pp 48 | | Or.inr np => False.elim (h (λ p => np p)) 49 | 50 | theorem cnfItePos1 [h : Decidable c] : ¬ (ite c a b) ∨ ¬ c ∨ a := by 51 | apply orImplies 52 | intro hite 53 | have hite' := notNotElim hite 54 | match h with 55 | | isTrue _ => exact Or.inr hite' 56 | | isFalse hnc => exact Or.inl hnc 57 | 58 | theorem cnfItePos2 [h : Decidable c] : ¬ (ite c a b) ∨ c ∨ b := by 59 | apply orImplies 60 | intro hite 61 | have hite' := notNotElim hite 62 | match h with 63 | | isFalse _ => exact Or.inr hite' 64 | | isTrue hc => exact Or.inl hc 65 | 66 | theorem cnfItePos3 [h : Decidable c] : ¬ (ite c a b) ∨ a ∨ b := by 67 | apply orImplies 68 | intro hite 69 | have hite' := notNotElim hite 70 | match h with 71 | | isFalse _ => exact Or.inr hite' 72 | | isTrue _ => exact Or.inl hite' 73 | 74 | theorem cnfIteNeg1 [h : Decidable c] : (ite c a b) ∨ ¬ c ∨ ¬ a := by 75 | apply orImplies 76 | intro hnite 77 | match h with 78 | | isTrue _ => exact Or.inr hnite 79 | | isFalse hnc => exact Or.inl hnc 80 | 81 | theorem cnfIteNeg2 [h : Decidable c] : (ite c a b) ∨ c ∨ ¬ b := by 82 | apply orImplies 83 | intro hnite 84 | match h with 85 | | isFalse _ => exact Or.inr hnite 86 | | isTrue hc => exact Or.inl hc 87 | 88 | theorem cnfIteNeg3 [h : Decidable c] : (ite c a b) ∨ ¬ a ∨ ¬ b := by 89 | apply orImplies 90 | intro hnite 91 | match h with 92 | | isFalse _ => exact Or.inr hnite 93 | | isTrue _ => exact Or.inl hnite 94 | 95 | theorem scopes : ∀ {ps : List Prop} {q : Prop}, impliesN ps q → andN ps → q := 96 | by intros ps q h 97 | match ps with 98 | | [] => intro; exact h 99 | | [p] => unfold andN 100 | unfold impliesN at h 101 | exact h 102 | | p₁::p₂::ps => unfold andN 103 | unfold impliesN at h 104 | intro ⟨hp₁, hps⟩ 105 | revert hps 106 | exact scopes (h hp₁) 107 | 108 | end Smt.Reconstruct.Builtin 109 | -------------------------------------------------------------------------------- /Smt/Reconstruct/Builtin/Rewrites.lean: -------------------------------------------------------------------------------- 1 | /- 2 | Copyright (c) 2021-2023 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: Abdalrhman Mohamed 6 | -/ 7 | 8 | namespace Smt.Reconstruct.Builtin 9 | 10 | -- https://github.com/cvc5/cvc5/blob/main/src/theory/builtin/rewrites 11 | 12 | -- ITE 13 | 14 | theorem ite_true_cond : ite True x y = x := rfl 15 | theorem ite_false_cond : ite False x y = y := rfl 16 | theorem ite_not_cond [h : Decidable c] : ite (¬c) x y = ite c y x := 17 | h.byCases (fun hc => if_pos hc ▸ if_neg (not_not_intro hc) ▸ rfl) 18 | (fun hnc => if_pos hnc ▸ if_neg hnc ▸ rfl) 19 | theorem ite_eq_branch [h : Decidable c] : ite c x x = x := 20 | h.byCases (if_pos · ▸ rfl) (if_neg · ▸ rfl) 21 | 22 | theorem ite_then_lookahead [h : Decidable c] : ite c (ite c x y) z = ite c x z := 23 | h.byCases (fun hc => if_pos hc ▸ if_pos hc ▸ if_pos hc ▸ rfl) 24 | (fun hc => if_neg hc ▸ if_neg hc ▸ rfl) 25 | theorem ite_else_lookahead [h : Decidable c] : ite c x (ite c y z) = ite c x z := 26 | h.byCases (fun hc => if_pos hc ▸ if_pos hc ▸ rfl) 27 | (fun hc => if_neg hc ▸ if_neg hc ▸ if_neg hc ▸ rfl) 28 | theorem ite_then_neg_lookahead [h : Decidable c] : ite c (ite (¬c) x y) z = ite c y z := 29 | h.byCases (fun hc => if_pos hc ▸ if_pos hc ▸ ite_not_cond (c := c) ▸ if_pos hc ▸ rfl) 30 | (fun hc => if_neg hc ▸ if_neg hc ▸ rfl) 31 | theorem ite_else_neg_lookahead [h : Decidable c] : ite c x (ite (¬c) y z) = ite c x y := 32 | h.byCases (fun hc => if_pos hc ▸ if_pos hc ▸ rfl) 33 | (fun hc => if_neg hc ▸ if_neg hc ▸ ite_not_cond (c := c) ▸ if_neg hc ▸ rfl) 34 | 35 | end Smt.Reconstruct.Builtin 36 | -------------------------------------------------------------------------------- /Smt/Reconstruct/Certified/Coe.lean: -------------------------------------------------------------------------------- 1 | import Smt.Reconstruct.Certified.Types 2 | import Smt.Reconstruct.Certified.Term 3 | 4 | open Types 5 | open Coe 6 | open proof 7 | open sort 8 | open term 9 | 10 | instance {Δ : SEnvironment} : CoeOut (interpSort Δ (atom 1)) Prop where 11 | coe e := e 12 | 13 | instance {Δ : SEnvironment} : CoeOut (interpSort Δ boolSort) Prop where 14 | coe e := e 15 | 16 | def coe' : (A = B) → A → B 17 | | rfl, a => a 18 | -------------------------------------------------------------------------------- /Smt/Reconstruct/Certified/Defs.lean: -------------------------------------------------------------------------------- 1 | import Smt.Reconstruct.Certified.Env 2 | import Smt.Reconstruct.Certified.Utils 3 | import Smt.Reconstruct.Certified.Types 4 | import Smt.Reconstruct.Certified.Coe 5 | import Smt.Reconstruct.Certified.Term 6 | 7 | open Types 8 | open proof 9 | open term 10 | open sort 11 | 12 | set_option maxHeartbeats 500000 13 | 14 | @[simp] def combineProps : Interpretation → Interpretation → (Prop → Prop → Prop) → Interpretation 15 | | some ⟨ boolSort, k₁ ⟩, some ⟨ boolSort, k₂ ⟩, f => 16 | some ⟨ boolSort, λ Γ Δ => f (k₁ Γ Δ) (k₂ Γ Δ) ⟩ -- maybe we can use state monad to avoid propagating the environments manually? 17 | | _, _, _ => none 18 | 19 | @[simp] def interpTerm : term → Interpretation 20 | | term.const n s => some ⟨ s, λ Γ Δ => Γ n Δ s ⟩ 21 | | term.not t₁ => match interpTerm t₁ with 22 | | some ⟨ boolSort, k ⟩ => some ⟨ boolSort, λ Γ Δ => ¬ (k Γ Δ) ⟩ 23 | | _ => none 24 | | term.and t₁ t₂ => combineProps (interpTerm t₁) (interpTerm t₂) (λ p₁ p₂ => p₁ ∧ p₂) 25 | | term.or t₁ t₂ => combineProps (interpTerm t₁) (interpTerm t₂) (λ p₁ p₂ => p₁ ∨ p₂) 26 | | term.implies t₁ t₂ => combineProps (interpTerm t₁) (interpTerm t₂) (λ p₁ p₂ => p₁ → p₂) 27 | | term.xor t₁ t₂ => combineProps (interpTerm t₁) (interpTerm t₂) (λ p₁ p₂ => p₁ ≠ p₂) 28 | | term.eq t₁ t₂ => match interpTerm t₁, interpTerm t₂ with 29 | | some ⟨ s₁, k₁ ⟩, some ⟨ s₂, k₂ ⟩ => 30 | if r: s₂ = s₁ 31 | then some ⟨ boolSort, λ Γ Δ => 32 | let coercion := coe' (congrArg (interpSort Δ) r) 33 | k₁ Γ Δ = coercion (k₂ Γ Δ) ⟩ 34 | else none 35 | | _, _ => none 36 | | term.bot => some ⟨ boolSort, λ _ _ => False ⟩ 37 | | term.top => some ⟨ boolSort, λ _ _ => True ⟩ 38 | | term.app t₁ t₂ => match interpTerm t₁, interpTerm t₂ with 39 | | some ⟨ arrow s₁₁ s₁₂, k₁ ⟩, some ⟨ s₂, k₂ ⟩ => 40 | if r: s₂ = s₁₁ 41 | then some ⟨ s₁₂ , λ Γ Δ => 42 | let coercion := coe' (congrArg (interpSort Δ) r) 43 | k₁ Γ Δ (coercion (k₂ Γ Δ)) ⟩ 44 | else none 45 | | _, _ => none 46 | | _ => none 47 | 48 | @[simp] def validWith (Γ : Environment) (Δ : SEnvironment) (t : term) : Prop := 49 | match interpTerm t with 50 | | some ⟨ boolSort, k ⟩ => CoeOut.coe (k Γ Δ) 51 | | _ => False 52 | 53 | def followsFrom (t₁ t₂ : term) : Prop := 54 | ∀ {Γ : Environment} {Δ : SEnvironment}, validWith Γ Δ t₁ → validWith Γ Δ t₂ 55 | 56 | @[simp] def wellTyped (t : term) : Prop := isSome (interpTerm t) 57 | 58 | @[simp] def getInterp (t : term) (h : wellTyped t) : Σ (s : sort), Environment → (Δ : SEnvironment) → interpSort Δ s := 59 | match r: interpTerm t with 60 | | some ⟨ s, k ⟩ => ⟨ s, k ⟩ 61 | | none => by simp at h; rewrite [r] at h; exact (False.elim h) 62 | -------------------------------------------------------------------------------- /Smt/Reconstruct/Certified/Env.lean: -------------------------------------------------------------------------------- 1 | import Smt.Reconstruct.Certified.Types 2 | import Smt.Reconstruct.Certified.Coe 3 | import Smt.Reconstruct.Certified.Term 4 | 5 | open Types 6 | 7 | open proof 8 | open sort 9 | 10 | open Nat 11 | 12 | def defaultSEnvironment: SEnvironment := λ _ => ⟨ Nat , default ⟩ 13 | 14 | def defaultValue (Δ : SEnvironment) (s : sort) : interpSort Δ s := 15 | match s with 16 | | arrow _ s₂ => λ _ => defaultValue Δ s₂ 17 | | atom 0 => Sigma.snd (Δ 0) 18 | | atom 1 => False 19 | | atom (succ (succ i)) => Sigma.snd (Δ (succ (succ i))) 20 | | sort.undef => False 21 | | sort.array _ _ => False 22 | | sort.bv _ => False 23 | | sort.dep => False 24 | 25 | def defaultEnvironment: Environment := λ _ Δ s => defaultValue Δ s 26 | -------------------------------------------------------------------------------- /Smt/Reconstruct/Certified/EufRules.lean: -------------------------------------------------------------------------------- 1 | import Smt.Reconstruct.Certified.Defs 2 | import Smt.Reconstruct.Certified.Term 3 | 4 | open Types 5 | open proof 6 | open term 7 | open sort 8 | open Nat 9 | 10 | namespace Rules 11 | 12 | theorem refl : ∀ {t : term} {Γ : Environment} {Δ : SEnvironment}, 13 | wellTyped t → validWith Γ Δ (eq t t) 14 | | t, Γ, Δ, wTt => 15 | match r: interpTerm t with 16 | | some ⟨ atom 0 , _ ⟩ => by simp; rewrite [r]; exact rfl 17 | | some ⟨ atom 1 , _ ⟩ => by simp; rewrite [r]; exact rfl 18 | | some ⟨ atom (succ (succ _)) , _ ⟩ => by simp; rewrite [r]; simp; exact rfl 19 | | some ⟨ sort.undef, _ ⟩ => by simp; rewrite [r]; exact rfl 20 | | some ⟨ sort.array _ _, _ ⟩ => by simp; rewrite [r]; simp; exact rfl 21 | | some ⟨ sort.bv _, _ ⟩ => by simp; rewrite [r]; simp; exact rfl 22 | | some ⟨ sort.arrow _ _, _ ⟩ => by simp; rewrite [r]; simp; exact rfl 23 | | some ⟨ sort.dep, _ ⟩ => by simp; rewrite [r]; exact rfl 24 | | none => by simp at wTt; rewrite [r] at wTt; exact False.elim wTt 25 | 26 | 27 | end Rules 28 | -------------------------------------------------------------------------------- /Smt/Reconstruct/Certified/ModusPonens.lean: -------------------------------------------------------------------------------- 1 | import Smt.Reconstruct.Certified.Env 2 | import Smt.Reconstruct.Certified.Rules 3 | import Smt.Reconstruct.Certified.Term 4 | 5 | open Types 6 | open Rules 7 | open proof 8 | open term 9 | open sort 10 | 11 | def p: term := const 1000 boolSort 12 | def q: term := const 1001 boolSort 13 | 14 | def mpDE: term := implies p (implies (implies p q) q) 15 | def notMpDE := not mpDE 16 | 17 | theorem th0: followsFrom notMpDE bot := λ lean_a0 => 18 | have lean_s0 := notImplies2 lean_a0 19 | have lean_s1 := notImplies1 lean_s0 20 | have lean_s2 := impliesElim lean_s1 21 | have lean_s4 := notImplies1 lean_a0 22 | have lean_s6 := R1 (conjunction lean_s2 lean_s4) 23 | have lean_s9 := notImplies2 lean_s0 24 | contradiction (conjunction lean_s9 lean_s6) 25 | 26 | noncomputable def envMP (x y : Prop) : Environment := λ i Δ s => 27 | match s with 28 | | boolSort => if i == 1000 then x else y 29 | | s' => defaultValue Δ s' 30 | 31 | theorem notMpDEFalse: ∀ {Γ: Environment} {Δ: SEnvironment}, ¬ validWith Γ Δ notMpDE := 32 | followsBot th0 33 | 34 | theorem mpDETrue: ∀ {Γ: Environment} {Δ: SEnvironment}, validWith Γ Δ mpDE := 35 | have mpDEIsBool : isBool mpDE := rfl 36 | interpNotTerm mpDEIsBool notMpDEFalse 37 | 38 | def curryModusPonens (x y: Prop) : Prop := x → (x → y) → y 39 | 40 | theorem mp: ∀ (x y: Prop), curryModusPonens x y := λ x y => 41 | @mpDETrue (envMP x y) defaultSEnvironment 42 | 43 | theorem mp' (x y : Prop) : curryModusPonens x y := λ h₁ h₂ => h₂ h₁ 44 | 45 | #check mp' 46 | -------------------------------------------------------------------------------- /Smt/Reconstruct/Certified/Types.lean: -------------------------------------------------------------------------------- 1 | import Smt.Reconstruct.Certified.Term 2 | 3 | open proof 4 | open sort 5 | open term 6 | 7 | namespace Types 8 | 9 | def SEnvironment : Type 1 := Nat → (Σ (A : Type) , A) 10 | 11 | def interpSort (Δ : SEnvironment) (s : sort) : Type := 12 | match s with 13 | | sort.arrow s₁ s₂ => (interpSort Δ s₁) → (interpSort Δ s₂) 14 | | sort.atom 1 => Prop 15 | | sort.atom n => Sigma.fst (Δ n) 16 | | _ => Prop 17 | 18 | -- takes the number of a constant and it's type, and returns its value 19 | def Environment : Type 1 := Nat → (Δ : SEnvironment) → (s : sort) → interpSort Δ s 20 | 21 | def Interpretation: Type 1 := 22 | Option (Σ (s : sort), Environment → (Δ : SEnvironment) → interpSort Δ s) 23 | 24 | end Types 25 | -------------------------------------------------------------------------------- /Smt/Reconstruct/Certified/Utils.lean: -------------------------------------------------------------------------------- 1 | @[simp] def isSome : Option α → Prop 2 | | none => False 3 | | some _ => True 4 | -------------------------------------------------------------------------------- /Smt/Reconstruct/Prop/Core.lean: -------------------------------------------------------------------------------- 1 | /- 2 | Copyright (c) 2021-2024 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: Abdalrhman Mohamed 6 | -/ 7 | 8 | /- abbrev Implies (p q : Prop) := p → q -/ 9 | 10 | inductive XOr (p q : Prop) : Prop where 11 | | inl : p → ¬q → XOr p q 12 | | inr : ¬p → q → XOr p q 13 | 14 | theorem XOr.elim {p q r : Prop} (h : XOr p q) (h₁ : p → ¬q → r) (h₂ : ¬p → q → r) : r := match h with 15 | | inl hp hnq => h₁ hp hnq 16 | | inr hnp hq => h₂ hnp hq 17 | 18 | theorem XOr.symm (h : XOr p q) : XOr q p := h.elim (flip inr) (flip inl) 19 | 20 | namespace XOr 21 | 22 | @[macro_inline] instance [dp : Decidable p] [dq : Decidable q] : Decidable (XOr p q) := 23 | match dp with 24 | | isTrue hp => 25 | match dq with 26 | | isTrue hq => isFalse (·.elim (fun _ hnq => hnq hq) (fun hnp _ => hnp hp)) 27 | | isFalse hnq => isTrue (.inl hp hnq) 28 | | isFalse hnp => 29 | match dq with 30 | | isTrue hq => isTrue (.inr hnp hq) 31 | | isFalse hnq => isFalse (·.elim (fun hp _ => hnp hp) (fun _ hq => hnq hq)) 32 | 33 | end XOr 34 | 35 | def andN : List Prop → Prop 36 | | [] => True 37 | | p :: [] => p 38 | | p :: ps => p ∧ andN ps 39 | 40 | @[simp] theorem andN_append : andN (ps ++ qs) = (andN ps ∧ andN qs) := by 41 | match ps, qs with 42 | | [], _ 43 | | [p], [] 44 | | [p], q :: qs => simp [andN] 45 | | p₁ :: p₂ :: ps, qs => 46 | rw [List.cons_append, andN, andN, andN_append, and_assoc] 47 | all_goals (intro h; nomatch h) 48 | 49 | @[simp] theorem andN_cons_append : andN (p :: ps) = (p ∧ andN ps) := by 50 | cases ps <;> simp only [andN, and_true] 51 | 52 | def orN : List Prop → Prop 53 | | [] => False 54 | | p :: [] => p 55 | | p :: qs => p ∨ orN qs 56 | 57 | @[simp] theorem orN_append : orN (ps ++ qs) = (orN ps ∨ orN qs) := by 58 | match ps, qs with 59 | | [], _ 60 | | [p], [] 61 | | [p], q :: qs => simp [orN] 62 | | p₁ :: p₂ :: ps, qs => 63 | rw [List.cons_append, orN, orN, orN_append, or_assoc] 64 | all_goals (intro h; nomatch h) 65 | 66 | @[simp] theorem orN_cons_append : orN (p :: ps) = (p ∨ orN ps) := by 67 | cases ps <;> simp only [orN, or_false] 68 | 69 | def impliesN (ps : List Prop) (q : Prop) : Prop := match ps with 70 | | [] => q 71 | | p :: ps => p → impliesN ps q 72 | 73 | def notN : List Prop → List Prop := List.map Not 74 | 75 | namespace Smt.Reconstruct.Prop 76 | 77 | theorem and_assoc_eq : ((p ∧ q) ∧ r) = (p ∧ (q ∧ r)) := by 78 | simp [and_assoc] 79 | 80 | theorem and_comm_eq : (p ∧ q) = (q ∧ p) := by 81 | simp [and_comm] 82 | 83 | theorem or_assoc_eq : ((a ∨ b) ∨ c) = (a ∨ (b ∨ c)) := by 84 | simp [or_assoc] 85 | 86 | theorem or_comm_eq : (p ∨ q) = (q ∨ p) := by 87 | simp [or_comm] 88 | 89 | instance : Std.Associative And := ⟨@and_assoc_eq⟩ 90 | 91 | instance : Std.Commutative And := ⟨@and_comm_eq⟩ 92 | 93 | instance : Std.IdempotentOp And := ⟨and_self⟩ 94 | 95 | instance : Std.LawfulIdentity And True where 96 | left_id := true_and 97 | right_id := and_true 98 | 99 | instance : Std.Associative Or := ⟨@or_assoc_eq⟩ 100 | 101 | instance : Std.Commutative Or := ⟨@or_comm_eq⟩ 102 | 103 | instance : Std.IdempotentOp Or := ⟨or_self⟩ 104 | 105 | instance : Std.LawfulIdentity Or False where 106 | left_id := false_or 107 | right_id := or_false 108 | 109 | end Smt.Reconstruct.Prop 110 | -------------------------------------------------------------------------------- /Smt/Reconstruct/Prop/LiftOrNToNeg.lean: -------------------------------------------------------------------------------- 1 | /- 2 | Copyright (c) 2021-2023 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: Tomaz Gomes Mascarenhas 6 | -/ 7 | 8 | import Lean 9 | 10 | import Smt.Reconstruct.Prop.Lemmas 11 | import Smt.Reconstruct.Prop.LiftOrNToImp 12 | import Smt.Reconstruct.Util 13 | 14 | namespace Smt.Reconstruct.Prop 15 | 16 | open Lean Elab Tactic Meta Expr 17 | open List 18 | 19 | theorem orFalse : ∀ {A : Prop}, A ∨ False → A := by 20 | intros a h 21 | cases h with 22 | | inl ha => exact ha 23 | | inr ff => exact False.elim ff 24 | /- 25 | removeFalse: takes an or-chain ended in `False` and proves 26 | that it implies the same or-chain without the `False` in the end. 27 | -/ 28 | syntax (name := removeFalse) "removeFalse" term "," term : tactic 29 | 30 | def removeFalseCore (val type : Expr) : MetaM Expr := do 31 | let length ← getLength type 32 | if length > 2 then 33 | let groupped ← groupPrefixCore val (length - 1) 34 | mkAppM ``orFalse #[groupped] 35 | else 36 | mkAppM ``orFalse #[val] 37 | 38 | @[tactic removeFalse] def evalRemoveFalse : Tactic := 39 | fun stx => withMainContext do 40 | let hyp ← elabTerm stx[1] none 41 | let hypType ← inferType hyp 42 | let answer ← removeFalseCore hyp hypType 43 | let mvar ← getMainGoal 44 | let type ← Meta.inferType answer 45 | let fname ← mkFreshId 46 | let (_, mvar') ← MVarId.intro1P $ ← mvar.assert fname type answer 47 | replaceMainGoal [mvar'] 48 | evalTactic (← `(tactic| exact $(mkIdent fname))) 49 | 50 | def liftOrNToNegMeta (val : Expr) : MetaM Expr := do 51 | let type ← inferType val 52 | let withoutFalse ← removeFalseCore val type 53 | let type' ← instantiateMVars (← inferType withoutFalse) 54 | let propsList ← collectPropsInOrChain type' 55 | let notPropsList := map notExpr propsList 56 | let propsListExpr := listExpr notPropsList $ Expr.sort Level.zero 57 | let deMorgan := mkApp (mkConst ``deMorgan₂) propsListExpr 58 | let modusTollens ← mkAppM ``mt #[deMorgan] 59 | let notNotHyp ← mkAppM ``notNotIntro #[withoutFalse] 60 | let answer := mkApp modusTollens notNotHyp 61 | return answer 62 | 63 | syntax (name := liftOrNToNeg) "liftOrNToNeg" term : tactic 64 | 65 | @[tactic liftOrNToNeg] def evalLiftOrNToNeg : Tactic := 66 | fun stx => withMainContext do 67 | trace[smt.profile.reconstruct] m!"[liftOrNToNeg] start time: {← IO.monoNanosNow}ns" 68 | let hyp ← elabTerm stx[1] none 69 | let answer ← liftOrNToNegMeta hyp 70 | let mvar ← getMainGoal 71 | let type ← Meta.inferType answer 72 | let fname ← mkFreshId 73 | let (_, mvar') ← MVarId.intro1P $ ← mvar.assert fname type answer 74 | replaceMainGoal [mvar'] 75 | evalTactic (← `(tactic| exact $(mkIdent fname))) 76 | trace[smt.profile.reconstruct] m!"[liftOrNToNeg] end time: {← IO.monoNanosNow}ns" 77 | 78 | end Smt.Reconstruct.Prop 79 | -------------------------------------------------------------------------------- /Smt/Reconstruct/Prop/PermutateOr.lean: -------------------------------------------------------------------------------- 1 | /- 2 | Copyright (c) 2021-2023 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: Tomaz Gomes Mascarenhas 6 | -/ 7 | 8 | import Lean 9 | 10 | import Smt.Reconstruct.Prop.Pull 11 | import Smt.Reconstruct.Util 12 | 13 | namespace Smt.Reconstruct.Prop 14 | 15 | open Lean Elab Tactic 16 | 17 | def permutateOrMeta (val : Expr) (perm : List Nat) 18 | (suffIdx : Option Nat) : MetaM Expr := do 19 | let type ← instantiateMVars (← Meta.inferType val) 20 | let suffIdx: Nat ← 21 | match suffIdx with 22 | | some i => pure i 23 | | none => pure $ (← getLength type) - 1 24 | let props ← collectPropsInOrChain' suffIdx type 25 | let props := permutateList props perm.reverse 26 | go props suffIdx val 27 | where go : List Expr → Nat → Expr → MetaM Expr 28 | | [], _, acc => return acc 29 | | e::es, suffIdx, acc => do 30 | let type ← Meta.inferType acc 31 | let pulled ← pullCore e acc type suffIdx 32 | go es suffIdx pulled 33 | 34 | def traceReorder (r : Except Exception Unit) : MetaM MessageData := 35 | return match r with 36 | | .ok _ => m!"{checkEmoji}" 37 | | _ => m!"{bombEmoji}" 38 | 39 | def reorder (mv : MVarId) (e : Expr) (is : Array Nat) (i : Option Nat) : MetaM Unit := withTraceNode `smt.reconstruct.reorder traceReorder do 40 | let answer ← permutateOrMeta e is.toList i 41 | mv.assign answer 42 | 43 | -- TODO: find a way to remove '?' without breaking the parser 44 | syntax (name := permutateOr) "permutateOr" term "," ("[" term,* "]")? ("," term)? : tactic 45 | 46 | def parsePermuteOr : Syntax → TacticM (List Nat × Option Nat) 47 | | `(tactic| permutateOr $_, [ $[$hs],* ]) => 48 | hs.toList.mapM stxToNat >>= λ li => return ⟨li, none⟩ 49 | | `(tactic| permutateOr $_, [ $[$hs],* ], $i) => 50 | hs.toList.mapM stxToNat >>= λ li => 51 | elabTerm i none >>= λ i' => 52 | return ⟨li, getNatLit? i'⟩ 53 | | _ => throwError "[permutateOr]: wrong usage" 54 | 55 | @[tactic permutateOr] def evalPermutateOr : Tactic := fun stx => 56 | withMainContext do 57 | trace[smt.debug] m!"[permutateOr] start time: {← IO.monoNanosNow}ns" 58 | let hyp ← elabTerm stx[1] none 59 | let ⟨hs, suffIdx⟩ ← parsePermuteOr stx 60 | let answer ← permutateOrMeta hyp hs suffIdx 61 | let mvar ← getMainGoal 62 | let type ← instantiateMVars (← Meta.inferType answer) 63 | let fname ← mkFreshId 64 | let (_, mvar') ← MVarId.intro1P $ ← mvar.assert fname type answer 65 | replaceMainGoal [mvar'] 66 | evalTactic (← `(tactic| exact $(mkIdent fname))) 67 | trace[smt.debug] m!"[permutateOr] end time: {← IO.monoNanosNow}ns" 68 | 69 | end Smt.Reconstruct.Prop 70 | -------------------------------------------------------------------------------- /Smt/Reconstruct/Rat/Rewrites.lean: -------------------------------------------------------------------------------- 1 | /- 2 | Copyright (c) 2021-2023 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: Abdalrhman Mohamed, Harun Khan 6 | -/ 7 | 8 | import Batteries.Data.Rat 9 | import Smt.Reconstruct.Rat.Lemmas 10 | 11 | namespace Smt.Reconstruct.Rat.Rewrite 12 | 13 | open Function 14 | 15 | -- https://github.com/cvc5/cvc5/blob/main/src/theory/arith/rewrites 16 | 17 | variable {t s x : Rat} 18 | 19 | theorem div_total_zero : t / 0 = 0 := 20 | Rat.div_def t 0 ▸ Rat.inv_zero.symm ▸ Rat.mul_zero t 21 | 22 | -- Eliminations 23 | 24 | theorem elim_gt : (t > s) = ¬(s ≥ t) := 25 | propext Rat.not_le.symm 26 | theorem elim_lt : (t < s) = ¬(t ≥ s) := 27 | propext Rat.not_le.symm 28 | theorem elim_leq : (t ≤ s) = (s ≥ t) := 29 | propext ge_iff_le 30 | 31 | theorem geq_norm1 : (t ≥ s) = (t - s ≥ 0) := by 32 | rw [←elim_leq, ←elim_leq, Rat.le_iff_sub_nonneg _ _] 33 | 34 | theorem eq_elim : (t = s) = (t ≥ s ∧ t ≤ s) := by 35 | apply propext 36 | rw [←elim_leq, And.comm] 37 | exact Rat.le_antisymm_iff _ _ 38 | 39 | theorem eq_conflict {t : Int} {c : Rat} (hcc : (↑c.floor = c) = False) : (t = c) = False := by 40 | simp only [eq_iff_iff, iff_false] 41 | intro htc 42 | have hcc : c.floor < c := ((Rat.le_iff_eq_or_lt c.floor c).mp (Rat.floor_le_self c)).resolve_left hcc.mp 43 | cases Int.lt_trichotomy t c.floor <;> rename_i htcf 44 | · have hntc : ↑t ≠ c := (Rat.lt_iff_le_and_ne.mp (Rat.lt_trans (Rat.cast_lt2 htcf) hcc)).right 45 | contradiction 46 | · cases htcf <;> rename_i htcf 47 | · simp_all 48 | · have h : c < t := by 49 | apply @Rat.lt_of_lt_of_le _ _ _ 50 | · exact Rat.self_le_floor_add_one c 51 | · exact Rat.cast_le2 htcf 52 | simp_all [Rat.lt_irrefl] 53 | 54 | theorem geq_tighten {t : Int} {c : Rat} {cc : Int} (hc : (↑c.floor = c) = False) (hcc : cc = Int.addN [c.floor, 1]) : (t ≥ c) = (t ≥ cc) := by 55 | have Int.floor_lt {z : Int} {a : Rat} : a.floor < z ↔ a < ↑z := sorry 56 | simp only [hcc, Int.addN, ge_iff_le, eq_iff_iff, Rat.le_iff_eq_or_lt, ← Int.floor_lt] 57 | have h : ↑t ≠ c := by simpa [Eq.symm] using eq_conflict hc 58 | apply Iff.intro <;> intro hct 59 | · have h := hct.resolve_left h.symm 60 | omega 61 | · omega 62 | 63 | -- Absolute value comparisons 64 | 65 | theorem abs_eq : (x.abs = y.abs) = (x = y ∨ x = -y) := by 66 | cases hx : decide (x < 0) <;> cases hy : decide (y < 0) <;> simp_all [Rat.abs] <;> sorry 67 | 68 | theorem abs_gt : (x.abs > y.abs) = ite (x ≥ 0) (ite (y ≥ 0) (x > y) (x > -y)) (ite (y ≥ 0) (-x > y) (-x > -y)) := by 69 | simp only [Rat.abs, gt_iff_lt, ge_iff_le, eq_iff_iff] <;> split <;> split <;> split <;> split <;> sorry 70 | 71 | -- ITE lifting 72 | 73 | theorem geq_ite_lift [h : Decidable c] {t s r : Rat} : (ite c t s ≥ r) = ite c (t ≥ r) (s ≥ r) := by 74 | cases h <;> simp_all 75 | 76 | theorem leq_ite_lift [h : Decidable c] {t s r : Rat} : (ite c t s ≤ r) = ite c (t ≤ r) (s ≤ r) := by 77 | cases h <;> simp_all 78 | 79 | -- min/max rules 80 | 81 | theorem min_lt1 : (ite (t < s) t s ≤ t) = True := by 82 | cases h : decide (t < s) <;> simp_all [Rat.not_lt.mp] 83 | 84 | theorem min_lt2 : (ite (t < s) t s ≤ s) = True := by 85 | cases h : decide (t < s) <;> simp_all [Rat.le_of_lt] 86 | 87 | theorem max_geq1 : (ite (t ≥ s) t s ≥ t) = True := by 88 | cases h : decide (t ≥ s) <;> simp_all [Rat.le_of_not_le] 89 | 90 | theorem max_geq2 : (ite (t ≥ s) t s ≥ s) = True := by 91 | cases h : decide (t ≥ s) <;> simp_all 92 | 93 | end Smt.Reconstruct.Rat.Rewrite 94 | -------------------------------------------------------------------------------- /Smt/Reconstruct/Real/Core.lean: -------------------------------------------------------------------------------- 1 | /- 2 | Copyright (c) 2021-2025 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: Abdalrhman Mohamed 6 | -/ 7 | 8 | import Mathlib.Data.Real.Basic 9 | 10 | namespace Real 11 | 12 | def addN : List Real → Real 13 | | [] => 0 14 | | [x] => x 15 | | x :: xs => x + addN xs 16 | 17 | @[simp] theorem addN_append : addN (xs ++ ys) = addN xs + addN ys := by 18 | match xs, ys with 19 | | [], _ 20 | | [x], [] 21 | | [x], y :: ys => simp [addN] 22 | | x₁ :: x₂ :: xs, ys => 23 | rw [List.cons_append, addN, addN, addN_append, add_assoc] 24 | all_goals (intro h; nomatch h) 25 | 26 | @[simp] theorem addN_cons_append : addN (x :: xs) = x + addN xs := by 27 | cases xs <;> simp only [addN, add_zero] 28 | 29 | def mulN : List Real → Real 30 | | [] => 1 31 | | [x] => x 32 | | x :: xs => x * mulN xs 33 | 34 | @[simp] theorem mulN_append : mulN (xs ++ ys) = mulN xs * mulN ys := by 35 | match xs, ys with 36 | | [], _ 37 | | [x], [] 38 | | [x], y :: ys => simp [mulN] 39 | | x₁ :: x₂ :: xs, ys => 40 | rw [List.cons_append, mulN, mulN, mulN_append, mul_assoc] 41 | all_goals (intro h; nomatch h) 42 | 43 | @[simp] theorem mulN_cons_append : mulN (x :: xs) = x * mulN xs := by 44 | cases xs <;> simp only [mulN, mul_one] 45 | 46 | end Real 47 | -------------------------------------------------------------------------------- /Smt/Reconstruct/Rewrite.lean: -------------------------------------------------------------------------------- 1 | /- 2 | Copyright (c) 2021-2024 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: Harun Khan 6 | -/ 7 | 8 | import Lean 9 | 10 | import Smt.Reconstruct.Prop.Rewrites 11 | 12 | theorem Eq.trans₂ {α} {a b c d : α} (h₁ : a = b) (h₂ : b = c) (h₃ : c = d) : a = d := 13 | h₁ ▸ h₂ ▸ h₃ 14 | 15 | namespace Smt.Reconstruct.Tactic 16 | 17 | open Lean Elab.Tactic 18 | open Smt.Reconstruct.Prop 19 | 20 | def traceSmtRw (r : Except Exception Unit) : MetaM MessageData := 21 | return match r with 22 | | .ok _ => m!"{checkEmoji}" 23 | | _ => m!"{bombEmoji}" 24 | 25 | def smtRw (mv : MVarId) (op : Expr) (id : Expr) (rr : Expr) (xss : Array (Array Expr)) : MetaM Unit := 26 | withTraceNode `smt.reconstruct.DSL_REWRITE.tac traceSmtRw do 27 | let xs := xss.map (fun xs => xs[1:].foldr (mkApp2 op) (xs[0]?.getD id)) 28 | let some (α, l, r) := (← mv.getType).eq? 29 | | throwError "[smt_rw] expected a top level equality with AC operator on lhs and/or rhs, got {← mv.getType}" 30 | let h₂ := mkAppN rr xs 31 | let some (_, ml, mr) := (← Meta.inferType h₂).eq? 32 | | throwError "[smt_rw] error applying rewrite rule to arguments: {mkAppN rr xs}" 33 | let lvl ← Meta.getLevel α 34 | let h₁ ← Meta.mkFreshExprMVar (mkApp3 (.const ``Eq [lvl]) α l ml) 35 | let h₃ ← Meta.mkFreshExprMVar (mkApp3 (.const ``Eq [lvl]) α mr r) 36 | Meta.AC.rewriteUnnormalizedRefl h₁.mvarId! 37 | Meta.AC.rewriteUnnormalizedRefl h₃.mvarId! 38 | mv.assign (mkApp8 (.const ``Eq.trans₂ [lvl]) α l ml mr r h₁ h₂ h₃) 39 | 40 | syntax inner := "[" term,* "]" 41 | syntax outer := "[" inner,* "]" 42 | 43 | syntax (name := smt_rw) "smt_rw" ident ident ident outer : tactic 44 | 45 | def parseInner : TSyntax ``inner → TacticM (Array Expr) 46 | | `(inner| [$ts,*]) => ts.getElems.mapM (elabTerm · none) 47 | | _ => throwError "[inner]: wrong usage" 48 | 49 | def parseOuter : TSyntax ``outer → TacticM (Array (Array Expr)) 50 | | `(outer| [$is,*]) => is.getElems.mapM parseInner 51 | | _ => throwError "[outer]: wrong usage" 52 | 53 | @[tactic smt_rw] def evalSMTRw : Tactic := fun stx => do 54 | let op ← elabTermForApply stx[1] 55 | let id ← elabTermForApply stx[2] 56 | let rr ← elabTermForApply stx[3] 57 | let xs ← parseOuter ⟨stx[4]⟩ 58 | let mv : MVarId ← getMainGoal 59 | smtRw mv op id rr xs 60 | 61 | end Smt.Reconstruct.Tactic 62 | -------------------------------------------------------------------------------- /Smt/Reconstruct/Timed.lean: -------------------------------------------------------------------------------- 1 | import Lean 2 | 3 | import Smt.Reconstruct.Prop.Lemmas 4 | 5 | open Lean Elab.Tactic 6 | open Meta hiding contradiction 7 | 8 | open Smt.Reconstruct 9 | 10 | open Smt.Reconstruct.Prop 11 | 12 | set_option trace.smt.profile.reconstruct true 13 | 14 | syntax (name := timed) "timed" term : tactic 15 | @[tactic timed] def evalTimed : Tactic := fun stx => do 16 | if stx[1].getNumArgs > 0 then 17 | trace[smt.profile.reconstruct] s!"[{stx[1][0]}] start time: {← IO.monoNanosNow}ns" 18 | else 19 | trace[smt.profile.reconstruct] s!"[{stx[1]}] start time: {← IO.monoNanosNow}ns" 20 | let tstx : Term := ⟨stx[1]⟩ 21 | evalTactic (← `(tactic| exact $tstx)) 22 | if stx[1].getNumArgs > 0 then 23 | trace[smt.profile.reconstruct] s!"[{stx[1][0]}] end time: {← IO.monoNanosNow}ns" 24 | else 25 | trace[smt.profile.reconstruct] s!"[{stx[1]}] end time: {← IO.monoNanosNow}ns" 26 | 27 | example (a b : Prop) : ¬ (a → b) → a := by 28 | intro h 29 | timed notImplies1 h 30 | 31 | example (a b : Prop) : a = a := by 32 | have _ : b = b := by timed rfl 33 | timed rfl 34 | -------------------------------------------------------------------------------- /Smt/Reconstruct/UF/Rewrites.lean: -------------------------------------------------------------------------------- 1 | /- 2 | Copyright (c) 2021-2023 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: Abdalrhman Mohamed 6 | -/ 7 | 8 | namespace Smt.Reconstruct.UF 9 | 10 | -- https://github.com/cvc5/cvc5/blob/main/src/theory/uf/rewrites 11 | 12 | variable {c : Prop} [h : Decidable c] {t s r : α} 13 | 14 | -- Equality 15 | 16 | theorem eq_refl : (t = t) = True := eq_self t 17 | theorem eq_symm : (t = s) = (s = t) := propext ⟨(· ▸ rfl), (· ▸ rfl)⟩ 18 | 19 | theorem eq_cond_deq (h : (s = r) = False) : ((t = s) = (t = r)) = (¬t = s ∧ ¬t = r) := 20 | propext <| Iff.intro 21 | (fun hsr => And.intro (fun hts => absurd (hts ▸ hsr ▸ hts) (of_eq_false h)) 22 | (fun htr => absurd (htr ▸ Eq.symm (hsr ▸ htr)) (of_eq_false h))) 23 | (fun hnsr => propext ⟨(absurd · hnsr.left), (absurd · hnsr.right)⟩) 24 | 25 | theorem eq_ite_lift : (ite c t s = r) = (ite c (t = r) (s = r)) := h.byCases 26 | (fun hc => if_pos hc ▸ if_pos hc ▸ rfl) 27 | (fun hnc => if_neg hnc ▸ if_neg hnc ▸ rfl) 28 | 29 | theorem distinct_binary_elim : (t ≠ s) = ¬(t = s) := rfl 30 | 31 | end Smt.Reconstruct.UF 32 | -------------------------------------------------------------------------------- /Smt/String.lean: -------------------------------------------------------------------------------- 1 | /- 2 | Copyright (c) 2021-2023 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: Abdalrhman Mohamed 6 | -/ 7 | 8 | import Smt.Translate.String 9 | -------------------------------------------------------------------------------- /Smt/Tactic.lean: -------------------------------------------------------------------------------- 1 | /- 2 | Copyright (c) 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 | -/ 6 | 7 | import Smt.Tactic.Smt 8 | -------------------------------------------------------------------------------- /Smt/Tactic/WHNFConfigurableRef.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 | -/ 7 | import Lean.Meta.Basic 8 | 9 | namespace Smt 10 | open Lean Meta 11 | 12 | /-- Controls which rules are applied during WHNF and reduction. -/ 13 | structure Reduction.Config where 14 | /-- When `zeta` is off, we add `let x : t := v; b` to the set of WHNFs. 15 | 16 | This requires an extra congruence rule in full reduction: 17 | ```lean 18 | Γ ⊢ t ⤳ t' Γ ⊢ v ⤳ v' Γ, x : t ⊢ b ⤳ b' 19 | -------------------------------------------- 20 | Γ ⊢ let x : t := v; b ⤳ let x : t' := v'; b' 21 | ``` -/ 22 | zeta : Bool := true 23 | beta : Bool := true 24 | eta : Bool := true 25 | iota : Bool := true 26 | proj : Bool := true 27 | /-- Add the WHNF rule 28 | ```lean 29 | Γ ⊢ ELIM[let x : t := v; b] ⤳ let x : t := v; ELIM[b] 30 | ``` 31 | where `ELIM` is either a function application or a projection. -/ 32 | letPushElim : Bool := false 33 | 34 | abbrev ReductionM := ReaderT Reduction.Config MetaM 35 | 36 | initialize whnfRef : IO.Ref (Expr → ReductionM Expr) ← 37 | IO.mkRef fun e => return e 38 | 39 | def whnf (e : Expr) : ReductionM Expr := do 40 | let fn ← whnfRef.get 41 | fn e 42 | 43 | def whnfD (e : Expr) : ReductionM Expr := 44 | withTransparency TransparencyMode.default <| whnf e 45 | 46 | end Smt 47 | -------------------------------------------------------------------------------- /Smt/Translate/Bool.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: Abdalrhman Mohamed, Wojciech Nawrocki 6 | -/ 7 | 8 | import Smt.Recognizers 9 | import Smt.Translate 10 | 11 | namespace Smt.Translate.Bool 12 | 13 | open Translator Term 14 | 15 | private def mkBool : Lean.Expr := 16 | .const ``Bool [] 17 | 18 | @[smt_translate] def translateType : Translator := fun e => match e with 19 | | .const ``Bool _ => return symbolT "Bool" 20 | | _ => return none 21 | 22 | @[smt_translate] def translateBool : Translator := fun e => do 23 | if let .const ``true _ := e then 24 | return symbolT "true" 25 | else if let .const ``false _ := e then 26 | return symbolT "false" 27 | else if let .some b := e.app1? ``not then 28 | return appT (symbolT "not") (← applyTranslators! b) 29 | else if let some (a, b) := e.app2? ``and then 30 | return mkApp2 (symbolT "and") (← applyTranslators! a) (← applyTranslators! b) 31 | else if let some (a, b) := e.app2? ``or then 32 | return mkApp2 (symbolT "or") (← applyTranslators! a) (← applyTranslators! b) 33 | else if let some (a, b) := e.app2? ``xor then 34 | return mkApp2 (symbolT "xor") (← applyTranslators! a) (← applyTranslators! b) 35 | else if let some (_, x, y) := e.beq? then 36 | return mkApp2 (symbolT "=") (← applyTranslators! x) (← applyTranslators! y) 37 | else if let some (_, x, y) := e.bne? then 38 | return mkApp2 (symbolT "distinct") (← applyTranslators! x) (← applyTranslators! y) 39 | else 40 | return none 41 | 42 | @[smt_translate] def translateProp : Translator := fun e => do 43 | if let some (.const ``Bool _, a, b) := e.eq? then 44 | return mkApp2 (symbolT "=") (← applyTranslators! a) (← applyTranslators! b) 45 | else 46 | return none 47 | 48 | end Smt.Translate.Bool 49 | -------------------------------------------------------------------------------- /Smt/Translate/Commands.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: Abdalrhman Mohamed, Tomaz Gomes Mascarenhas, Wojciech Nawrocki 6 | -/ 7 | 8 | import Smt.Data.Sexp 9 | import Smt.Dsl.Sexp 10 | import Smt.Translate.Term 11 | 12 | namespace Smt.Translate 13 | 14 | open Smt.Translate 15 | 16 | /-- An SMT-LIBv2 command that we can emit. -/ 17 | inductive Command where 18 | | setLogic (l : String) 19 | | setOption (k : String) (v : String) 20 | | declareSort (nm : String) (arity : Nat) 21 | | defineSort (nm : String) (ps : List Term) (tm : Term) 22 | | declare (nm : String) (st : Term) 23 | | defineFun (nm : String) (ps : List (String × Term)) (cod : Term) (tm : Term) (rec : Bool) 24 | | assert (tm : Term) 25 | | checkSat 26 | | getModel 27 | | getProof 28 | | exit 29 | 30 | namespace Command 31 | 32 | open Lean Term 33 | 34 | def defNat : Command := .defineSort "Nat" [] (symbolT "Int") 35 | 36 | def defNatSub : Command := 37 | .defineFun "Nat.sub" [("x", symbolT "Nat"), ("y", symbolT "Nat")] (symbolT "Nat") 38 | (Term.mkApp3 (symbolT "ite") 39 | (Term.mkApp2 (symbolT "<") (symbolT "x") (symbolT "y")) 40 | (literalT "0") 41 | (Term.mkApp2 (symbolT "-") (symbolT "x") (symbolT "y"))) 42 | false 43 | 44 | open ToSexp in 45 | protected def toSexp : Command → Sexp 46 | | .setLogic l => sexp!{(set-logic {l})} 47 | | .setOption k v => sexp!{(set-option {s!":{k}"} {v})} 48 | | .assert val => sexp!{(assert {val})} 49 | | .declare nm st@(arrowT ..) => 50 | let sts := arrowToList st 51 | sexp!{(declare-fun {quoteSymbol nm} (...{sts.dropLast.map toSexp}) {sts.getLast!})} 52 | | .declare nm st => sexp!{(declare-const {quoteSymbol nm} {st})} 53 | | .declareSort nm arity => 54 | sexp!{(declare-sort {quoteSymbol nm} {toString arity})} 55 | | .defineSort nm ps tm => 56 | sexp!{(define-sort {quoteSymbol nm} (...{ps.map toSexp}) {tm})} 57 | | .defineFun nm ps cod tm false => 58 | sexp!{(define-fun {quoteSymbol nm} {paramsToSexp ps} {cod} {tm})} 59 | | .defineFun nm ps cod tm true => 60 | sexp!{(define-fun-rec {quoteSymbol nm} {paramsToSexp ps} {cod} {tm})} 61 | | .checkSat => sexp!{(check-sat)} 62 | | .getModel => sexp!{(get-model)} 63 | | .getProof => sexp!{(get-proof)} 64 | | .exit => sexp!{(exit)} 65 | where 66 | arrowToList : Term → List Term 67 | | arrowT d c => d :: arrowToList c 68 | | s => [s] 69 | paramToSexp (p : String × Term) : Sexp := sexp!{({quoteSymbol p.fst} {p.snd})} 70 | paramsToSexp (ps : List (String × Term)) : Sexp := sexp!{(...{ps.map paramToSexp})} 71 | 72 | instance : ToSexp Command := ⟨Command.toSexp⟩ 73 | 74 | instance : ToString Command := ⟨toString ∘ ToSexp.toSexp⟩ 75 | 76 | def cmdsAsQuery : List Command → String := .intercalate "\n" ∘ (.map toString) 77 | 78 | end Smt.Translate.Command 79 | -------------------------------------------------------------------------------- /Smt/Translate/Int.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: Abdalrhman Mohamed, Wojciech Nawrocki 6 | -/ 7 | 8 | import Smt.Recognizers 9 | import Smt.Translate 10 | 11 | namespace Smt.Translate.Int 12 | 13 | open Translator Term 14 | 15 | private def mkInt : Lean.Expr := 16 | .const ``Int [] 17 | 18 | @[smt_translate] def translateType : Translator := fun e => match e with 19 | | .const ``Int _ => return symbolT "Int" 20 | | _ => return none 21 | 22 | @[smt_translate] def translateInt : Translator := fun e => do 23 | if let some n := e.natLitOf? mkInt then 24 | return literalT (toString n) 25 | else if let some x := e.negOf? mkInt then 26 | return appT (symbolT "-") (← applyTranslators! x) 27 | else if let some (x, y) := e.hAddOf? mkInt mkInt then 28 | return mkApp2 (symbolT "+") (← applyTranslators! x) (← applyTranslators! y) 29 | else if let some (x, y) := e.hSubOf? mkInt mkInt then 30 | return mkApp2 (symbolT "-") (← applyTranslators! x) (← applyTranslators! y) 31 | else if let some (x, y) := e.hMulOf? mkInt mkInt then 32 | return mkApp2 (symbolT "*") (← applyTranslators! x) (← applyTranslators! y) 33 | else if let some (x, y) := e.hDivOf? mkInt mkInt then 34 | return mkApp2 (symbolT "div") (← applyTranslators! x) (← applyTranslators! y) 35 | else if let some (x, y) := e.hModOf? mkInt mkInt then 36 | return mkApp2 (symbolT "mod") (← applyTranslators! x) (← applyTranslators! y) 37 | else 38 | return none 39 | 40 | @[smt_translate] def translateProp : Translator := fun e => do 41 | if let some (x, y) := e.ltOf? mkInt then 42 | return mkApp2 (symbolT "<") (← applyTranslators! x) (← applyTranslators! y) 43 | else if let some (x, y) := e.leOf? mkInt then 44 | return mkApp2 (symbolT "<=") (← applyTranslators! x) (← applyTranslators! y) 45 | else if let some (x, y) := e.geOf? mkInt then 46 | return mkApp2 (symbolT ">=") (← applyTranslators! x) (← applyTranslators! y) 47 | else if let some (x, y) := e.gtOf? mkInt then 48 | return mkApp2 (symbolT ">") (← applyTranslators! x) (← applyTranslators! y) 49 | else 50 | return none 51 | 52 | end Smt.Translate.Int 53 | -------------------------------------------------------------------------------- /Smt/Translate/Nat.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: Abdalrhman Mohamed, Wojciech Nawrocki 6 | -/ 7 | 8 | import Smt.Recognizers 9 | import Smt.Translate 10 | 11 | namespace Smt.Translate.Nat 12 | 13 | open Lean Expr 14 | open Translator Term 15 | 16 | private def mkNat : Lean.Expr := 17 | .const ``Nat [] 18 | 19 | @[smt_translate] def translateNat : Translator := fun e => do 20 | if let some n := e.natLitOf? mkNat then 21 | return literalT (toString n) 22 | else if let .const ``Nat.zero _ := e then 23 | return literalT "0" 24 | else if let some n := e.app1? ``Nat.succ then 25 | return mkApp2 (symbolT "+") (← applyTranslators! n) (literalT "1") 26 | else if let some (m, n) := e.hAddOf? mkNat mkNat then 27 | return mkApp2 (symbolT "+") (← applyTranslators! m) (← applyTranslators! n) 28 | else if let some (m, n) := e.hSubOf? mkNat mkNat then 29 | modify fun st => { st with depConstants := st.depConstants.insert ``Nat.sub } 30 | return mkApp2 (symbolT "Nat.sub") (← applyTranslators! m) (← applyTranslators! n) 31 | else if let some (m, n) := e.hMulOf? mkNat mkNat then 32 | return mkApp2 (symbolT "*") (← applyTranslators! m) (← applyTranslators! n) 33 | else if let some (m, n) := e.hDivOf? mkNat mkNat then 34 | return mkApp2 (symbolT "div") (← applyTranslators! m) (← applyTranslators! n) 35 | else if let some (m, n) := e.hModOf? mkNat mkNat then 36 | return mkApp2 (symbolT "mod") (← applyTranslators! m) (← applyTranslators! n) 37 | else 38 | return none 39 | 40 | @[smt_translate] def translateProp : Translator := fun e => do 41 | if let some (m, n) := e.ltOf? mkNat then 42 | return mkApp2 (symbolT "<") (← applyTranslators! m) (← applyTranslators! n) 43 | else if let some (m, n) := e.leOf? mkNat then 44 | return mkApp2 (symbolT "<=") (← applyTranslators! m) (← applyTranslators! n) 45 | else if let some (m, n) := e.geOf? mkNat then 46 | return mkApp2 (symbolT ">=") (← applyTranslators! m) (← applyTranslators! n) 47 | else if let some (m, n) := e.gtOf? mkNat then 48 | return mkApp2 (symbolT ">") (← applyTranslators! m) (← applyTranslators! n) 49 | else 50 | return none 51 | 52 | /- Translates quantified expressions over natural numbers for with versions that 53 | ensure the quantified variables are greater than or equal to 0. For 54 | example, given `∀ x : Nat, p(x)`, this method returns the expr 55 | `∀ x : Nat, x ≥ 0 → p(x)`. -/ 56 | @[smt_translate] def translateForalls : Translator 57 | | e@(forallE n t@(const ``Nat _) b bi) => do 58 | if e.isArrow then return none 59 | let tmB ← Meta.withLocalDecl n bi t (translateBody b) 60 | let tmGeqZero := Term.mkApp2 (symbolT ">=") (symbolT n.toString) (literalT "0") 61 | let tmProp := Term.mkApp2 (symbolT "=>") tmGeqZero tmB 62 | return forallT n.toString (symbolT "Int") tmProp 63 | | _ => return none 64 | where 65 | translateBody (b : Expr) (x : Expr) : TranslationM Term := do 66 | modify fun s => { s with localFVars := s.localFVars.insert x.fvarId! } 67 | let tmB ← applyTranslators! (b.instantiate #[x]) 68 | modify fun s => { s with localFVars := s.localFVars.erase x.fvarId! } 69 | return tmB 70 | 71 | end Smt.Translate.Nat 72 | -------------------------------------------------------------------------------- /Smt/Translate/Real.lean: -------------------------------------------------------------------------------- 1 | /- 2 | Copyright (c) 2021-2023 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: Abdalrhman Mohamed 6 | -/ 7 | 8 | import Smt.Recognizers 9 | import Mathlib.Data.Real.Archimedean 10 | 11 | import Smt.Translate 12 | 13 | namespace Smt.Translate.Rat 14 | 15 | open Translator Term 16 | 17 | private def mkReal : Lean.Expr := 18 | .const ``Real [] 19 | 20 | @[smt_translate] def translateType : Translator := fun e => match e with 21 | | .const ``Real _ => return symbolT "Real" 22 | | _ => return none 23 | 24 | @[smt_translate] def translateInt : Translator := fun e => do 25 | if let some x := e.intFloorOf? mkReal then 26 | return appT (symbolT "to_int") (← applyTranslators! x) 27 | else 28 | return none 29 | 30 | @[smt_translate] def translateReal : Translator := fun e => do 31 | if let some n := e.natLitOf? mkReal then 32 | return literalT s!"{n}.0" 33 | else if let some x := e.intCastOf? mkReal then 34 | return appT (symbolT "to_real") (← applyTranslators! x) 35 | else if let some x := e.negOf? mkReal then 36 | return appT (symbolT "-") (← applyTranslators! x) 37 | else if let some x := e.absOf? mkReal then 38 | return appT (symbolT "abs") (← applyTranslators! x) 39 | else if let some (x, y) := e.hAddOf? mkReal mkReal then 40 | return mkApp2 (symbolT "+") (← applyTranslators! x) (← applyTranslators! y) 41 | else if let some (x, y) := e.hSubOf? mkReal mkReal then 42 | return mkApp2 (symbolT "-") (← applyTranslators! x) (← applyTranslators! y) 43 | else if let some (x, y) := e.hMulOf? mkReal mkReal then 44 | return mkApp2 (symbolT "*") (← applyTranslators! x) (← applyTranslators! y) 45 | else if let some (x, y) := e.hDivOf? mkReal mkReal then 46 | return mkApp2 (symbolT "/") (← applyTranslators! x) (← applyTranslators! y) 47 | else 48 | return none 49 | 50 | @[smt_translate] def translateProp : Translator := fun e => do 51 | if let some (x, y) := e.ltOf? mkReal then 52 | return mkApp2 (symbolT "<") (← applyTranslators! x) (← applyTranslators! y) 53 | else if let some (x, y) := e.leOf? mkReal then 54 | return mkApp2 (symbolT "<=") (← applyTranslators! x) (← applyTranslators! y) 55 | else if let some (x, y) := e.geOf? mkReal then 56 | return mkApp2 (symbolT ">=") (← applyTranslators! x) (← applyTranslators! y) 57 | else if let some (x, y) := e.gtOf? mkReal then 58 | return mkApp2 (symbolT ">") (← applyTranslators! x) (← applyTranslators! y) 59 | else 60 | return none 61 | 62 | end Smt.Translate.Rat 63 | -------------------------------------------------------------------------------- /Smt/Translate/String.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: Abdalrhman Mohamed, Wojciech Nawrocki 6 | -/ 7 | 8 | import Smt.Recognizers 9 | import Smt.Translate 10 | 11 | namespace Smt.Translate.String 12 | 13 | open Lean Expr 14 | open Translator Term 15 | 16 | private def mkString : Expr := 17 | .const ``String [] 18 | 19 | @[smt_translate] def translateType : Translator := fun e => match e with 20 | | .const ``String _ => return symbolT "String" 21 | | _ => return none 22 | 23 | @[smt_translate] def translateInt : Translator := fun e => do 24 | if let some s := e.app1? ``String.length then 25 | return appT (symbolT "str.len") (← applyTranslators! s) 26 | else 27 | return none 28 | 29 | @[smt_translate] def translateString : Translator := fun e => do 30 | if let .lit (.strVal s) := e then 31 | return literalT s!"\"{s}\"" 32 | else if let some (z, x, y) := e.app3? ``String.replace then 33 | return mkApp3 (symbolT "str.replace_all") (← applyTranslators! z) (← applyTranslators! x) (← applyTranslators! y) 34 | else if let some (x, y) := e.hAppendOf? mkString mkString then 35 | return mkApp2 (symbolT "str.++") (← applyTranslators! x) (← applyTranslators! y) 36 | else 37 | return none 38 | 39 | @[smt_translate] def translateProp : Translator := fun e => do 40 | if let some (x, y) := e.ltOf? mkString then 41 | return mkApp2 (symbolT "str.<") (← applyTranslators! x) (← applyTranslators! y) 42 | else if let some (x, y) := e.gtOf? mkString then 43 | return mkApp2 (symbolT "str.>") (← applyTranslators! x) (← applyTranslators! y) 44 | else 45 | return none 46 | 47 | end Smt.Translate.String 48 | -------------------------------------------------------------------------------- /Smt/UF.lean: -------------------------------------------------------------------------------- 1 | /- 2 | Copyright (c) 2021-2023 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: Abdalrhman Mohamed 6 | -/ 7 | 8 | import Smt.Reconstruct.UF 9 | -------------------------------------------------------------------------------- /Test/Auto.expected: -------------------------------------------------------------------------------- https://raw.githubusercontent.com/ufmg-smite/lean-smt/db6a7caf8685b33897ceebc9159bb1e180e2c568/Test/Auto.expected -------------------------------------------------------------------------------- /Test/Auto.lean: -------------------------------------------------------------------------------- 1 | import Mathlib.Algebra.Group.Defs 2 | import Mathlib.Data.Set.Function 3 | 4 | import Smt 5 | import Smt.Auto 6 | 7 | set_option auto.native true 8 | 9 | attribute [rebind Auto.Native.solverFunc] Smt.smtSolverFunc 10 | 11 | example (x : α) : List.head? [x] = .some x := by 12 | have h₂ : ∀ (x : α) (ys : _), List.head? (x :: ys) = .some x := fun _ _ => rfl 13 | auto 14 | 15 | example (x y : α) : [x] ++ [y] = [x, y] := by 16 | -- Invoke definition unfolding 17 | have h : ∀ (x y : List α), x ++ y = x.append y := fun _ _ => rfl 18 | auto [h] d[List.append] 19 | 20 | variable [Group G] 21 | 22 | theorem inverse' : ∀ (a : G), a * a⁻¹ = 1 := by 23 | auto [mul_assoc, one_mul, inv_mul_cancel] 24 | 25 | theorem identity' : ∀ (a : G), a * 1 = a := by 26 | auto [mul_assoc, one_mul, inv_mul_cancel, inverse'] 27 | 28 | theorem unique_identity (e : G) : (∀ z, e * z = z) → e = 1 := by 29 | auto [mul_assoc, one_mul, inv_mul_cancel] 30 | 31 | -- TODO: pre-process Nat away 32 | -- example (x y : α) : List.get? [x, y] 1 = .some y := by 33 | -- auto d[List.get?] 34 | 35 | set_option auto.mono.mode "fol" 36 | set_option auto.mono.ciInstDefEq.mode "reducible" 37 | set_option auto.mono.termLikeDefEq.mode "reducible" 38 | 39 | variable [Nonempty α] [Nonempty β] {f : α → β} {s : Set α} {v u : Set β} 40 | 41 | example : f '' s ⊆ v ↔ s ⊆ f ⁻¹' v := by 42 | auto [Set.subset_def, Set.mem_image, Set.mem_preimage] 43 | 44 | example (h : Function.Injective f) : f ⁻¹' (f '' s) ⊆ s := by 45 | auto [Set.subset_def, Set.mem_preimage, Function.Injective.mem_set_image, h] 46 | 47 | example : f '' (f ⁻¹' u) ⊆ u := by 48 | auto [Set.subset_def, Set.mem_image, Set.mem_preimage] 49 | 50 | example (h : Function.Surjective f) : u ⊆ f '' (f ⁻¹' u) := by 51 | unfold Function.Surjective at h 52 | auto [Set.subset_def, Set.mem_image, Set.mem_preimage, h] 53 | -------------------------------------------------------------------------------- /Test/BitVec/Add.expected: -------------------------------------------------------------------------------- https://raw.githubusercontent.com/ufmg-smite/lean-smt/db6a7caf8685b33897ceebc9159bb1e180e2c568/Test/BitVec/Add.expected -------------------------------------------------------------------------------- /Test/BitVec/Add.lean: -------------------------------------------------------------------------------- 1 | import Smt 2 | 3 | open BitVec 4 | 5 | example {x y : BitVec 4} : ¬(x = 11#4 ∧ y = 11#4 ∧ x + y = 15#4) := by 6 | smt 7 | 8 | example {x y : BitVec 8} : ¬(x = 11#8 ∧ y = 11#8 ∧ x + y = 15#8) := by 9 | smt 10 | 11 | example {x y : BitVec 16} : ¬(x = 11#16 ∧ y = 11#16 ∧ x + y = 15#16) := by 12 | smt 13 | -------------------------------------------------------------------------------- /Test/BitVec/Shift.expected: -------------------------------------------------------------------------------- 1 | goal: x ++ y = zeroExtend 4 x <<< 2#2 ||| zeroExtend 4 y 2 | 3 | query: 4 | (set-logic ALL) 5 | (declare-const y (_ BitVec 2)) 6 | (declare-const x (_ BitVec 2)) 7 | (assert (not (= (concat x y) (bvor (bvshl ((_ zero_extend 2) x) #b10) ((_ zero_extend 2) y))))) 8 | (check-sat) 9 | Test/BitVec/Shift.lean:5:8: warning: declaration uses 'sorry' 10 | goal: x ++ y = zeroExtend 6 x <<< 3#2 ||| zeroExtend 6 y 11 | 12 | query: 13 | (set-logic ALL) 14 | (declare-const x (_ BitVec 3)) 15 | (declare-const y (_ BitVec 3)) 16 | (assert (not (= (concat x y) (bvor (bvshl ((_ zero_extend 3) x) #b11) ((_ zero_extend 3) y))))) 17 | (check-sat) 18 | Test/BitVec/Shift.lean:10:8: warning: declaration uses 'sorry' 19 | -------------------------------------------------------------------------------- /Test/BitVec/Shift.lean: -------------------------------------------------------------------------------- 1 | import Smt 2 | 3 | open BitVec 4 | 5 | theorem append_eq_shl_or_2 (x y : BitVec 2) 6 | : x ++ y = (x.zeroExtend 4 <<< 2#2) ||| y.zeroExtend 4 := by 7 | smt_show 8 | sorry 9 | 10 | theorem append_eq_shl_or_3 (x y : BitVec 3) 11 | : x ++ y = (x.zeroExtend 6 <<< 3#2) ||| y.zeroExtend 6 := by 12 | smt_show 13 | sorry 14 | -------------------------------------------------------------------------------- /Test/BitVec/XorComm.expected: -------------------------------------------------------------------------------- 1 | goal: x ^^^ y = y ^^^ x 2 | 3 | query: 4 | (set-logic ALL) 5 | (declare-const x (_ BitVec 2)) 6 | (declare-const y (_ BitVec 2)) 7 | (assert (not (= (bvxor x y) (bvxor y x)))) 8 | (check-sat) 9 | Test/BitVec/XorComm.lean:3:8: warning: declaration uses 'sorry' 10 | goal: x ^^^ y = y ^^^ x 11 | 12 | query: 13 | (set-logic ALL) 14 | (declare-const x (_ BitVec 8)) 15 | (declare-const y (_ BitVec 8)) 16 | (assert (not (= (bvxor x y) (bvxor y x)))) 17 | (check-sat) 18 | Test/BitVec/XorComm.lean:7:8: warning: declaration uses 'sorry' 19 | -------------------------------------------------------------------------------- /Test/BitVec/XorComm.lean: -------------------------------------------------------------------------------- 1 | import Smt 2 | 3 | theorem xor_comm_2 (x y : BitVec 2) : x ^^^ y = y ^^^ x := by 4 | smt_show 5 | sorry 6 | 7 | theorem xor_comm_4p4 (x y : BitVec (4+4)) : x ^^^ y = y ^^^ x := by 8 | smt_show 9 | sorry 10 | -------------------------------------------------------------------------------- /Test/Bool/Addition.expected: -------------------------------------------------------------------------------- 1 | goal: p = true → (p || q) = true 2 | 3 | query: 4 | (set-logic ALL) 5 | (declare-const p Bool) 6 | (declare-const q Bool) 7 | (assert (not (=> (= p true) (= (or p q) true)))) 8 | (check-sat) 9 | -------------------------------------------------------------------------------- /Test/Bool/Addition.lean: -------------------------------------------------------------------------------- 1 | import Smt 2 | 3 | theorem addition (p q : Bool) : p → p || q := by 4 | smt_show 5 | simp_all 6 | -------------------------------------------------------------------------------- /Test/Bool/Assoc.expected: -------------------------------------------------------------------------------- 1 | goal: (f p (f q r) == f (f p q) r) = true 2 | 3 | query: 4 | (set-logic ALL) 5 | (declare-const q Bool) 6 | (declare-const r Bool) 7 | (declare-fun f (Bool Bool) Bool) 8 | (declare-const p Bool) 9 | (assert (not (= (= (f p (f q r)) (f (f p q) r)) true))) 10 | (check-sat) 11 | Test/Bool/Assoc.lean:3:8: warning: declaration uses 'sorry' 12 | -------------------------------------------------------------------------------- /Test/Bool/Assoc.lean: -------------------------------------------------------------------------------- 1 | import Smt 2 | 3 | theorem assoc (f : Bool → Bool → Bool) (p q r : Bool) : 4 | f p (f q r) == f (f p q) r := by 5 | smt_show 6 | admit 7 | -------------------------------------------------------------------------------- /Test/Bool/Comm.expected: -------------------------------------------------------------------------------- 1 | goal: (f p q == f q p) = true 2 | 3 | query: 4 | (set-logic ALL) 5 | (declare-fun f (Bool Bool) Bool) 6 | (declare-const q Bool) 7 | (declare-const p Bool) 8 | (assert (not (= (= (f p q) (f q p)) true))) 9 | (check-sat) 10 | Test/Bool/Comm.lean:3:0: warning: declaration uses 'sorry' 11 | -------------------------------------------------------------------------------- /Test/Bool/Comm.lean: -------------------------------------------------------------------------------- 1 | import Smt 2 | 3 | example (f : Bool → Bool → Bool) (p q : Bool) : f p q == f q p := by 4 | smt_show 5 | admit 6 | -------------------------------------------------------------------------------- /Test/Bool/Cong.expected: -------------------------------------------------------------------------------- 1 | goal: (p == q) = true → (f p == f q) = true 2 | 3 | query: 4 | (set-logic ALL) 5 | (declare-const p Bool) 6 | (declare-fun f (Bool) Bool) 7 | (declare-const q Bool) 8 | (assert (not (=> (= (= p q) true) (= (= (f p) (f q)) true)))) 9 | (check-sat) 10 | -------------------------------------------------------------------------------- /Test/Bool/Cong.lean: -------------------------------------------------------------------------------- 1 | import Smt 2 | 3 | theorem cong (p q : Bool) (f : Bool → Bool) : p == q → f p == f q := by 4 | smt_show 5 | cases p <;> cases q <;> cases f true <;> cases f false <;> simp_all 6 | -------------------------------------------------------------------------------- /Test/Bool/Conjunction.expected: -------------------------------------------------------------------------------- 1 | goal: p = true → q = true → (p && q) = true 2 | 3 | query: 4 | (set-logic ALL) 5 | (declare-const q Bool) 6 | (declare-const p Bool) 7 | (assert (not (=> (= p true) (=> (= q true) (= (and p q) true))))) 8 | (check-sat) 9 | -------------------------------------------------------------------------------- /Test/Bool/Conjunction.lean: -------------------------------------------------------------------------------- 1 | import Smt 2 | 3 | theorem conjunction (p q : Bool) : p → q → p && q := by 4 | smt_show 5 | simp_all 6 | -------------------------------------------------------------------------------- /Test/Bool/DisjunctiveSyllogism.expected: -------------------------------------------------------------------------------- 1 | goal: (p || q) = true → (!p) = true → q = true 2 | 3 | query: 4 | (set-logic ALL) 5 | (declare-const q Bool) 6 | (declare-const p Bool) 7 | (assert (not (=> (= (or p q) true) (=> (= (not p) true) (= q true))))) 8 | (check-sat) 9 | -------------------------------------------------------------------------------- /Test/Bool/DisjunctiveSyllogism.lean: -------------------------------------------------------------------------------- 1 | import Smt 2 | 3 | theorem disjunctive_syllogism (p q : Bool) : p || q → !p → q := by 4 | smt_show 5 | intro hpq hnp 6 | cases p <;> simp_all 7 | -------------------------------------------------------------------------------- /Test/Bool/Exists'.expected: -------------------------------------------------------------------------------- 1 | goal: ∃ p, p = true 2 | 3 | query: 4 | (set-logic ALL) 5 | (assert (not (exists ((p Bool)) (= p true)))) 6 | (check-sat) 7 | -------------------------------------------------------------------------------- /Test/Bool/Exists'.lean: -------------------------------------------------------------------------------- 1 | import Smt 2 | 3 | theorem exists' : ∃ p : Bool, p := by 4 | smt_show 5 | exact Exists.intro true rfl 6 | -------------------------------------------------------------------------------- /Test/Bool/Falsum.expected: -------------------------------------------------------------------------------- 1 | goal: (!false) = true 2 | 3 | query: 4 | (set-logic ALL) 5 | (assert (not (= (not false) true))) 6 | (check-sat) 7 | -------------------------------------------------------------------------------- /Test/Bool/Falsum.lean: -------------------------------------------------------------------------------- 1 | import Smt 2 | 3 | theorem falsum : !false := by 4 | smt_show 5 | simp_all 6 | -------------------------------------------------------------------------------- /Test/Bool/HypotheticalSyllogism.expected: -------------------------------------------------------------------------------- 1 | goal: (p = true → q = true) → (q = true → r = true) → p = true → r = true 2 | 3 | query: 4 | (set-logic ALL) 5 | (declare-const r Bool) 6 | (declare-const q Bool) 7 | (declare-const p Bool) 8 | (assert (not (=> (=> (= p true) (= q true)) (=> (=> (= q true) (= r true)) (=> (= p true) (= r true)))))) 9 | (check-sat) 10 | -------------------------------------------------------------------------------- /Test/Bool/HypotheticalSyllogism.lean: -------------------------------------------------------------------------------- 1 | import Smt 2 | 3 | theorem hypothetical_syllogism (p q r : Bool) : (p → q) → (q → r) → p → r := by 4 | smt_show 5 | simp_all 6 | -------------------------------------------------------------------------------- /Test/Bool/ModusPonens'.expected: -------------------------------------------------------------------------------- 1 | goal: q = true 2 | 3 | query: 4 | (set-logic ALL) 5 | (declare-const p Bool) 6 | (declare-const q Bool) 7 | (assert (=> (= p true) (= q true))) 8 | (assert (= p true)) 9 | (assert (not (= q true))) 10 | (check-sat) 11 | -------------------------------------------------------------------------------- /Test/Bool/ModusPonens'.lean: -------------------------------------------------------------------------------- 1 | import Smt 2 | 3 | theorem modus_ponens' (p q : Bool) (hp : p) (hpq : p → q) : q := by 4 | smt_show [hp, hpq] 5 | simp_all 6 | -------------------------------------------------------------------------------- /Test/Bool/ModusPonens.expected: -------------------------------------------------------------------------------- 1 | goal: p = true → (p = true → q = true) → q = true 2 | 3 | query: 4 | (set-logic ALL) 5 | (declare-const p Bool) 6 | (declare-const q Bool) 7 | (assert (not (=> (= p true) (=> (=> (= p true) (= q true)) (= q true))))) 8 | (check-sat) 9 | -------------------------------------------------------------------------------- /Test/Bool/ModusPonens.lean: -------------------------------------------------------------------------------- 1 | import Smt 2 | 3 | theorem modus_ponens (p q : Bool) : p → (p → q) → q := by 4 | smt_show 5 | simp_all 6 | -------------------------------------------------------------------------------- /Test/Bool/ModusTollens.expected: -------------------------------------------------------------------------------- 1 | goal: (!q) = true → (p = true → q = true) → (!p) = true 2 | 3 | query: 4 | (set-logic ALL) 5 | (declare-const p Bool) 6 | (declare-const q Bool) 7 | (assert (not (=> (= (not q) true) (=> (=> (= p true) (= q true)) (= (not p) true))))) 8 | (check-sat) 9 | -------------------------------------------------------------------------------- /Test/Bool/ModusTollens.lean: -------------------------------------------------------------------------------- 1 | import Smt 2 | 3 | theorem modus_tollens (p q : Bool) : !q → (p → q) → !p := by 4 | smt_show 5 | intro hnq hpq 6 | cases p <;> simp_all 7 | -------------------------------------------------------------------------------- /Test/Bool/PropExt.expected: -------------------------------------------------------------------------------- 1 | goal: (p = true) = (q = true) → (p == q) = true 2 | 3 | query: 4 | (set-logic ALL) 5 | (declare-const q Bool) 6 | (declare-const p Bool) 7 | (assert (not (=> (= (= p true) (= q true)) (= (= p q) true)))) 8 | (check-sat) 9 | -------------------------------------------------------------------------------- /Test/Bool/PropExt.lean: -------------------------------------------------------------------------------- 1 | import Smt 2 | 3 | theorem prop_ext (p q : Bool) : (p ↔ q) → p == q := by 4 | smt_show 5 | intro ⟨hpq, hqp⟩ 6 | cases p <;> cases q <;> simp_all 7 | -------------------------------------------------------------------------------- /Test/Bool/Refl.expected: -------------------------------------------------------------------------------- 1 | goal: (p == p) = true 2 | 3 | query: 4 | (set-logic ALL) 5 | (declare-const p Bool) 6 | (assert (not (= (= p p) true))) 7 | (check-sat) 8 | -------------------------------------------------------------------------------- /Test/Bool/Refl.lean: -------------------------------------------------------------------------------- 1 | import Smt 2 | 3 | example (p : Bool) : p == p := by 4 | smt_show 5 | cases p <;> simp_all 6 | -------------------------------------------------------------------------------- /Test/Bool/Resolution.expected: -------------------------------------------------------------------------------- 1 | goal: (p || q) = true → (!p || r) = true → (q || r) = true 2 | 3 | query: 4 | (set-logic ALL) 5 | (declare-const r Bool) 6 | (declare-const p Bool) 7 | (declare-const q Bool) 8 | (assert (not (=> (= (or p q) true) (=> (= (or (not p) r) true) (= (or q r) true))))) 9 | (check-sat) 10 | -------------------------------------------------------------------------------- /Test/Bool/Resolution.lean: -------------------------------------------------------------------------------- 1 | import Smt 2 | 3 | theorem resolution (p q r : Bool) : p || q → !p || r → q || r := by 4 | smt_show 5 | intro hpq 6 | intro hnpr 7 | cases p <;> cases r <;> simp_all 8 | -------------------------------------------------------------------------------- /Test/Bool/Simplification.expected: -------------------------------------------------------------------------------- 1 | goal: (p && q) = true → p = true 2 | 3 | query: 4 | (set-logic ALL) 5 | (declare-const p Bool) 6 | (declare-const q Bool) 7 | (assert (not (=> (= (and p q) true) (= p true)))) 8 | (check-sat) 9 | -------------------------------------------------------------------------------- /Test/Bool/Simplification.lean: -------------------------------------------------------------------------------- 1 | import Smt 2 | 3 | theorem simplification (p q : Bool) : p && q → p := by 4 | smt_show 5 | cases p <;> simp_all 6 | -------------------------------------------------------------------------------- /Test/Bool/Symm.expected: -------------------------------------------------------------------------------- 1 | goal: (p == q) = true → (q == p) = true 2 | 3 | query: 4 | (set-logic ALL) 5 | (declare-const q Bool) 6 | (declare-const p Bool) 7 | (assert (not (=> (= (= p q) true) (= (= q p) true)))) 8 | (check-sat) 9 | -------------------------------------------------------------------------------- /Test/Bool/Symm.lean: -------------------------------------------------------------------------------- 1 | import Smt 2 | 3 | example (p q : Bool) : p == q → q == p := by 4 | smt_show 5 | cases p <;> cases q <;> simp_all 6 | -------------------------------------------------------------------------------- /Test/Bool/Trans.expected: -------------------------------------------------------------------------------- 1 | goal: (p == q) = true → (q == r) = true → (p == r) = true 2 | 3 | query: 4 | (set-logic ALL) 5 | (declare-const r Bool) 6 | (declare-const p Bool) 7 | (declare-const q Bool) 8 | (assert (not (=> (= (= p q) true) (=> (= (= q r) true) (= (= p r) true))))) 9 | (check-sat) 10 | -------------------------------------------------------------------------------- /Test/Bool/Trans.lean: -------------------------------------------------------------------------------- 1 | import Smt 2 | 3 | example (p q r : Bool) : p == q → q == r → p == r := by 4 | smt_show 5 | cases p <;> cases q <;> cases r <;> simp_all 6 | -------------------------------------------------------------------------------- /Test/Bool/Triv'''.expected: -------------------------------------------------------------------------------- https://raw.githubusercontent.com/ufmg-smite/lean-smt/db6a7caf8685b33897ceebc9159bb1e180e2c568/Test/Bool/Triv'''.expected -------------------------------------------------------------------------------- /Test/Bool/Triv''.expected: -------------------------------------------------------------------------------- 1 | goal: (!p) = true → (!p) = true 2 | 3 | query: 4 | (set-logic ALL) 5 | (declare-const p Bool) 6 | (assert (not (=> (= (not p) true) (= (not p) true)))) 7 | (check-sat) 8 | -------------------------------------------------------------------------------- /Test/Bool/Triv''.lean: -------------------------------------------------------------------------------- 1 | import Smt 2 | 3 | theorem triv'' (p : Bool) : !p → !p := by 4 | smt_show 5 | simp_all 6 | -------------------------------------------------------------------------------- /Test/Bool/Triv'.expected: -------------------------------------------------------------------------------- 1 | goal: ∀ (p : Bool), p = true → p = true 2 | 3 | query: 4 | (set-logic ALL) 5 | (assert (not (forall ((p Bool)) (=> (= p true) (= p true))))) 6 | (check-sat) 7 | -------------------------------------------------------------------------------- /Test/Bool/Triv'.lean: -------------------------------------------------------------------------------- 1 | import Smt 2 | 3 | theorem triv': ∀ p : Bool, ∀ _ : p, p := by 4 | smt_show 5 | intro p 6 | simp_all 7 | -------------------------------------------------------------------------------- /Test/Bool/Triv.expected: -------------------------------------------------------------------------------- 1 | goal: p = true → p = true 2 | 3 | query: 4 | (set-logic ALL) 5 | (declare-const p Bool) 6 | (assert (not (=> (= p true) (= p true)))) 7 | (check-sat) 8 | -------------------------------------------------------------------------------- /Test/Bool/Triv.lean: -------------------------------------------------------------------------------- 1 | import Smt 2 | 3 | theorem triv (p : Bool) : p → p := by 4 | smt_show 5 | simp_all 6 | -------------------------------------------------------------------------------- /Test/Bool/Verum.expected: -------------------------------------------------------------------------------- 1 | goal: true = true 2 | 3 | query: 4 | (set-logic ALL) 5 | (assert (not (= true true))) 6 | (check-sat) 7 | -------------------------------------------------------------------------------- /Test/Bool/Verum.lean: -------------------------------------------------------------------------------- 1 | import Smt 2 | 3 | theorem verum : true := by 4 | smt_show 5 | simp_all 6 | -------------------------------------------------------------------------------- /Test/Concretize/Add.expected: -------------------------------------------------------------------------------- https://raw.githubusercontent.com/ufmg-smite/lean-smt/db6a7caf8685b33897ceebc9159bb1e180e2c568/Test/Concretize/Add.expected -------------------------------------------------------------------------------- /Test/Concretize/Add.lean: -------------------------------------------------------------------------------- 1 | -- import Smt.Tactic.Concretize 2 | 3 | -- def generalAdd [Add α] (a b : α) := a + b 4 | 5 | -- example : @generalAdd Int _ 3 3 = (6 : Int) := by 6 | -- concretize [generalAdd] 7 | -- rfl 8 | -------------------------------------------------------------------------------- /Test/Concretize/BitVec.expected: -------------------------------------------------------------------------------- https://raw.githubusercontent.com/ufmg-smite/lean-smt/db6a7caf8685b33897ceebc9159bb1e180e2c568/Test/Concretize/BitVec.expected -------------------------------------------------------------------------------- /Test/Concretize/BitVec.lean: -------------------------------------------------------------------------------- 1 | -- import Smt.Tactic.Concretize 2 | 3 | -- def polyAdd {w : Nat} : BitVec w → BitVec w → BitVec w := 4 | -- BitVec.xor 5 | -- def polyDouble {w : Nat} (x : BitVec w) : BitVec w := 6 | -- polyAdd x x 7 | 8 | -- example (x y : BitVec 2) : polyDouble (polyAdd x y) = polyDouble (polyAdd y x) := by 9 | -- concretize [polyDouble, polyAdd] 10 | -- trace_state 11 | -- sorry 12 | -------------------------------------------------------------------------------- /Test/EUF/Example.expected: -------------------------------------------------------------------------------- https://raw.githubusercontent.com/ufmg-smite/lean-smt/db6a7caf8685b33897ceebc9159bb1e180e2c568/Test/EUF/Example.expected -------------------------------------------------------------------------------- /Test/EUF/Example.lean: -------------------------------------------------------------------------------- 1 | import Smt 2 | 3 | example [Nonempty U] {f : U → U → U} {a b c d : U} : (a = b) → (c = d) → p1 ∧ True → (¬ p1) ∨ (p2 ∧ p3) → (¬ p3) ∨ (¬ (f a c = f b d)) → False := by 4 | smt 5 | 6 | example [Nonempty U] {f : U → U → U} {a b c d : U} 7 | (h0 : a = b) (h1 : c = d) (h2 : p1 ∧ True) (h3 : (¬ p1) ∨ (p2 ∧ p3)) (h4 : (¬ p3) ∨ (¬ (f a c = f b d))) : False := by 8 | smt [*] 9 | -------------------------------------------------------------------------------- /Test/EUF/Issue100.expected: -------------------------------------------------------------------------------- 1 | [smt] goal: True 2 | [smt] 3 | query: 4 | (set-logic ALL) 5 | (declare-sort value 0) 6 | (declare-sort round 0) 7 | (declare-sort node 0) 8 | (declare-fun upd_rel (node round round value) Bool) 9 | (declare-fun rel (node round round value) Bool) 10 | (assert (forall ((x node)) (forall ((x_1 round)) (forall ((x_2 round)) (forall ((x_3 value)) (= (upd_rel x x_1 x_2 x_3) (rel x x_1 x_2 x_3))))))) 11 | (assert (not true)) 12 | (check-sat) 13 | -------------------------------------------------------------------------------- /Test/EUF/Issue100.lean: -------------------------------------------------------------------------------- 1 | import Smt 2 | 3 | theorem funextEq {α β : Type} (f g : α → β) : 4 | (f = g) = ∀ x, f x = g x := by 5 | apply propext 6 | constructor 7 | { intro h ; simp only [h, implies_true] } 8 | { intro h ; apply funext h } 9 | 10 | set_option trace.smt true in 11 | theorem extracted {node round value : Type} 12 | (rel : node → round → round → value → Prop) 13 | (upd_rel : node → round → round → value → Prop) 14 | (hnext : upd_rel = rel) 15 | : True := by 16 | simp only [funextEq] at hnext 17 | -- After the simplification, `hnext` is pretty-printed as: 18 | -- (hnext : (∀ (x : node) (x_1 x_2 : round) (x_3 : value), upd_rel x x_1 x_2 x_3 = rel x x_1 x_2 x_3)) 19 | -- but the binders actually all have name `x` 20 | smt [hnext] 21 | -------------------------------------------------------------------------------- /Test/EUF/Issue126.expected: -------------------------------------------------------------------------------- https://raw.githubusercontent.com/ufmg-smite/lean-smt/db6a7caf8685b33897ceebc9159bb1e180e2c568/Test/EUF/Issue126.expected -------------------------------------------------------------------------------- /Test/EUF/Issue97.expected: -------------------------------------------------------------------------------- https://raw.githubusercontent.com/ufmg-smite/lean-smt/db6a7caf8685b33897ceebc9159bb1e180e2c568/Test/EUF/Issue97.expected -------------------------------------------------------------------------------- /Test/EUF/Issue97.lean: -------------------------------------------------------------------------------- 1 | import Smt 2 | 3 | theorem extracted {round : Type} [round_dec : DecidableEq round] 4 | (st_one_a : round → Prop) (st'_one_a : round → Prop) 5 | (hnext : ∃ r, (∀ (x : round), st'_one_a x = if x = r then True else st_one_a x)) 6 | : True := by 7 | smt [hnext] 8 | -------------------------------------------------------------------------------- /Test/EqnDef/Extract.expected: -------------------------------------------------------------------------------- https://raw.githubusercontent.com/ufmg-smite/lean-smt/db6a7caf8685b33897ceebc9159bb1e180e2c568/Test/EqnDef/Extract.expected -------------------------------------------------------------------------------- /Test/EqnDef/Extract.lean: -------------------------------------------------------------------------------- 1 | import Smt 2 | 3 | open Lean Meta Elab Tactic in 4 | elab "extract_def" i:ident : tactic => do 5 | let nm ← Elab.realizeGlobalConstNoOverloadWithInfo i 6 | let _ ← Smt.addEqnDefForConst nm 7 | 8 | def foo : Int := 10 9 | 10 | example : foo = 10 := by 11 | extract_def foo 12 | exact foo.def 13 | 14 | def bar (a b : Int) : Int := a + b 15 | 16 | example : bar 2 3 = 2 + 3 := by 17 | extract_def bar 18 | exact bar.def 2 3 19 | 20 | def baz (a : Int) : Int → Int := (a + ·) 21 | 22 | example : baz 2 3 = 2 + 3 := by 23 | extract_def baz 24 | exact baz.def 2 3 25 | 26 | def baw (a : Int) : Int → Int := Int.add a 27 | 28 | example : baw 2 3 = 2 + 3 := by 29 | extract_def baw 30 | exact baw.def 2 3 31 | -------------------------------------------------------------------------------- /Test/Examples/Demo.expected: -------------------------------------------------------------------------------- https://raw.githubusercontent.com/ufmg-smite/lean-smt/db6a7caf8685b33897ceebc9159bb1e180e2c568/Test/Examples/Demo.expected -------------------------------------------------------------------------------- /Test/Examples/Demo.lean: -------------------------------------------------------------------------------- 1 | import Smt 2 | 3 | variable [Nonempty G] (op : G → G → G) (inv : G → G) (e : G) 4 | 5 | axiom assoc : ∀ a b c, op a (op b c) = op (op a b) c 6 | axiom inverse : ∀ a, op (inv a) a = e 7 | axiom ident : ∀ a, op e a = a 8 | 9 | theorem inverse' : ∀ a, op a (inv a) = e := by 10 | smt [assoc op, inverse op inv e, ident op e] 11 | 12 | theorem identity' {inv : G → G} : ∀ a, op a e = a := by 13 | smt [assoc op, inverse op inv e, ident op e, inverse' op inv e] 14 | 15 | theorem unique_identity {inv : G → G} : ∀ e', (∀ z, op e' z = z) → e' = e := by 16 | smt [assoc op, inverse op inv e, ident op e] 17 | -------------------------------------------------------------------------------- /Test/Int/Arith.expected: -------------------------------------------------------------------------------- 1 | Test/Int/Arith.lean:3:47: error: unsolved goals 2 | n m : Int 3 | h : 0 < m 4 | ⊢ 0 < m → 5 | ¬n % m < m → 6 | ¬(m + -1 * if m = 0 then (fun x => x % 0) n else n % m) ≥ 1 → 7 | n = (if m = 0 then (fun x => x % 0) n else n % m) + m * (n / m) → 8 | m + -1 * n + m * (n / m) ≥ 1 → n % m = if m = 0 then (fun x => x % 0) n else n % m 9 | 10 | n m : Int 11 | h : 0 < m 12 | ⊢ 0 < m → 13 | ¬n % m < m → 14 | n = (if m = 0 then (fun x => x % 0) n else n % m) + m * (n / m) → 15 | ¬(m + -1 * if m = 0 then (fun x => x % 0) n else n % m) ≥ 1 → 16 | m + -1 * n + m * (n / m) ≥ 1 → 17 | m + -1 * n + m * (n / m) ≥ 1 → n % m = if m = 0 then (fun x => x % 0) n else n % m 18 | 19 | n m : Int 20 | h : 0 < m 21 | ⊢ 0 < m → 22 | ¬n % m < m → 23 | n % m = n - m * (n / m) ∧ 24 | (m > 0 → m * (n / m) ≤ n ∧ n < m * (n / m + 1)) ∧ (m < 0 → m * (n / m) ≤ n ∧ n < m * (n / m + -1)) 25 | 26 | n m : Int 27 | h : 0 < m 28 | ⊢ 0 < m → ¬n % m < m → n % m = if m = 0 then (fun x => x % 0) n else n % m 29 | -------------------------------------------------------------------------------- /Test/Int/Arith.lean: -------------------------------------------------------------------------------- 1 | import Smt 2 | 3 | example (n m : Int) (h : 0 < m) : n % m < m := by 4 | smt [h] 5 | 6 | example (n m k l : Int) : (n - m) * k + l = n*k - m*k + l := by 7 | smt 8 | 9 | example (n m k l : Int) (hN : n ≤ m) (hK : k ≤ l) : n + k ≤ m + l := by 10 | smt [hN, hK] 11 | -------------------------------------------------------------------------------- /Test/Int/Binders.expected: -------------------------------------------------------------------------------- 1 | goal: curryAdd a b = curryAdd b a 2 | 3 | query: 4 | (set-logic ALL) 5 | (define-fun curryAdd ((a._@.Test.Int.Binders._hyg.4 Int) (a._@.Test.Int.Binders._hyg.6 Int)) Int (+ a._@.Test.Int.Binders._hyg.4 a._@.Test.Int.Binders._hyg.6)) 6 | (declare-const a Int) 7 | (declare-const b Int) 8 | (assert (not (= (curryAdd a b) (curryAdd b a)))) 9 | (check-sat) 10 | Test/Int/Binders.lean:5:0: warning: declaration uses 'sorry' 11 | goal: partCurryAdd a b = partCurryAdd b a 12 | 13 | query: 14 | (set-logic ALL) 15 | (define-fun partCurryAdd ((a Int) (a._@.Test.Int.Binders._hyg.50 Int)) Int (+ a a._@.Test.Int.Binders._hyg.50)) 16 | (declare-const b Int) 17 | (declare-const a Int) 18 | (assert (not (= (partCurryAdd a b) (partCurryAdd b a)))) 19 | (check-sat) 20 | Test/Int/Binders.lean:11:0: warning: declaration uses 'sorry' 21 | goal: partCurryAdd' a b = partCurryAdd' b a 22 | 23 | query: 24 | (set-logic ALL) 25 | (define-fun |partCurryAdd'| ((a Int) (a._@.Init.Prelude._hyg.2075 Int)) Int (+ a a._@.Init.Prelude._hyg.2075)) 26 | (declare-const a Int) 27 | (declare-const b Int) 28 | (assert (not (= (|partCurryAdd'| a b) (|partCurryAdd'| b a)))) 29 | (check-sat) 30 | Test/Int/Binders.lean:15:0: warning: declaration uses 'sorry' 31 | goal: mismatchNamesAdd a b = mismatchNamesAdd b a 32 | 33 | query: 34 | (set-logic ALL) 35 | (define-fun mismatchNamesAdd ((a Int) (b Int)) Int (+ a b)) 36 | (declare-const b Int) 37 | (declare-const a Int) 38 | (assert (not (= (mismatchNamesAdd a b) (mismatchNamesAdd b a)))) 39 | (check-sat) 40 | Test/Int/Binders.lean:25:0: warning: declaration uses 'sorry' 41 | -------------------------------------------------------------------------------- /Test/Int/Binders.lean: -------------------------------------------------------------------------------- 1 | import Smt 2 | 3 | def curryAdd : Int → Int → Int := HAdd.hAdd 4 | 5 | example (a b : Int) : curryAdd a b = curryAdd b a := by 6 | smt_show [curryAdd] 7 | sorry 8 | 9 | def partCurryAdd (a : Int) : Int → Int := HAdd.hAdd a 10 | 11 | example (a b : Int) : partCurryAdd a b = partCurryAdd b a := by 12 | smt_show [partCurryAdd] 13 | sorry 14 | 15 | example (a b : Int) 16 | : let partCurryAdd' := fun a => HAdd.hAdd a; 17 | partCurryAdd' a b = partCurryAdd' b a := by 18 | intro partCurryAdd' 19 | smt_show [partCurryAdd'] 20 | sorry 21 | 22 | set_option linter.unusedVariables false in 23 | def mismatchNamesAdd : ∀ (a b : Int), Int := fun c d => c + d 24 | 25 | example (a b : Int) : mismatchNamesAdd a b = mismatchNamesAdd b a := by 26 | smt_show [mismatchNamesAdd] 27 | sorry 28 | -------------------------------------------------------------------------------- /Test/Int/Cong.expected: -------------------------------------------------------------------------------- https://raw.githubusercontent.com/ufmg-smite/lean-smt/db6a7caf8685b33897ceebc9159bb1e180e2c568/Test/Int/Cong.expected -------------------------------------------------------------------------------- /Test/Int/Cong.lean: -------------------------------------------------------------------------------- 1 | import Smt 2 | 3 | theorem cong (x y : Int) (f : Int → Int) : x = y → f x = f y := by 4 | smt 5 | -------------------------------------------------------------------------------- /Test/Int/DefineSort.expected: -------------------------------------------------------------------------------- 1 | goal: a.add b = b.add a 2 | 3 | query: 4 | (set-logic ALL) 5 | (define-sort MyInt () Int) 6 | (define-fun MyInt.add ((a MyInt) (b MyInt)) MyInt (+ a b)) 7 | (declare-const b MyInt) 8 | (declare-const a MyInt) 9 | (assert (not (= (MyInt.add a b) (MyInt.add b a)))) 10 | (check-sat) 11 | Test/Int/DefineSort.lean:6:0: warning: declaration uses 'sorry' 12 | -------------------------------------------------------------------------------- /Test/Int/DefineSort.lean: -------------------------------------------------------------------------------- 1 | import Smt 2 | 3 | def MyInt := Int 4 | def MyInt.add (a b : MyInt) : MyInt := @HAdd.hAdd Int Int _ _ a b 5 | 6 | example (a b : MyInt) : a.add b = b.add a := by 7 | smt_show [MyInt, MyInt.add] 8 | sorry 9 | -------------------------------------------------------------------------------- /Test/Int/Dvd.expected: -------------------------------------------------------------------------------- 1 | goal: dvd a c = true 2 | 3 | query: 4 | (set-logic ALL) 5 | (declare-fun dvd (Int Int) Bool) 6 | (assert (forall ((x Int)) (forall ((y Int)) (forall ((z Int)) (=> (= (dvd x y) true) (=> (= (dvd y z) true) (= (dvd x z) true))))))) 7 | (declare-const c Int) 8 | (declare-const b Int) 9 | (assert (= (dvd b c) true)) 10 | (declare-const a Int) 11 | (assert (= (dvd a b) true)) 12 | (assert (not (= (dvd a c) true))) 13 | (check-sat) 14 | goal: dvd c e = true 15 | 16 | query: 17 | (set-logic ALL) 18 | (declare-fun dvd (Int Int) Bool) 19 | (assert (forall ((x Int)) (forall ((y Int)) (forall ((z Int)) (=> (= (dvd x y) true) (=> (= (dvd y z) true) (= (dvd x z) true))))))) 20 | (declare-const d Int) 21 | (declare-const e Int) 22 | (assert (= (dvd d e) true)) 23 | (declare-const c Int) 24 | (assert (= (dvd c d) true)) 25 | (assert (not (= (dvd c e) true))) 26 | (check-sat) 27 | goal: dvd a e = true 28 | 29 | query: 30 | (set-logic ALL) 31 | (declare-fun dvd (Int Int) Bool) 32 | (assert (forall ((x Int)) (forall ((y Int)) (forall ((z Int)) (=> (= (dvd x y) true) (=> (= (dvd y z) true) (= (dvd x z) true))))))) 33 | (declare-const c Int) 34 | (declare-const e Int) 35 | (assert (= (dvd c e) true)) 36 | (declare-const a Int) 37 | (assert (= (dvd a c) true)) 38 | (assert (not (= (dvd a e) true))) 39 | (check-sat) 40 | -------------------------------------------------------------------------------- /Test/Int/Dvd.lean: -------------------------------------------------------------------------------- 1 | import Smt 2 | 3 | opaque dvd : Int → Int → Bool 4 | 5 | example 6 | (a b c d e : Int) 7 | (dvdax : ∀ x y z, dvd x y → dvd y z → dvd x z) 8 | (h1 : dvd a b) 9 | (h2 : dvd b c) 10 | (h3 : dvd c d) 11 | (h4 : dvd d e) : 12 | dvd a e := by 13 | have h5 : dvd a c := by 14 | smt_show [h1, h2, dvdax] 15 | exact dvdax _ _ _ h1 h2 16 | have h6 : dvd c e := by 17 | smt_show [h3, h4, dvdax] 18 | exact dvdax _ _ _ h3 h4 19 | smt_show [h5, h6, dvdax] 20 | exact dvdax _ _ _ h5 h6 21 | -------------------------------------------------------------------------------- /Test/Int/DvdTimeout.expected: -------------------------------------------------------------------------------- https://raw.githubusercontent.com/ufmg-smite/lean-smt/db6a7caf8685b33897ceebc9159bb1e180e2c568/Test/Int/DvdTimeout.expected -------------------------------------------------------------------------------- /Test/Int/DvdTimeout.lean: -------------------------------------------------------------------------------- 1 | import Smt 2 | 3 | opaque dvd : Int → Int → Prop 4 | 5 | example 6 | (a b c d : Int) 7 | (dvdax : ∀ x y z, dvd x y → dvd x z → dvd x (y + z)) 8 | (h1 : dvd a b) 9 | (h2 : dvd a c) 10 | (h3 : dvd a d) 11 | : dvd a (b + c + d) := by 12 | smt [dvdax, h1, h2, h3] 13 | -------------------------------------------------------------------------------- /Test/Int/ForallExists.expected: -------------------------------------------------------------------------------- 1 | Test/Int/ForallExists.lean:4:2: error: incorrect number of universe levels Nat.cast 2 | -------------------------------------------------------------------------------- /Test/Int/ForallExists.lean: -------------------------------------------------------------------------------- 1 | import Smt 2 | 3 | theorem forallExists : ∀ x : Nat, ∃ y : Int, x ≤ y := by 4 | smt_show 5 | intro x 6 | apply Exists.intro 7 | case w => exact Int.ofNat x 8 | case h => 9 | induction x with 10 | | zero => decide 11 | | succ x _ => 12 | simp only [LE.le, Int.le, HSub.hSub, Sub.sub, Int.sub, Neg.neg, 13 | Int.neg, Int.negOfNat, HAdd.hAdd, Add.add, Int.add] 14 | simp only [Int.subNatNat, Nat.sub_self, Int.NonNeg.mk] 15 | -------------------------------------------------------------------------------- /Test/Int/Let.expected: -------------------------------------------------------------------------------- 1 | goal: f z = z 2 | 3 | query: 4 | (set-logic ALL) 5 | (define-fun z () Int 10) 6 | (declare-fun f (Int) Int) 7 | (assert (= (f 10) 10)) 8 | (assert (not (= (f z) z))) 9 | (check-sat) 10 | goal: f y = y 11 | 12 | query: 13 | (set-logic ALL) 14 | (define-fun z () Int 10) 15 | (define-fun y () Int z) 16 | (declare-fun f (Int) Int) 17 | (assert (= (f 10) 10)) 18 | (assert (not (= (f y) y))) 19 | (check-sat) 20 | goal: z 3 = 10 21 | 22 | query: 23 | (set-logic ALL) 24 | (declare-fun f (Int) Int) 25 | (define-fun z ((x._@.Test.Int.Let._hyg.322 Int)) Int (f 10)) 26 | (assert (= (f 10) 10)) 27 | (assert (not (= (z 3) 10))) 28 | (check-sat) 29 | -------------------------------------------------------------------------------- /Test/Int/Let.lean: -------------------------------------------------------------------------------- 1 | import Smt 2 | 3 | variable (f : Int → Int) 4 | 5 | example (h : f 10 = 10) : let y := 10; f y = 10 := by 6 | smt [h] 7 | 8 | example (h : let y := 10; f y = 10) : f 10 = 10 := by 9 | smt [h] 10 | 11 | example (h : f 10 = 10) : f 10 = 10 := by 12 | let z : Int := 10 13 | have : 10 = z := rfl 14 | rw [this] 15 | smt_show [h, z] 16 | exact h 17 | 18 | example (h : f 10 = 10) : f 10 = 10 := by 19 | let z : Int := 10 20 | let y : Int := z 21 | have : 10 = y := rfl 22 | rw [this] 23 | smt_show [h, y, z] 24 | exact h 25 | 26 | example (h : f 10 = 10) : f 10 = 10 := by 27 | let z (_ : Int) : Int := f 10 28 | have : f 10 = z 3 := rfl 29 | rw [this] 30 | smt_show [h, z] 31 | exact h 32 | -------------------------------------------------------------------------------- /Test/Int/Neg.expected: -------------------------------------------------------------------------------- https://raw.githubusercontent.com/ufmg-smite/lean-smt/db6a7caf8685b33897ceebc9159bb1e180e2c568/Test/Int/Neg.expected -------------------------------------------------------------------------------- /Test/Int/Neg.lean: -------------------------------------------------------------------------------- 1 | import Smt 2 | 3 | theorem neg (x : Int) : - -x = x := by 4 | smt 5 | -------------------------------------------------------------------------------- /Test/Int/linarith.expected: -------------------------------------------------------------------------------- 1 | Test/Int/linarith.lean:39:5: warning: unused variable `h₁` 2 | note: this linter can be disabled with `set_option linter.unusedVariables false` 3 | Test/Int/linarith.lean:46:9: warning: unused variable `e` 4 | note: this linter can be disabled with `set_option linter.unusedVariables false` 5 | Test/Int/linarith.lean:65:0: warning: declaration uses 'sorry' 6 | Test/Int/linarith.lean:79:62: warning: unused variable `h3` 7 | note: this linter can be disabled with `set_option linter.unusedVariables false` 8 | Test/Int/linarith.lean:83:36: warning: unused variable `z` 9 | note: this linter can be disabled with `set_option linter.unusedVariables false` 10 | Test/Int/linarith.lean:84:5: warning: unused variable `h5` 11 | note: this linter can be disabled with `set_option linter.unusedVariables false` 12 | -------------------------------------------------------------------------------- /Test/Nat/Cong.expected: -------------------------------------------------------------------------------- 1 | goal: x = y → f x = f y 2 | 3 | query: 4 | (set-logic ALL) 5 | (define-sort Nat () Int) 6 | (declare-const y Nat) 7 | (assert (>= y 0)) 8 | (declare-const x Nat) 9 | (assert (>= x 0)) 10 | (declare-fun f (Nat) Nat) 11 | (assert (forall ((_uniq.33 Nat)) (=> (>= _uniq.33 0) (>= (f _uniq.33) 0)))) 12 | (assert (not (=> (= x y) (= (f x) (f y))))) 13 | (check-sat) 14 | -------------------------------------------------------------------------------- /Test/Nat/Cong.lean: -------------------------------------------------------------------------------- 1 | import Smt 2 | 3 | theorem cong (x y : Nat) (f : Nat → Nat) : x = y → f x = f y := by 4 | smt_show 5 | simp_all 6 | -------------------------------------------------------------------------------- /Test/Nat/Exists''.expected: -------------------------------------------------------------------------------- 1 | goal: ∃ x, x ≥ 0 2 | 3 | query: 4 | (set-logic ALL) 5 | (define-sort Nat () Int) 6 | (assert (not (exists ((x Nat)) (>= x 0)))) 7 | (check-sat) 8 | -------------------------------------------------------------------------------- /Test/Nat/Exists''.lean: -------------------------------------------------------------------------------- 1 | import Smt 2 | 3 | theorem exists'' : ∃ x : Nat, x ≥ 0 := by 4 | smt_show 5 | exact ⟨0, by decide⟩ 6 | -------------------------------------------------------------------------------- /Test/Nat/Exists'.expected: -------------------------------------------------------------------------------- 1 | goal: ∃ x, x = 1 2 | 3 | query: 4 | (set-logic ALL) 5 | (define-sort Nat () Int) 6 | (assert (not (exists ((x Nat)) (= x 1)))) 7 | (check-sat) 8 | -------------------------------------------------------------------------------- /Test/Nat/Exists'.lean: -------------------------------------------------------------------------------- 1 | import Smt 2 | 3 | theorem exists' : ∃ x : Nat, x = 1 := by 4 | smt_show 5 | exact ⟨1, rfl⟩ 6 | -------------------------------------------------------------------------------- /Test/Nat/Forall'.expected: -------------------------------------------------------------------------------- 1 | goal: ∀ (x : Nat), x ≥ 0 2 | 3 | query: 4 | (set-logic ALL) 5 | (assert (not (forall ((x Int)) (=> (>= x 0) (>= x 0))))) 6 | (check-sat) 7 | -------------------------------------------------------------------------------- /Test/Nat/Forall'.lean: -------------------------------------------------------------------------------- 1 | import Smt 2 | 3 | theorem forall' : ∀ x : Nat, x ≥ 0 := by 4 | smt_show 5 | simp [Nat.zero_le] 6 | -------------------------------------------------------------------------------- /Test/Nat/Max.expected: -------------------------------------------------------------------------------- 1 | goal: x ≤ x.max' y ∧ y ≤ x.max' y 2 | 3 | query: 4 | (set-logic ALL) 5 | (define-sort Nat () Int) 6 | (declare-fun |Nat.max'| (Nat Nat) Nat) 7 | (assert (forall ((_uniq.88 Nat)) (=> (>= _uniq.88 0) (forall ((_uniq.89 Nat)) (=> (>= _uniq.89 0) (>= (|Nat.max'| _uniq.88 _uniq.89) 0)))))) 8 | (declare-const y Nat) 9 | (assert (>= y 0)) 10 | (declare-const x Nat) 11 | (assert (>= x 0)) 12 | (assert (not (and (<= x (|Nat.max'| x y)) (<= y (|Nat.max'| x y))))) 13 | (check-sat) 14 | goal: x ≤ x.max' y ∧ y ≤ x.max' y 15 | 16 | query: 17 | (set-logic ALL) 18 | (define-sort Nat () Int) 19 | (define-fun |Nat.max'| ((x Nat) (y Nat)) Nat (ite (<= x y) y x)) 20 | (declare-const y Nat) 21 | (assert (>= y 0)) 22 | (declare-const x Nat) 23 | (assert (>= x 0)) 24 | (assert (not (and (<= x (|Nat.max'| x y)) (<= y (|Nat.max'| x y))))) 25 | (check-sat) 26 | -------------------------------------------------------------------------------- /Test/Nat/Max.lean: -------------------------------------------------------------------------------- 1 | import Smt 2 | 3 | def Nat.max' (x y : Nat) : Nat := if x ≤ y then y else x 4 | 5 | theorem Nat.not_le_of_reverse_le {m n : Nat} : ¬ m ≤ n → n ≤ m := fun hn => 6 | match Nat.le_total m n with 7 | | Or.inl h => absurd h hn 8 | | Or.inr h => h 9 | 10 | theorem Nat.max'_ge : ∀ x y : Nat, x ≤ max' x y ∧ y ≤ max' x y := by 11 | intro x y 12 | smt_show 13 | by_cases h : x ≤ y <;> simp [max', h] 14 | apply not_le_of_reverse_le h 15 | 16 | theorem Nat.max'_ge' : ∀ x y : Nat, x ≤ max' x y ∧ y ≤ max' x y := by 17 | intro x y 18 | smt_show [max'] 19 | by_cases h : x ≤ y <;> simp [max', h] 20 | apply not_le_of_reverse_le h 21 | -------------------------------------------------------------------------------- /Test/Nat/NeqZero.expected: -------------------------------------------------------------------------------- 1 | goal: ∀ (x : Nat), x ≠ 0 2 | 3 | query: 4 | (set-logic ALL) 5 | (assert (not (forall ((x Int)) (=> (>= x 0) (distinct x 0))))) 6 | (check-sat) 7 | Test/Nat/NeqZero.lean:3:8: warning: declaration uses 'sorry' 8 | goal: ∀ (x : Nat), x + 1 ≠ 0 9 | 10 | query: 11 | (set-logic ALL) 12 | (assert (not (forall ((x Int)) (=> (>= x 0) (distinct (+ x 1) 0))))) 13 | (check-sat) 14 | Test/Nat/NeqZero.lean:7:8: warning: declaration uses 'sorry' 15 | -------------------------------------------------------------------------------- /Test/Nat/NeqZero.lean: -------------------------------------------------------------------------------- 1 | import Smt 2 | 3 | theorem neq_zero : ∀ (x : Nat), x ≠ 0 := by 4 | smt_show 5 | admit 6 | 7 | theorem succ_neq_zero : ∀ (x : Nat), x + 1 ≠ 0 := by 8 | smt_show 9 | admit 10 | -------------------------------------------------------------------------------- /Test/Nat/Sum'.expected: -------------------------------------------------------------------------------- https://raw.githubusercontent.com/ufmg-smite/lean-smt/db6a7caf8685b33897ceebc9159bb1e180e2c568/Test/Nat/Sum'.expected -------------------------------------------------------------------------------- /Test/Nat/Sum'.lean: -------------------------------------------------------------------------------- 1 | -- import Smt 2 | 3 | -- def sum (n : Nat) : Nat := if n = 0 then 0 else n + sum (n - 1) 4 | 5 | -- theorem sum_formula : sum n = n * (n + 1) / 2 := by 6 | -- induction n with 7 | -- | zero => smt_show [sum]; rfl 8 | -- | succ n ih => 9 | -- smt_show [sum, ih] 10 | -- sorry 11 | -------------------------------------------------------------------------------- /Test/Nat/Triv'.expected: -------------------------------------------------------------------------------- 1 | goal: 0 + 1 = 1 2 | 3 | query: 4 | (set-logic ALL) 5 | (assert (not (= (+ 0 1) 1))) 6 | (check-sat) 7 | -------------------------------------------------------------------------------- /Test/Nat/Triv'.lean: -------------------------------------------------------------------------------- 1 | import Smt 2 | 3 | theorem triv' : 0 + 1 = 1 := by 4 | smt_show 5 | simp_all 6 | -------------------------------------------------------------------------------- /Test/Nat/Triv.expected: -------------------------------------------------------------------------------- 1 | goal: Nat.zero + Nat.zero.succ = Nat.zero.succ 2 | 3 | query: 4 | (set-logic ALL) 5 | (assert (not (= (+ 0 (+ 0 1)) (+ 0 1)))) 6 | (check-sat) 7 | -------------------------------------------------------------------------------- /Test/Nat/Triv.lean: -------------------------------------------------------------------------------- 1 | import Smt 2 | 3 | theorem triv : Nat.zero + Nat.succ Nat.zero = Nat.succ Nat.zero := by 4 | smt_show 5 | simp_all 6 | -------------------------------------------------------------------------------- /Test/Nat/ZeroSub'.expected: -------------------------------------------------------------------------------- 1 | Test/Nat/ZeroSub'.lean:3:0: warning: declaration uses 'sorry' 2 | -------------------------------------------------------------------------------- /Test/Nat/ZeroSub'.lean: -------------------------------------------------------------------------------- 1 | import Smt 2 | 3 | example : ∀ x, 0 - x = 0 := by 4 | intro x 5 | induction x with 6 | | zero => smt +trust 7 | | succ x ih => smt +trust [ih] 8 | -------------------------------------------------------------------------------- /Test/Nat/ZeroSub.expected: -------------------------------------------------------------------------------- 1 | goal: 0 - x = 0 2 | 3 | query: 4 | (set-logic ALL) 5 | (define-sort Nat () Int) 6 | (define-fun Nat.sub ((x Nat) (y Nat)) Nat (ite (< x y) 0 (- x y))) 7 | (declare-const x Nat) 8 | (assert (>= x 0)) 9 | (assert (not (= (Nat.sub 0 x) 0))) 10 | (check-sat) 11 | -------------------------------------------------------------------------------- /Test/Nat/ZeroSub.lean: -------------------------------------------------------------------------------- 1 | import Smt 2 | 3 | example : 0 - x = 0 := by 4 | smt_show 5 | induction x <;> simp_all [← Nat.sub_succ] 6 | -------------------------------------------------------------------------------- /Test/Prop/Addition.expected: -------------------------------------------------------------------------------- https://raw.githubusercontent.com/ufmg-smite/lean-smt/db6a7caf8685b33897ceebc9159bb1e180e2c568/Test/Prop/Addition.expected -------------------------------------------------------------------------------- /Test/Prop/Addition.lean: -------------------------------------------------------------------------------- 1 | import Smt 2 | 3 | theorem addition (p q : Prop) : p → p ∨ q := by 4 | smt 5 | -------------------------------------------------------------------------------- /Test/Prop/Assoc.expected: -------------------------------------------------------------------------------- 1 | Test/Prop/Assoc.lean:5:2: error: unable to prove goal, either it is false or you need to define more symbols with `smt [foo, bar]` 2 | -------------------------------------------------------------------------------- /Test/Prop/Assoc.lean: -------------------------------------------------------------------------------- 1 | import Smt 2 | 3 | theorem assoc (f : Prop → Prop → Prop) (p q r : Prop) : 4 | f p (f q r) = f (f p q) r := by 5 | smt 6 | admit 7 | -------------------------------------------------------------------------------- /Test/Prop/Comm.expected: -------------------------------------------------------------------------------- 1 | Test/Prop/Comm.lean:4:2: error: unable to prove goal, either it is false or you need to define more symbols with `smt [foo, bar]` 2 | -------------------------------------------------------------------------------- /Test/Prop/Comm.lean: -------------------------------------------------------------------------------- 1 | import Smt 2 | 3 | example (f : Prop → Prop → Prop) (p q : Prop) : f p q = f q p := by 4 | smt 5 | admit 6 | -------------------------------------------------------------------------------- /Test/Prop/Cong.expected: -------------------------------------------------------------------------------- https://raw.githubusercontent.com/ufmg-smite/lean-smt/db6a7caf8685b33897ceebc9159bb1e180e2c568/Test/Prop/Cong.expected -------------------------------------------------------------------------------- /Test/Prop/Cong.lean: -------------------------------------------------------------------------------- 1 | import Smt 2 | 3 | example (p q : Prop) (f : Prop → Prop) : p = q → f p = f q := by 4 | smt 5 | -------------------------------------------------------------------------------- /Test/Prop/Conjunction.expected: -------------------------------------------------------------------------------- https://raw.githubusercontent.com/ufmg-smite/lean-smt/db6a7caf8685b33897ceebc9159bb1e180e2c568/Test/Prop/Conjunction.expected -------------------------------------------------------------------------------- /Test/Prop/Conjunction.lean: -------------------------------------------------------------------------------- 1 | import Smt 2 | 3 | theorem conjunction (p q : Prop) : p → q → p ∧ q := by 4 | smt 5 | -------------------------------------------------------------------------------- /Test/Prop/DisjunctiveSyllogism.expected: -------------------------------------------------------------------------------- https://raw.githubusercontent.com/ufmg-smite/lean-smt/db6a7caf8685b33897ceebc9159bb1e180e2c568/Test/Prop/DisjunctiveSyllogism.expected -------------------------------------------------------------------------------- /Test/Prop/DisjunctiveSyllogism.lean: -------------------------------------------------------------------------------- 1 | import Smt 2 | 3 | theorem disjunctive_syllogism (p q : Prop) : p ∨ q → ¬p → q := by 4 | smt 5 | -------------------------------------------------------------------------------- /Test/Prop/Exists'.expected: -------------------------------------------------------------------------------- https://raw.githubusercontent.com/ufmg-smite/lean-smt/db6a7caf8685b33897ceebc9159bb1e180e2c568/Test/Prop/Exists'.expected -------------------------------------------------------------------------------- /Test/Prop/Exists'.lean: -------------------------------------------------------------------------------- 1 | import Smt 2 | 3 | theorem exists' : ∃ p : Prop, p := by 4 | smt 5 | -------------------------------------------------------------------------------- /Test/Prop/Falsum.expected: -------------------------------------------------------------------------------- https://raw.githubusercontent.com/ufmg-smite/lean-smt/db6a7caf8685b33897ceebc9159bb1e180e2c568/Test/Prop/Falsum.expected -------------------------------------------------------------------------------- /Test/Prop/Falsum.lean: -------------------------------------------------------------------------------- 1 | import Smt 2 | 3 | theorem falsum : ¬False := by 4 | smt 5 | -------------------------------------------------------------------------------- /Test/Prop/HypotheticalSyllogism.expected: -------------------------------------------------------------------------------- https://raw.githubusercontent.com/ufmg-smite/lean-smt/db6a7caf8685b33897ceebc9159bb1e180e2c568/Test/Prop/HypotheticalSyllogism.expected -------------------------------------------------------------------------------- /Test/Prop/HypotheticalSyllogism.lean: -------------------------------------------------------------------------------- 1 | import Smt 2 | 3 | theorem hypothetical_syllogism (p q r : Prop) : (p → q) → (q → r) → p → r := by 4 | smt 5 | -------------------------------------------------------------------------------- /Test/Prop/Iff.expected: -------------------------------------------------------------------------------- https://raw.githubusercontent.com/ufmg-smite/lean-smt/db6a7caf8685b33897ceebc9159bb1e180e2c568/Test/Prop/Iff.expected -------------------------------------------------------------------------------- /Test/Prop/Iff.lean: -------------------------------------------------------------------------------- 1 | import Smt 2 | 3 | variable {node : Type} [Nonempty node] 4 | 5 | example (a : node → Prop) (h : ∀ n, a n ↔ a n) : ∀ n, a n ↔ a n := by 6 | smt [h] 7 | -------------------------------------------------------------------------------- /Test/Prop/ModusPonens'.expected: -------------------------------------------------------------------------------- https://raw.githubusercontent.com/ufmg-smite/lean-smt/db6a7caf8685b33897ceebc9159bb1e180e2c568/Test/Prop/ModusPonens'.expected -------------------------------------------------------------------------------- /Test/Prop/ModusPonens'.lean: -------------------------------------------------------------------------------- 1 | import Smt 2 | 3 | theorem modus_ponens' (p q : Prop) (hp : p) (hpq : p → q) : q := by 4 | smt [hp, hpq] 5 | -------------------------------------------------------------------------------- /Test/Prop/ModusPonens.expected: -------------------------------------------------------------------------------- 1 | goal: q 2 | 3 | query: 4 | (set-logic ALL) 5 | (declare-const q Bool) 6 | (assert (not q)) 7 | (check-sat) 8 | -------------------------------------------------------------------------------- /Test/Prop/ModusPonens.lean: -------------------------------------------------------------------------------- 1 | import Smt 2 | 3 | theorem modus_ponens (p q : Prop) (hp : p) (hpq : p → q) : q := by 4 | smt +showQuery 5 | exact hpq hp 6 | 7 | example (p q : Prop) (hp : p) (hpq : p → q) : q := by 8 | smt [hp, hpq] 9 | 10 | example (p q : Prop) (hp : p) (hpq : p → q) : q := by 11 | smt [*] 12 | 13 | example (p q : Prop) (hp : p) (hpq : p → q) : q := by 14 | smt (timeout := .some 1) [*] 15 | -------------------------------------------------------------------------------- /Test/Prop/ModusTollens.expected: -------------------------------------------------------------------------------- https://raw.githubusercontent.com/ufmg-smite/lean-smt/db6a7caf8685b33897ceebc9159bb1e180e2c568/Test/Prop/ModusTollens.expected -------------------------------------------------------------------------------- /Test/Prop/ModusTollens.lean: -------------------------------------------------------------------------------- 1 | import Smt 2 | 3 | theorem modus_tollens (p q : Prop) : ¬q → (p → q) → ¬p := by 4 | smt 5 | -------------------------------------------------------------------------------- /Test/Prop/PropExt.expected: -------------------------------------------------------------------------------- https://raw.githubusercontent.com/ufmg-smite/lean-smt/db6a7caf8685b33897ceebc9159bb1e180e2c568/Test/Prop/PropExt.expected -------------------------------------------------------------------------------- /Test/Prop/PropExt.lean: -------------------------------------------------------------------------------- 1 | import Smt 2 | 3 | theorem prop_ext (p q : Prop) : (p ↔ q) → p = q := by 4 | smt 5 | -------------------------------------------------------------------------------- /Test/Prop/Refl.expected: -------------------------------------------------------------------------------- https://raw.githubusercontent.com/ufmg-smite/lean-smt/db6a7caf8685b33897ceebc9159bb1e180e2c568/Test/Prop/Refl.expected -------------------------------------------------------------------------------- /Test/Prop/Refl.lean: -------------------------------------------------------------------------------- 1 | import Smt 2 | 3 | example (p : Prop) : p = p := by 4 | smt 5 | -------------------------------------------------------------------------------- /Test/Prop/Resolution.expected: -------------------------------------------------------------------------------- https://raw.githubusercontent.com/ufmg-smite/lean-smt/db6a7caf8685b33897ceebc9159bb1e180e2c568/Test/Prop/Resolution.expected -------------------------------------------------------------------------------- /Test/Prop/Resolution.lean: -------------------------------------------------------------------------------- 1 | import Smt 2 | 3 | theorem resolution (p q r : Prop) : p ∨ q → ¬p ∨ r → q ∨ r := by 4 | smt 5 | -------------------------------------------------------------------------------- /Test/Prop/Simplification.expected: -------------------------------------------------------------------------------- https://raw.githubusercontent.com/ufmg-smite/lean-smt/db6a7caf8685b33897ceebc9159bb1e180e2c568/Test/Prop/Simplification.expected -------------------------------------------------------------------------------- /Test/Prop/Simplification.lean: -------------------------------------------------------------------------------- 1 | import Smt 2 | 3 | theorem simplification (p q : Prop) : p ∧ q → p := by 4 | smt 5 | -------------------------------------------------------------------------------- /Test/Prop/Symm.expected: -------------------------------------------------------------------------------- https://raw.githubusercontent.com/ufmg-smite/lean-smt/db6a7caf8685b33897ceebc9159bb1e180e2c568/Test/Prop/Symm.expected -------------------------------------------------------------------------------- /Test/Prop/Symm.lean: -------------------------------------------------------------------------------- 1 | import Smt 2 | 3 | example (p q : Prop) : p = q → q = p := by 4 | smt 5 | -------------------------------------------------------------------------------- /Test/Prop/Trans.expected: -------------------------------------------------------------------------------- https://raw.githubusercontent.com/ufmg-smite/lean-smt/db6a7caf8685b33897ceebc9159bb1e180e2c568/Test/Prop/Trans.expected -------------------------------------------------------------------------------- /Test/Prop/Trans.lean: -------------------------------------------------------------------------------- 1 | import Smt 2 | 3 | example (p q r : Prop) : p = q → q = r → p = r := by 4 | smt 5 | -------------------------------------------------------------------------------- /Test/Prop/Triv''.expected: -------------------------------------------------------------------------------- https://raw.githubusercontent.com/ufmg-smite/lean-smt/db6a7caf8685b33897ceebc9159bb1e180e2c568/Test/Prop/Triv''.expected -------------------------------------------------------------------------------- /Test/Prop/Triv''.lean: -------------------------------------------------------------------------------- 1 | import Smt 2 | 3 | theorem triv'' (p : Prop) : ¬p → ¬p := by 4 | smt 5 | -------------------------------------------------------------------------------- /Test/Prop/Triv'.expected: -------------------------------------------------------------------------------- https://raw.githubusercontent.com/ufmg-smite/lean-smt/db6a7caf8685b33897ceebc9159bb1e180e2c568/Test/Prop/Triv'.expected -------------------------------------------------------------------------------- /Test/Prop/Triv'.lean: -------------------------------------------------------------------------------- 1 | import Smt 2 | 3 | example : ∀ p : Prop, p → p := by 4 | smt 5 | -------------------------------------------------------------------------------- /Test/Prop/Triv.expected: -------------------------------------------------------------------------------- https://raw.githubusercontent.com/ufmg-smite/lean-smt/db6a7caf8685b33897ceebc9159bb1e180e2c568/Test/Prop/Triv.expected -------------------------------------------------------------------------------- /Test/Prop/Triv.lean: -------------------------------------------------------------------------------- 1 | import Smt 2 | 3 | theorem triv (p : Prop) : p → p := by 4 | smt 5 | -------------------------------------------------------------------------------- /Test/Prop/Verum.expected: -------------------------------------------------------------------------------- https://raw.githubusercontent.com/ufmg-smite/lean-smt/db6a7caf8685b33897ceebc9159bb1e180e2c568/Test/Prop/Verum.expected -------------------------------------------------------------------------------- /Test/Prop/Verum.lean: -------------------------------------------------------------------------------- 1 | import Smt 2 | 3 | theorem verum : True := by 4 | smt 5 | -------------------------------------------------------------------------------- /Test/Quant.expected: -------------------------------------------------------------------------------- https://raw.githubusercontent.com/ufmg-smite/lean-smt/db6a7caf8685b33897ceebc9159bb1e180e2c568/Test/Quant.expected -------------------------------------------------------------------------------- /Test/Quant.lean: -------------------------------------------------------------------------------- 1 | import Smt 2 | 3 | example {U : Type} [Nonempty U] {p : U → Prop} : 4 | (∀ x y z, p x ∧ p y ∧ p z) = ((∀ x, p x) ∧ (∀ y, p y) ∧ (∀ z, p z)) := by 5 | smt 6 | -------------------------------------------------------------------------------- /Test/Rat/linarith.expected: -------------------------------------------------------------------------------- 1 | Test/Rat/linarith.lean:106:9: warning: unused variable `a` 2 | note: this linter can be disabled with `set_option linter.unusedVariables false` 3 | Test/Rat/linarith.lean:106:13: warning: unused variable `c` 4 | note: this linter can be disabled with `set_option linter.unusedVariables false` 5 | Test/Rat/linarith.lean:116:9: warning: unused variable `a` 6 | note: this linter can be disabled with `set_option linter.unusedVariables false` 7 | Test/Rat/linarith.lean:116:13: warning: unused variable `c` 8 | note: this linter can be disabled with `set_option linter.unusedVariables false` 9 | -------------------------------------------------------------------------------- /Test/Real/Density.expected: -------------------------------------------------------------------------------- 1 | goal: x < z → ∃ y, x < y ∧ y < z 2 | 3 | query: 4 | (set-logic ALL) 5 | (declare-const z Real) 6 | (declare-const x Real) 7 | (assert (not (=> (< x z) (exists ((y Real)) (and (< x y) (< y z)))))) 8 | (check-sat) 9 | Test/Real/Density.lean:4:8: warning: declaration uses 'sorry' 10 | Test/Real/Density.lean:5:2: warning: 'smt_show' tactic does nothing 11 | note: this linter can be disabled with `set_option linter.unusedTactic false` 12 | -------------------------------------------------------------------------------- /Test/Real/Density.lean: -------------------------------------------------------------------------------- 1 | import Smt 2 | import Smt.Real 3 | 4 | theorem density (x z : Real) : x < z → ∃ y, x < y ∧ y < z := by 5 | smt_show 6 | sorry 7 | -------------------------------------------------------------------------------- /Test/Real/Inverse.expected: -------------------------------------------------------------------------------- 1 | goal: x ≠ 0 → ∃ y, x * y = 1 2 | 3 | query: 4 | (set-logic ALL) 5 | (declare-const x Real) 6 | (assert (not (=> (distinct x 0.0) (exists ((y Real)) (= (* x y) 1.0))))) 7 | (check-sat) 8 | Test/Real/Inverse.lean:4:8: warning: declaration uses 'sorry' 9 | Test/Real/Inverse.lean:5:2: warning: 'smt_show' tactic does nothing 10 | note: this linter can be disabled with `set_option linter.unusedTactic false` 11 | -------------------------------------------------------------------------------- /Test/Real/Inverse.lean: -------------------------------------------------------------------------------- 1 | import Smt 2 | import Smt.Real 3 | 4 | theorem inverse (x : Real) : x ≠ 0 → ∃ y, x * y = 1 := by 5 | smt_show 6 | sorry 7 | -------------------------------------------------------------------------------- /Test/Reconstruction/AndElim.expected: -------------------------------------------------------------------------------- https://raw.githubusercontent.com/ufmg-smite/lean-smt/db6a7caf8685b33897ceebc9159bb1e180e2c568/Test/Reconstruction/AndElim.expected -------------------------------------------------------------------------------- /Test/Reconstruction/AndElim.lean: -------------------------------------------------------------------------------- 1 | import Smt 2 | 3 | open Smt.Reconstruct 4 | 5 | example : A ∧ B ∧ C ∧ D → B := by 6 | intro h 7 | andElim h, 1 8 | -------------------------------------------------------------------------------- /Test/Reconstruction/Arith/ArithMulSign.expected: -------------------------------------------------------------------------------- 1 | Test/Reconstruction/Arith/ArithMulSign.lean:6:2: error: type mismatch 2 | _uniq.436 3 | has type 4 | a + 3 > Int.ofNat 0 → (a + 3) ^ 2 > 0 : Prop 5 | but is expected to have type 6 | a + 3 > 0 → (a + 3) ^ 2 > 0 : Prop 7 | Test/Reconstruction/Arith/ArithMulSign.lean:12:2: error: application type mismatch 8 | combineSigns₃ (powNegOdd _uniq.1842 (Nat.odd_iff.mpr rfl)) (powPos _uniq.1813) 9 | argument 10 | powPos _uniq.1813 11 | has type 12 | a ^ 2 > 0 : Prop 13 | but is expected to have type 14 | Pow.pow a 2 > 0 : Prop 15 | Test/Reconstruction/Arith/ArithMulSign.lean:21:2: error: application type mismatch 16 | combineSigns₂ (powPos _uniq.3537) 17 | argument 18 | powPos _uniq.3537 19 | has type 20 | b ^ 2 > 0 : Prop 21 | but is expected to have type 22 | Pow.pow b 2 > 0 : Prop 23 | Test/Reconstruction/Arith/ArithMulSign.lean:24:2: error: application type mismatch 24 | combineSigns₄ (powNegOdd _uniq.6270 (Nat.odd_iff.mpr rfl)) 25 | (castNeg (Pow.pow a 3) (powNegOdd _uniq.6211 (Nat.odd_iff.mpr rfl))) 26 | argument 27 | castNeg (Pow.pow a 3) (powNegOdd _uniq.6211 (Nat.odd_iff.mpr rfl)) 28 | has type 29 | Pow.pow a 3 < 0 : Prop 30 | but is expected to have type 31 | zToR (Pow.pow a 3) < 0 : Prop 32 | Test/Reconstruction/Arith/ArithMulSign.lean:33:2: error: application type mismatch 33 | combineSigns₂ (powPos _uniq.9015) 34 | argument 35 | powPos _uniq.9015 36 | has type 37 | b ^ 2 > 0 : Prop 38 | but is expected to have type 39 | Pow.pow b 2 > 0 : Prop 40 | -------------------------------------------------------------------------------- /Test/Reconstruction/Arith/ArithMulSign.lean: -------------------------------------------------------------------------------- 1 | import Smt.Reconstruct.Arith 2 | 3 | open Smt.Reconstruct.Arith 4 | 5 | example (a : Int) : (a + 3) > 0 → (a + 3) ^ 2 > 0 := by 6 | arithMulSign [(a + 3)], [1], [2] 7 | 8 | example (a : Int) : a < 0 → a ^ 3 < 0 := by 9 | arithMulSign [a], [-1], [3] 10 | 11 | example (a b : Real) : a > 0 → b < 0 → a ^ 2 * b ^ 3 < 0 := by 12 | arithMulSign [a,b], [1,-1], [2,3] 13 | 14 | example (a b c d e : Int) : 15 | a < 0 → 16 | b > 0 → 17 | c < 0 → 18 | d > 0 → 19 | e < 0 → 20 | a * (b ^ 2) * (c ^ 2) * (d ^ 4) * (e ^ 5) > 0 := by 21 | arithMulSign [a,b,c,d,e], [-1,1,-1,1,-1], [1,2,2,4,5] 22 | 23 | example (a : Int) (b : Real) : a < 0 → b < 0 → a ^ 3 * b ^ 3 > 0 := by 24 | arithMulSign [a,b], [-1,-1], [3,3] 25 | 26 | example (a e : Int) (b c d : Real) : 27 | a < 0 → 28 | b > 0 → 29 | c < 0 → 30 | d > 0 → 31 | e < 0 → 32 | a * (b ^ 2) * (c ^ 2) * (d ^ 4) * (e ^ 5) > 0 := by 33 | arithMulSign [a,b,c,d,e], [-1,1,-1,1,-1], [1,2,2,4,5] 34 | -------------------------------------------------------------------------------- /Test/Reconstruction/Arith/MulPosNeg.expected: -------------------------------------------------------------------------------- https://raw.githubusercontent.com/ufmg-smite/lean-smt/db6a7caf8685b33897ceebc9159bb1e180e2c568/Test/Reconstruction/Arith/MulPosNeg.expected -------------------------------------------------------------------------------- /Test/Reconstruction/Arith/MulPosNeg.lean: -------------------------------------------------------------------------------- 1 | import Smt.Reconstruct.Arith 2 | 3 | open Smt.Reconstruct.Arith 4 | 5 | example {a b c : Int} : c > 0 ∧ a < b → c * a < c * b := by 6 | arithMulPos [a,b,c], 0 7 | 8 | example {a b c : Real} : 0 > c ∧ a ≤ b → c * a ≥ c * b := by 9 | arithMulNeg [a,b,c], 1 10 | 11 | example {a b c : Int} : 0 < 2 * c ∧ a < b → (2 * c) * a < (2 * c) * b := by 12 | arithMulPos [a, b, 2 * c], 0 13 | 14 | /- example {a b : Int} : 0 < (2 : Int) ∧ a < b → 2 * a < 2 * b := by -/ 15 | /- arithMulPos [a,b,2], 0 -/ 16 | 17 | example {a b c : Int} : 0 < 2 * c ∧ a < b → (2 * c) * a < (2 * c) * b := by 18 | arithMulPos [a, b, 2 * c], 0 19 | 20 | example {a b c : Int} : 0 < c ∧ a < b → c * a < c * b := by 21 | arithMulPos [a, b, c], 0 22 | 23 | example {a b c : Real} : 0 < c ∧ a < b → c * a < c * b := by 24 | arithMulPos [a, b, c], 0 25 | 26 | example {a b : Real} {c : Int} : 0 < c ∧ a < b → c * a < c * b := by 27 | arithMulPos [a, b, c], 0 28 | 29 | example {a b : Real} {c : Int} : 0 < c ∧ a < b → c * a < c * b := by 30 | arithMulPos [a, b, c], 0 31 | 32 | example {a c : Real} {b : Int} : 0 < c ∧ a < b → c * a < c * b := by 33 | arithMulPos [a, b, c], 0 34 | 35 | example {a : Real} {b c : Int} : 0 < c ∧ a < b → c * a < c * b := by 36 | arithMulPos [a, b, c], 0 37 | 38 | example {b c : Real} {a : Int} : 0 < c ∧ a < b → c * a < c * b := by 39 | arithMulPos [a, b, c], 0 40 | 41 | example {b : Real} {a c : Int} : 0 < c ∧ a < b → c * a < c * b := by 42 | arithMulPos [a, b, c], 0 43 | 44 | example {c : Real} {a b : Int} : 0 < c ∧ a < b → c * a < c * b := by 45 | arithMulPos [a, b, c], 0 46 | 47 | example {a b c : Int} : 0 < c ∧ a < b → c * a < c * b := by 48 | arithMulPos [a, b, c], 0 49 | -------------------------------------------------------------------------------- /Test/Reconstruction/Arith/SumBounds.expected: -------------------------------------------------------------------------------- 1 | Test/Reconstruction/Arith/SumBounds.lean:7:2: error: application type mismatch 2 | sumBounds₁ h₃ 3 | argument 4 | h₃ 5 | has type 6 | @LT.lt ℕ instLTNat c f : Prop 7 | but is expected to have type 8 | @LT.lt ℕ Preorder.toLT c f : Prop 9 | Test/Reconstruction/Arith/SumBounds.lean:12:4: error: type mismatch 10 | _uniq.452 11 | has type 12 | w + c + b + a ≤ z + f + e + d : Prop 13 | but is expected to have type 14 | a + (b + (c + w)) ≤ d + (e + (f + z)) : Prop 15 | Test/Reconstruction/Arith/SumBounds.lean:16:2: error: type mismatch 16 | _uniq.1647 17 | has type 18 | c + b + ↑a < f + e + ↑d : Prop 19 | but is expected to have type 20 | ↑a + (b + c) < ↑d + (e + f) : Prop 21 | Test/Reconstruction/Arith/SumBounds.lean:20:2: error: application type mismatch 22 | sumBounds₁ h₃ 23 | argument 24 | h₃ 25 | has type 26 | @LT.lt ℕ instLTNat c f : Prop 27 | but is expected to have type 28 | @LT.lt ℕ Preorder.toLT c f : Prop 29 | -------------------------------------------------------------------------------- /Test/Reconstruction/Arith/SumBounds.lean: -------------------------------------------------------------------------------- 1 | import Smt.Reconstruct.Arith 2 | 3 | open Smt.Reconstruct.Arith 4 | 5 | example {a b c d e f : Nat} : a < d → b < e → c < f → a + (b + c) < d + (e + f) := by 6 | intros h₁ h₂ h₃ 7 | sum_bounds [h₁, h₂, h₃] 8 | 9 | example {a b c d e f w z : Int} : 10 | a ≤ d → b ≤ e → c ≤ f → w ≤ z → a + (b + (c + w)) ≤ d + (e + (f + z)) := by 11 | intros h₁ h₂ h₃ h₄ 12 | sum_bounds [h₁, h₂, h₃, h₄] 13 | 14 | example {a d : Int} {b c e f : Real} : a = d → b < e → c ≤ f → a + (b + c) < d + (e + f) := by 15 | intros h₁ h₂ h₃ 16 | sum_bounds [h₁, h₂, h₃] 17 | 18 | example {a b c d e f : Nat} : a < d → b < e → c < f → a + (b + c) < d + (e + f) := by 19 | intros h₁ h₂ h₃ 20 | sum_bounds [h₁, h₂, h₃] 21 | -------------------------------------------------------------------------------- /Test/Reconstruction/Arith/TightBounds.expected: -------------------------------------------------------------------------------- https://raw.githubusercontent.com/ufmg-smite/lean-smt/db6a7caf8685b33897ceebc9159bb1e180e2c568/Test/Reconstruction/Arith/TightBounds.expected -------------------------------------------------------------------------------- /Test/Reconstruction/Arith/TightBounds.lean: -------------------------------------------------------------------------------- 1 | import Smt.Reconstruct.Arith 2 | 3 | open Smt.Reconstruct.Arith 4 | 5 | example {a b : Int} : a < b → (a : Int) ≤ Int.ceil (Real.instIntCast.intCast b) - 1 := by 6 | intro h 7 | intTightUb h 8 | 9 | example {a : Int} : Real.instIntCast.intCast a < (7.2 : Real) → a ≤ Int.ceil (7.2 : Real) - 1 := by 10 | intro h 11 | intTightUb h 12 | 13 | example {a b : Int} : a > b → (a : Int) ≥ Int.floor (Real.instIntCast.intCast b) + 1 := by 14 | intro h 15 | intTightLb h 16 | -------------------------------------------------------------------------------- /Test/Reconstruction/Arith/Trichotomy.expected: -------------------------------------------------------------------------------- https://raw.githubusercontent.com/ufmg-smite/lean-smt/db6a7caf8685b33897ceebc9159bb1e180e2c568/Test/Reconstruction/Arith/Trichotomy.expected -------------------------------------------------------------------------------- /Test/Reconstruction/Arith/Trichotomy.lean: -------------------------------------------------------------------------------- 1 | import Smt.Reconstruct.Arith 2 | 3 | open Smt.Reconstruct.Arith 4 | 5 | example {a b : Nat} : a ≥ b → ¬ a = b → a > b := by 6 | intros h₁ h₂ 7 | trichotomy h₁, h₂ 8 | -------------------------------------------------------------------------------- /Test/Reconstruction/Factor.expected: -------------------------------------------------------------------------------- 1 | Test/Reconstruction/Factor.lean:7:3: error: unknown tactic 2 | Test/Reconstruction/Factor.lean:5:39: error: unsolved goals 3 | A B C : Prop 4 | h : A ∨ B ∨ C ∨ B 5 | ⊢ A ∨ B ∨ C 6 | Test/Reconstruction/Factor.lean:11:3: error: unknown tactic 7 | Test/Reconstruction/Factor.lean:9:31: error: unsolved goals 8 | A B : Prop 9 | h : A ∨ B ∨ B 10 | ⊢ A ∨ B 11 | Test/Reconstruction/Factor.lean:15:6: error: unknown tactic 12 | Test/Reconstruction/Factor.lean:14:2: error: unsolved goals 13 | A B C : Prop 14 | h : A ∨ A ∨ A ∨ A ∨ B ∨ A ∨ B ∨ A ∨ C ∨ B ∨ C ∨ B ∨ A 15 | ⊢ A ∨ B ∨ C 16 | Test/Reconstruction/Factor.lean:19:3: error: unknown tactic 17 | Test/Reconstruction/Factor.lean:17:61: error: unsolved goals 18 | A B C D : Prop 19 | h : (A ∨ B) ∨ C ∨ D ∨ C ∨ A ∨ B 20 | ⊢ (A ∨ B) ∨ C ∨ D 21 | Test/Reconstruction/Factor.lean:23:3: error: unknown tactic 22 | Test/Reconstruction/Factor.lean:21:51: error: unsolved goals 23 | A B C : Prop 24 | h : (A ∨ B ∨ C) ∨ A ∨ B ∨ C 25 | ⊢ A ∨ B ∨ C 26 | Test/Reconstruction/Factor.lean:28:6: error: unknown tactic 27 | Test/Reconstruction/Factor.lean:27:2: error: unsolved goals 28 | A B E F : Prop 29 | h : A ∨ B ∨ (E ∨ F) ∨ A ∨ B ∨ E ∨ F 30 | ⊢ A ∨ B ∨ E ∨ F 31 | Test/Reconstruction/Factor.lean:33:6: error: unknown tactic 32 | Test/Reconstruction/Factor.lean:32:2: error: unsolved goals 33 | A B C E F : Prop 34 | h : (A ∨ B ∨ C) ∨ (E ∨ F) ∨ (A ∨ B ∨ C) ∨ E ∨ F 35 | ⊢ (A ∨ B ∨ C) ∨ E ∨ F 36 | Test/Reconstruction/Factor.lean:37:3: error: unknown tactic 37 | Test/Reconstruction/Factor.lean:35:61: error: unsolved goals 38 | A B C D : Prop 39 | h : (A ∨ B) ∨ C ∨ D ∨ C ∨ A ∨ B 40 | ⊢ (A ∨ B) ∨ C ∨ D 41 | -------------------------------------------------------------------------------- /Test/Reconstruction/Factor.lean: -------------------------------------------------------------------------------- 1 | import Smt 2 | 3 | open Smt.Reconstruct 4 | 5 | example : A ∨ B ∨ C ∨ B → A ∨ B ∨ C := by 6 | intro h 7 | factor h 8 | 9 | example : A ∨ B ∨ B → A ∨ B := by 10 | intro h 11 | factor h 12 | 13 | example : A ∨ A ∨ A ∨ A ∨ B ∨ A ∨ B ∨ A ∨ C ∨ B ∨ C ∨ B ∨ A → A ∨ B ∨ C := 14 | by intro h 15 | factor h 16 | 17 | example : (A ∨ B) ∨ C ∨ D ∨ C ∨ (A ∨ B) → (A ∨ B) ∨ C ∨ D := by 18 | intro h 19 | factor h, 4 20 | 21 | example : (A ∨ B ∨ C) ∨ (A ∨ B ∨ C) → A ∨ B ∨ C := by 22 | intro h 23 | factor h, 1 24 | 25 | example : 26 | A ∨ B ∨ (E ∨ F) ∨ A ∨ B ∨ (E ∨ F) → A ∨ B ∨ (E ∨ F) := 27 | by intro h 28 | factor h, 5 29 | 30 | example : 31 | (A ∨ B ∨ C) ∨ (E ∨ F) ∨ (A ∨ B ∨ C) ∨ (E ∨ F) → (A ∨ B ∨ C) ∨ (E ∨ F) := 32 | by intro h 33 | factor h, 3 34 | 35 | example : (A ∨ B) ∨ C ∨ D ∨ C ∨ (A ∨ B) → (A ∨ B) ∨ C ∨ D := by 36 | intro h 37 | factor h, 4 38 | -------------------------------------------------------------------------------- /Test/Reconstruction/LiftOrNToImp.expected: -------------------------------------------------------------------------------- https://raw.githubusercontent.com/ufmg-smite/lean-smt/db6a7caf8685b33897ceebc9159bb1e180e2c568/Test/Reconstruction/LiftOrNToImp.expected -------------------------------------------------------------------------------- /Test/Reconstruction/LiftOrNToImp.lean: -------------------------------------------------------------------------------- 1 | import Smt 2 | 3 | open Smt.Reconstruct 4 | 5 | example : ¬ A ∨ ¬ B ∨ C ∨ D ∨ E → A ∧ B → C ∨ D ∨ E := by 6 | intro h 7 | liftOrNToImp h, 2 8 | 9 | example : ¬ A ∨ ¬ B ∨ C ∨ D ∨ E → ((A ∧ B) → C ∨ D ∨ E) := fun h => by 10 | liftOrNToImp h, 2 11 | 12 | example : 13 | let p1 := ¬ A ∨ ¬ B ∨ C ∨ D ∨ E 14 | let p2 := ((A ∧ B) → C ∨ D ∨ E) 15 | p1 → p2 := fun h => by 16 | liftOrNToImp h, 2 17 | -------------------------------------------------------------------------------- /Test/Reconstruction/LiftOrNToNeg.expected: -------------------------------------------------------------------------------- 1 | Test/Reconstruction/LiftOrNToNeg.lean:7:3: error: unknown tactic 2 | Test/Reconstruction/LiftOrNToNeg.lean:5:51: error: unsolved goals 3 | A B C : Prop 4 | h : ¬A ∨ B ∨ ¬C ∨ False 5 | ⊢ ¬A ∨ B ∨ ¬C 6 | Test/Reconstruction/LiftOrNToNeg.lean:11:3: error: unknown tactic 7 | Test/Reconstruction/LiftOrNToNeg.lean:9:63: error: unsolved goals 8 | A B C D : Prop 9 | h : ¬A ∨ ¬B ∨ ¬C ∨ ¬D ∨ False 10 | ⊢ ¬(A ∧ B ∧ C ∧ D) 11 | -------------------------------------------------------------------------------- /Test/Reconstruction/LiftOrNToNeg.lean: -------------------------------------------------------------------------------- 1 | import Smt 2 | 3 | open Smt.Reconstruct 4 | 5 | example : ¬ A ∨ B ∨ ¬ C ∨ False → ¬ A ∨ B ∨ ¬ C := by 6 | intro h 7 | removeFalse h, h₁ 8 | 9 | example : ¬ A ∨ ¬ B ∨ ¬ C ∨ ¬ D ∨ False → ¬ (A ∧ B ∧ C ∧ D) := by 10 | intro h 11 | liftOrNToNeg h 12 | -------------------------------------------------------------------------------- /Test/Reconstruction/NotOrElim.expected: -------------------------------------------------------------------------------- https://raw.githubusercontent.com/ufmg-smite/lean-smt/db6a7caf8685b33897ceebc9159bb1e180e2c568/Test/Reconstruction/NotOrElim.expected -------------------------------------------------------------------------------- /Test/Reconstruction/NotOrElim.lean: -------------------------------------------------------------------------------- 1 | import Smt 2 | 3 | open Smt.Reconstruct 4 | 5 | example : ¬ (A ∨ B ∨ C ∨ D) → ¬ C := by 6 | intro h 7 | notOrElim h, 2 8 | -------------------------------------------------------------------------------- /Test/Reconstruction/PermutateOr.expected: -------------------------------------------------------------------------------- https://raw.githubusercontent.com/ufmg-smite/lean-smt/db6a7caf8685b33897ceebc9159bb1e180e2c568/Test/Reconstruction/PermutateOr.expected -------------------------------------------------------------------------------- /Test/Reconstruction/PermutateOr.lean: -------------------------------------------------------------------------------- 1 | import Smt 2 | 3 | open Smt.Reconstruct 4 | 5 | example : A ∨ B ∨ C ∨ D ∨ E → A ∨ C ∨ D ∨ B ∨ E := by 6 | intro h 7 | permutateOr h, [0, 2, 3, 1, 4] 8 | 9 | example : A ∨ B ∨ C ∨ (D ∨ E ∨ F) → (D ∨ E ∨ F) ∨ B ∨ C ∨ A := by 10 | intro h 11 | permutateOr h, [3, 1, 2, 0], 3 12 | 13 | example : (D ∨ E ∨ F ∨ G) ∨ (K ∨ I) ∨ (A ∨ B ∨ C) → (A ∨ B ∨ C) ∨ (K ∨ I) ∨ (D ∨ E ∨ F ∨ G) := by 14 | intro h 15 | permutateOr h, [2, 1, 0], 2 16 | 17 | example : (D ∨ E ∨ F ∨ G) ∨ (A ∨ B ∨ C ∨ Z ∨ W ∨ J ∨ L) ∨ (K ∨ I) → (A ∨ B ∨ C ∨ Z ∨ W ∨ J ∨ L) ∨ (K ∨ I) ∨ (D ∨ E ∨ F ∨ G) := by 18 | intro h 19 | permutateOr h, [1, 2, 0], 2 20 | -------------------------------------------------------------------------------- /Test/Reconstruction/Pull.expected: -------------------------------------------------------------------------------- https://raw.githubusercontent.com/ufmg-smite/lean-smt/db6a7caf8685b33897ceebc9159bb1e180e2c568/Test/Reconstruction/Pull.expected -------------------------------------------------------------------------------- /Test/Reconstruction/Pull.lean: -------------------------------------------------------------------------------- 1 | import Smt 2 | 3 | open Smt.Reconstruct 4 | 5 | example : A ∨ B ∨ C ∨ D ∨ E → A ∨ B ∨ D ∨ C ∨ E := by 6 | intro h 7 | pullToMiddle 2, 3, h 8 | 9 | example : A ∨ B ∨ C ∨ D ∨ E ∨ F ∨ G → A ∨ B ∨ C ∨ F ∨ D ∨ E ∨ G := by 10 | intro h 11 | pullToMiddle 3, 5, h 12 | 13 | example : A ∨ B ∨ C ∨ D ∨ E ∨ F ∨ G ∨ H ∨ I ∨ J → 14 | A ∨ J ∨ B ∨ C ∨ D ∨ E ∨ F ∨ G ∨ H ∨ I := by 15 | intro h 16 | pullToMiddle 1, 9, h 17 | 18 | example : A ∨ B ∨ C ∨ D ∨ E ∨ F ∨ G ∨ H → A ∨ B ∨ C ∨ G ∨ D ∨ E ∨ F ∨ H := by 19 | intro h 20 | pullToMiddle 3, 6, h 21 | 22 | example : A ∨ B ∨ C ∨ D ∨ E ∨ F ∨ G ∨ H → A ∨ B ∨ E ∨ C ∨ D ∨ F ∨ G ∨ H := by 23 | intro h 24 | pullToMiddle 2, 4, h 25 | 26 | example : A ∨ B ∨ C ∨ D ∨ E ∨ F ∨ G ∨ H → E ∨ A ∨ B ∨ C ∨ D ∨ F ∨ G ∨ H := by 27 | intro h 28 | pullToMiddle 0, 4, h 29 | 30 | example : A ∨ B ∨ C ∨ D ∨ E ∨ F ∨ G ∨ H → A ∨ G ∨ B ∨ C ∨ D ∨ E ∨ F ∨ H := by 31 | intro h 32 | pullToMiddle 1, 6, h 33 | 34 | example : A ∨ B ∨ C ∨ D ∨ E ∨ F → A ∨ B ∨ C ∨ F ∨ D ∨ E := by 35 | intro h 36 | pullToMiddle 3, 5, h 37 | 38 | example : A ∨ B ∨ C ∨ D → A ∨ D ∨ B ∨ C := by 39 | intro h 40 | pullToMiddle 1, 3, h 41 | 42 | example : A ∨ B ∨ C ∨ D ∨ E → (D ∨ E) ∨ A ∨ B ∨ C := by 43 | intro h 44 | pull h, (D ∨ E), 3 45 | 46 | example : A ∨ B ∨ C ∨ D ∨ E ∨ F ∨ G ∨ H → F ∨ A ∨ B ∨ C ∨ D ∨ E ∨ G ∨ H := by 47 | intro h 48 | pull h, F 49 | -------------------------------------------------------------------------------- /Test/Reconstruction/Resolution.expected: -------------------------------------------------------------------------------- 1 | Test/Reconstruction/Resolution.lean:7:3: error: unknown tactic 2 | Test/Reconstruction/Resolution.lean:5:69: error: unsolved goals 3 | A B C D E F G : Prop 4 | h₁ : A ∨ B ∨ C ∨ D 5 | h₂ : E ∨ ¬C ∨ F ∨ G 6 | ⊢ A ∨ B ∨ D ∨ E ∨ F ∨ G 7 | Test/Reconstruction/Resolution.lean:11:3: error: unknown tactic 8 | Test/Reconstruction/Resolution.lean:9:55: error: unsolved goals 9 | A B C D E : Prop 10 | h₁ : A ∨ B ∨ C ∨ D 11 | h₂ : E ∨ ¬B 12 | ⊢ A ∨ (C ∨ D) ∨ E 13 | Test/Reconstruction/Resolution.lean:15:3: error: unknown tactic 14 | Test/Reconstruction/Resolution.lean:13:37: error: unsolved goals 15 | A B C : Prop 16 | h₁ : ¬A 17 | h₂ : B ∨ A ∨ C 18 | ⊢ B ∨ C 19 | Test/Reconstruction/Resolution.lean:19:3: error: unknown tactic 20 | Test/Reconstruction/Resolution.lean:17:89: error: unsolved goals 21 | A B C D : Prop 22 | h₁ : ¬(A ∧ B) ∨ C ∨ ¬D ∨ ¬A 23 | h₂ : A ∨ ¬(A ∧ B) 24 | ⊢ ¬(A ∧ B) ∨ C ∨ ¬D ∨ ¬(A ∧ B) 25 | Test/Reconstruction/Resolution.lean:23:3: error: unknown tactic 26 | Test/Reconstruction/Resolution.lean:21:45: error: unsolved goals 27 | A B C D : Prop 28 | h₁ : A ∨ B ∨ C ∨ D 29 | h₂ : ¬A 30 | ⊢ B ∨ C ∨ D 31 | Test/Reconstruction/Resolution.lean:27:3: error: unknown tactic 32 | Test/Reconstruction/Resolution.lean:25:82: error: unsolved goals 33 | B A C D E F : Prop 34 | h₁ : B ∨ A ∨ C ∨ A 35 | h₂ : D ∨ ¬A ∨ E ∨ ¬A ∨ F 36 | ⊢ B ∨ C ∨ A ∨ D ∨ E ∨ ¬A ∨ F 37 | Test/Reconstruction/Resolution.lean:31:3: error: unknown tactic 38 | Test/Reconstruction/Resolution.lean:29:29: error: unsolved goals 39 | A : Prop 40 | h₁ : A 41 | h₂ : ¬A 42 | ⊢ False 43 | -------------------------------------------------------------------------------- /Test/Reconstruction/Resolution.lean: -------------------------------------------------------------------------------- 1 | import Smt 2 | 3 | open Smt.Reconstruct 4 | 5 | example : A ∨ B ∨ C ∨ D → E ∨ ¬ C ∨ F ∨ G → A ∨ B ∨ D ∨ E ∨ F ∨ G := by 6 | intros h₁ h₂ 7 | R1 h₁, h₂, C, [-1, -1] 8 | 9 | example : A ∨ B ∨ C ∨ D → E ∨ ¬ B → A ∨ (C ∨ D) ∨ E := by 10 | intros h₁ h₂ 11 | R1 h₁, h₂, B, [2, 1] 12 | 13 | example : ¬ A → B ∨ A ∨ C → B ∨ C := by 14 | intros h₁ h₂ 15 | R2 h₁, h₂, A, [0, 2] 16 | 17 | example : ¬ (A ∧ B) ∨ C ∨ ¬ D ∨ ¬ A → A ∨ ¬ (A ∧ B) → ¬ (A ∧ B) ∨ C ∨ ¬ D ∨ ¬ (A ∧ B) := by 18 | intros h₁ h₂ 19 | R2 h₁, h₂, A 20 | 21 | example : A ∨ B ∨ C ∨ D → ¬ A → B ∨ C ∨ D := by 22 | intros h₁ h₂ 23 | R1 h₁, h₂, A 24 | 25 | example : B ∨ A ∨ C ∨ A → D ∨ ¬ A ∨ E ∨ ¬ A ∨ F → B ∨ C ∨ A ∨ D ∨ E ∨ ¬ A ∨ F := by 26 | intros h₁ h₂ 27 | R1 h₁, h₂, A 28 | 29 | example : A → ¬ A → False := by 30 | intros h₁ h₂ 31 | R1 h₁, h₂, A 32 | -------------------------------------------------------------------------------- /Test/Solver/Error.expected: -------------------------------------------------------------------------------- 1 | Test/Solver/Error.lean:26:0: error: Cannot bind x to symbol of type Int, maybe the symbol has already been defined? 2 | -------------------------------------------------------------------------------- /Test/Solver/Error.lean: -------------------------------------------------------------------------------- 1 | import Smt.Translate.Solver 2 | 3 | open Smt Translate Solver 4 | 5 | def cvc5.os := 6 | if System.Platform.isWindows then "Win64" 7 | else if System.Platform.isOSX then "macOS" 8 | else "Linux" 9 | 10 | def cvc5.arch := 11 | if System.Platform.target.startsWith "x86_64" then "x86_64" 12 | else "arm64" 13 | 14 | def cvc5.target := s!"{os}-{arch}-static" 15 | 16 | def query : SolverM Result := do 17 | setLogic "LIA" 18 | declareConst "x" (.symbolT "Int") 19 | declareConst "x" (.symbolT "Int") 20 | checkSat 21 | 22 | def main : IO Unit := do 23 | let ss ← createFromKind .cvc5 s!".lake/packages/cvc5/cvc5-{cvc5.target}/bin/cvc5" none 24 | _ ← StateT.run query ss 25 | 26 | #eval main 27 | -------------------------------------------------------------------------------- /Test/Solver/Interactive.expected: -------------------------------------------------------------------------------- 1 | query: 2 | (set-logic LIA) 3 | (declare-const x Int) 4 | (declare-const y Int) 5 | (assert (< x y)) 6 | (check-sat) 7 | (assert (> (+ x 1) y)) 8 | (check-sat) 9 | 10 | res: unsat 11 | -------------------------------------------------------------------------------- /Test/Solver/Interactive.lean: -------------------------------------------------------------------------------- 1 | import Smt.Translate.Solver 2 | 3 | open Smt Translate Solver 4 | 5 | def cvc5.os := 6 | if System.Platform.isWindows then "Win64" 7 | else if System.Platform.isOSX then "macOS" 8 | else "Linux" 9 | 10 | def cvc5.arch := 11 | if System.Platform.target.startsWith "x86_64" then "x86_64" 12 | else "arm64" 13 | 14 | def cvc5.target := s!"{os}-{arch}-static" 15 | 16 | open Term in 17 | def query : SolverM Result := do 18 | setLogic "LIA" 19 | declareConst "x" (symbolT "Int") 20 | declareConst "y" (symbolT "Int") 21 | assert (mkApp2 (symbolT "<") (symbolT "x") (symbolT "y")) 22 | let mut res ← checkSat 23 | if res = .sat then 24 | assert (mkApp2 (symbolT ">") 25 | ((mkApp2 (symbolT "+") (symbolT "x") (literalT "1"))) 26 | (symbolT "y")) 27 | res ← checkSat 28 | return res 29 | 30 | def main : IO Unit := do 31 | let ss ← createFromKind .cvc5 s!".lake/packages/cvc5/cvc5-{cvc5.target}/bin/cvc5" none 32 | let (res, ss) ← StateT.run query ss 33 | _ ← StateT.run exit ss 34 | println! "query:\n{Command.cmdsAsQuery ss.commands.reverse}\n\nres: {res}" 35 | 36 | #eval main 37 | -------------------------------------------------------------------------------- /Test/Solver/Model.expected: -------------------------------------------------------------------------------- 1 | query: 2 | (set-logic LIA) 3 | (set-option :produce-models true) 4 | (declare-const x Int) 5 | (check-sat) 6 | (get-model) 7 | 8 | res: sat 9 | 10 | model: 11 | ( 12 | (define-fun x () Int 0) 13 | ) 14 | -------------------------------------------------------------------------------- /Test/Solver/Model.lean: -------------------------------------------------------------------------------- 1 | import Smt.Translate.Solver 2 | 3 | open Smt Translate Solver 4 | 5 | def cvc5.os := 6 | if System.Platform.isWindows then "Win64" 7 | else if System.Platform.isOSX then "macOS" 8 | else "Linux" 9 | 10 | def cvc5.arch := 11 | if System.Platform.target.startsWith "x86_64" then "x86_64" 12 | else "arm64" 13 | 14 | def cvc5.target := s!"{os}-{arch}-static" 15 | 16 | def query : SolverM (Result × Sexp) := do 17 | setLogic "LIA" 18 | setOption "produce-models" "true" 19 | declareConst "x" (.symbolT "Int") 20 | return (← checkSat, ← getModel) 21 | 22 | def main : IO Unit := do 23 | let ss ← createFromKind .cvc5 s!".lake/packages/cvc5/cvc5-{cvc5.target}/bin/cvc5" none 24 | let ((res, model), ss) ← StateT.run query ss 25 | _ ← StateT.run exit ss 26 | println! "query:\n{Command.cmdsAsQuery ss.commands.reverse}\n\nres: {res}\n\nmodel:\n{model}" 27 | 28 | #eval main 29 | -------------------------------------------------------------------------------- /Test/Solver/Proof.expected: -------------------------------------------------------------------------------- 1 | query: 2 | (set-logic QF_UF) 3 | (assert false) 4 | (check-sat) 5 | (get-proof) 6 | 7 | res: unsat 8 | 9 | proof: 10 | ((assume @p1 false)) 11 | -------------------------------------------------------------------------------- /Test/Solver/Proof.lean: -------------------------------------------------------------------------------- 1 | import Smt.Translate.Solver 2 | 3 | open Smt Translate Solver 4 | 5 | def cvc5.os := 6 | if System.Platform.isWindows then "Win64" 7 | else if System.Platform.isOSX then "macOS" 8 | else "Linux" 9 | 10 | def cvc5.arch := 11 | if System.Platform.target.startsWith "x86_64" then "x86_64" 12 | else "arm64" 13 | 14 | def cvc5.target := s!"{os}-{arch}-static" 15 | 16 | def query : SolverM Sexp := do 17 | setLogic "QF_UF" 18 | assert (.symbolT "false") 19 | _ ← checkSat 20 | getProof 21 | 22 | def main : IO Unit := do 23 | let ss ← createFromKind .cvc5 s!".lake/packages/cvc5/cvc5-{cvc5.target}/bin/cvc5" none 24 | let (res, ss) ← StateT.run query ss 25 | _ ← StateT.run exit ss 26 | println! "query:\n{Command.cmdsAsQuery ss.commands.reverse}\n\nres: unsat\n\nproof:\n{res}" 27 | 28 | #eval main 29 | -------------------------------------------------------------------------------- /Test/Solver/Triv.expected: -------------------------------------------------------------------------------- 1 | query: 2 | (check-sat) 3 | 4 | res: sat 5 | -------------------------------------------------------------------------------- /Test/Solver/Triv.lean: -------------------------------------------------------------------------------- 1 | import Smt.Translate.Solver 2 | 3 | open Smt Translate Solver 4 | 5 | def cvc5.os := 6 | if System.Platform.isWindows then "Win64" 7 | else if System.Platform.isOSX then "macOS" 8 | else "Linux" 9 | 10 | def cvc5.arch := 11 | if System.Platform.target.startsWith "x86_64" then "x86_64" 12 | else "arm64" 13 | 14 | def cvc5.target := s!"{os}-{arch}-static" 15 | 16 | def query : SolverM Result := checkSat 17 | 18 | def main : IO Unit := do 19 | let ss ← createFromKind .cvc5 s!".lake/packages/cvc5/cvc5-{cvc5.target}/bin/cvc5" none 20 | let (res, ss) ← StateT.run query ss 21 | _ ← StateT.run exit ss 22 | println! "query:\n{Command.cmdsAsQuery ss.commands.reverse}\n\nres: {res}" 23 | 24 | #eval main 25 | -------------------------------------------------------------------------------- /Test/String/Append.expected: -------------------------------------------------------------------------------- 1 | goal: "a" ++ "b" = "ab" 2 | 3 | query: 4 | (set-logic ALL) 5 | (assert (not (= (str.++ "a" "b") "ab"))) 6 | (check-sat) 7 | -------------------------------------------------------------------------------- /Test/String/Append.lean: -------------------------------------------------------------------------------- 1 | import Smt 2 | 3 | theorem append : "a" ++ "b" = "ab" := by 4 | smt_show 5 | rfl 6 | -------------------------------------------------------------------------------- /Test/String/Contains.expected: -------------------------------------------------------------------------------- 1 | Test/String/Contains.lean:4:2: error: cannot translate 97 2 | -------------------------------------------------------------------------------- /Test/String/Contains.lean: -------------------------------------------------------------------------------- 1 | import Smt 2 | 3 | theorem contains : "a".contains 'a' := by 4 | smt_show 5 | sorry 6 | -------------------------------------------------------------------------------- /Test/String/Empty.expected: -------------------------------------------------------------------------------- 1 | goal: "" = "" 2 | 3 | query: 4 | (set-logic ALL) 5 | (assert (not (= "" ""))) 6 | (check-sat) 7 | -------------------------------------------------------------------------------- /Test/String/Empty.lean: -------------------------------------------------------------------------------- 1 | import Smt 2 | 3 | theorem empty : "" = "" := by 4 | smt_show 5 | rfl 6 | -------------------------------------------------------------------------------- /Test/String/GetOp.expected: -------------------------------------------------------------------------------- 1 | Test/String/GetOp.lean:4:2: error: cannot translate 0 2 | -------------------------------------------------------------------------------- /Test/String/GetOp.lean: -------------------------------------------------------------------------------- 1 | import Smt 2 | 3 | theorem index : "a".get 0 = 'a' := by 4 | smt_show 5 | rfl 6 | -------------------------------------------------------------------------------- /Test/String/Length.expected: -------------------------------------------------------------------------------- 1 | goal: "a".length = 1 2 | 3 | query: 4 | (set-logic ALL) 5 | (assert (not (= (str.len "a") 1))) 6 | (check-sat) 7 | -------------------------------------------------------------------------------- /Test/String/Length.lean: -------------------------------------------------------------------------------- 1 | import Smt 2 | 3 | theorem length : "a".length = 1 := by 4 | smt_show 5 | rfl 6 | -------------------------------------------------------------------------------- /Test/String/Lt.expected: -------------------------------------------------------------------------------- 1 | goal: "a" < "b" 2 | 3 | query: 4 | (set-logic ALL) 5 | (assert (not (str.< "a" "b"))) 6 | (check-sat) 7 | -------------------------------------------------------------------------------- /Test/String/Lt.lean: -------------------------------------------------------------------------------- 1 | import Smt 2 | 3 | theorem lt : "a" < "b" := by 4 | smt_show 5 | decide 6 | -------------------------------------------------------------------------------- /Test/String/Replace.expected: -------------------------------------------------------------------------------- 1 | goal: "a".replace "a" "b" = "b" 2 | 3 | query: 4 | (set-logic ALL) 5 | (assert (not (= (str.replace_all "a" "a" "b") "b"))) 6 | (check-sat) 7 | Test/String/Replace.lean:3:8: warning: declaration uses 'sorry' 8 | -------------------------------------------------------------------------------- /Test/String/Replace.lean: -------------------------------------------------------------------------------- 1 | import Smt 2 | 3 | theorem replace : "a".replace "a" "b" = "b" := by 4 | smt_show 5 | admit 6 | -------------------------------------------------------------------------------- /Test/Unit/Int.expected: -------------------------------------------------------------------------------- https://raw.githubusercontent.com/ufmg-smite/lean-smt/db6a7caf8685b33897ceebc9159bb1e180e2c568/Test/Unit/Int.expected -------------------------------------------------------------------------------- /Test/Unit/Int.lean: -------------------------------------------------------------------------------- 1 | import Smt 2 | 3 | def Lean.Expr.intAbs? (e : Lean.Expr) : Option Expr := do 4 | let .app (.const ``Int.abs _) x := e | none 5 | return x 6 | 7 | namespace Smt.Translate.Int 8 | 9 | open Translator Term 10 | 11 | @[smt_translate] def translateAbs : Translator := fun e => do 12 | if let some x := e.intAbs? then 13 | return appT (symbolT "abs") (← applyTranslators! x) 14 | else 15 | return none 16 | 17 | end Smt.Translate.Int 18 | 19 | example : (1 : Int) < 2 := by 20 | smt 21 | 22 | example {x : Int} : x = 1 → x < 2 := by 23 | smt 24 | 25 | example {x : Int} : x * x ≥ 0 := by 26 | smt 27 | 28 | example {x y : Int} : x = y → x * y ≥ 0 := by 29 | smt 30 | 31 | example {x₁ x₂ y₁ y₂ : Int} (h₁ : x₁.abs > y₁.abs) (h₂ : x₂.abs > y₂.abs) : (x₁ * x₂).abs > (y₁ * y₂).abs := by 32 | smt [h₁, h₂] 33 | 34 | example {x₁ x₂ x₃ y₁ y₂ y₃ : Int} (h₁ : x₁.abs > y₁.abs) (h₂ : x₂.abs > y₂.abs) (h₃ : x₃.abs > y₃.abs) : (x₁ * x₂ * x₃).abs > (y₁ * y₂ * y₃).abs := by 35 | smt [h₁, h₂, h₃] 36 | -------------------------------------------------------------------------------- /Test/Unit/Quant.expected: -------------------------------------------------------------------------------- https://raw.githubusercontent.com/ufmg-smite/lean-smt/db6a7caf8685b33897ceebc9159bb1e180e2c568/Test/Unit/Quant.expected -------------------------------------------------------------------------------- /Test/Unit/Quant.lean: -------------------------------------------------------------------------------- 1 | import Smt 2 | 3 | example {p q r : U → Prop} : (∀ x, p x ∧ q x ∧ r x) = ((∀ x, p x) ∧ (∀ x, q x) ∧ (∀ x, r x)) := by 4 | smt 5 | 6 | example {p q : U → U → Prop} : (∀ x y, p x y ∧ q x y) = ((∀ x y, p x y) ∧ (∀ x y, q x y)) := by 7 | smt 8 | 9 | example {p : U → Prop} {q : Prop} : (∀ x, p x ∨ q) = ((∀ x, p x) ∨ q) := by 10 | smt 11 | 12 | example {p q : U → Prop} : (∀ x, (∀ y, p y) ∨ q x) = ((∀ y, p y) ∨ (∀ x, q x)) := by 13 | smt 14 | 15 | example {p : Prop} {q : U → Prop} : (∀ x, p ∨ q x) = (p ∨ (∀ x, q x)) := by 16 | smt 17 | 18 | example {p q r : U → Prop} : (∀ x y z, p x ∨ q y ∨ r z) = ((∀ x, p x) ∨ (∀ y, q y) ∨ (∀ z, r z)) := by 19 | smt 20 | 21 | example {p : U → Prop} {q : U → U → Prop} {r : Prop} : (∀ x y₁ y₂, p x ∨ q y₁ y₂ ∨ r) = ((∀ x, p x) ∨ (∀ y₁ y₂, q y₁ y₂) ∨ r) := by 22 | smt 23 | 24 | example [Decidable c] {p q : U → Prop} : (∀ x, ite c (p x) (q x)) = ite c (∀ x, p x) (∀ x, q x) := by 25 | smt 26 | 27 | example [Decidable c] {p q : U → U → Prop} : (∀ x y, ite c (p x y) (q x y)) = ite c (∀ x y, p x y) (∀ x y, q x y) := by 28 | smt 29 | 30 | example {p : U → U → U → Prop} : (∀ x y z, p x y z) = (∀ y z x, p x y z) := by 31 | smt 32 | 33 | example {t : U} {p : U → Prop} : (∀ y, y ≠ t ∨ p y) = p t := by 34 | smt 35 | 36 | example {t : U} : (∀ x, x ≠ t) = False := by 37 | smt 38 | -------------------------------------------------------------------------------- /Test/Unit/Rat.expected: -------------------------------------------------------------------------------- https://raw.githubusercontent.com/ufmg-smite/lean-smt/db6a7caf8685b33897ceebc9159bb1e180e2c568/Test/Unit/Rat.expected -------------------------------------------------------------------------------- /Test/Unit/Rat.lean: -------------------------------------------------------------------------------- 1 | import Smt 2 | import Smt.Rat 3 | 4 | def Lean.Expr.ratAbs? (e : Lean.Expr) : Option Expr := do 5 | let .app (.const ``Rat.abs _) x := e | none 6 | return x 7 | 8 | def Lean.Expr.ratFloor? (e : Lean.Expr) : Option Expr := do 9 | let .app (.const ``Rat.floor _) x := e | none 10 | return x 11 | 12 | namespace Smt.Translate.Rat 13 | 14 | open Translator Term 15 | 16 | private def mkRat : Lean.Expr := 17 | .const ``Rat [] 18 | 19 | @[smt_translate] def translateType : Translator := fun e => match e with 20 | | .const ``Rat _ => return symbolT "Real" 21 | | _ => return none 22 | 23 | @[smt_translate] def translateInt : Translator := fun e => do 24 | if let some x := e.ratFloor? then 25 | return appT (symbolT "to_int") (← applyTranslators! x) 26 | else 27 | return none 28 | 29 | @[smt_translate] def translateRat : Translator := fun e => do 30 | if let some n := e.natLitOf? mkRat then 31 | return literalT s!"{n}.0" 32 | else if let some x := e.intCastOf? mkRat then 33 | return appT (symbolT "to_real") (← applyTranslators! x) 34 | else if let some x := e.negOf? mkRat then 35 | return appT (symbolT "-") (← applyTranslators! x) 36 | else if let some x := e.ratAbs? then 37 | return appT (symbolT "abs") (← applyTranslators! x) 38 | else if let some (x, y) := e.hAddOf? mkRat mkRat then 39 | return mkApp2 (symbolT "+") (← applyTranslators! x) (← applyTranslators! y) 40 | else if let some (x, y) := e.hSubOf? mkRat mkRat then 41 | return mkApp2 (symbolT "-") (← applyTranslators! x) (← applyTranslators! y) 42 | else if let some (x, y) := e.hMulOf? mkRat mkRat then 43 | return mkApp2 (symbolT "*") (← applyTranslators! x) (← applyTranslators! y) 44 | else if let some (x, y) := e.hDivOf? mkRat mkRat then 45 | return mkApp2 (symbolT "/") (← applyTranslators! x) (← applyTranslators! y) 46 | else 47 | return none 48 | 49 | @[smt_translate] def translateProp : Translator := fun e => do 50 | if let some (x, y) := e.ltOf? mkRat then 51 | return mkApp2 (symbolT "<") (← applyTranslators! x) (← applyTranslators! y) 52 | else if let some (x, y) := e.leOf? mkRat then 53 | return mkApp2 (symbolT "<=") (← applyTranslators! x) (← applyTranslators! y) 54 | else if let some (x, y) := e.geOf? mkRat then 55 | return mkApp2 (symbolT ">=") (← applyTranslators! x) (← applyTranslators! y) 56 | else if let some (x, y) := e.gtOf? mkRat then 57 | return mkApp2 (symbolT ">") (← applyTranslators! x) (← applyTranslators! y) 58 | else 59 | return none 60 | 61 | end Smt.Translate.Rat 62 | 63 | example : (1 : Rat) < 2 := by 64 | smt 65 | 66 | example {x : Rat} : x = 1 → x < 2 := by 67 | smt 68 | 69 | example {x : Rat} : x * x ≥ 0 := by 70 | smt 71 | 72 | example {x y : Rat} : x = y → x * y ≥ 0 := by 73 | smt 74 | 75 | example {x₁ x₂ y₁ y₂ : Rat} (h₁ : x₁.abs > y₁.abs) (h₂ : x₂.abs > y₂.abs) : (x₁ * x₂).abs > (y₁ * y₂).abs := by 76 | smt [h₁, h₂] 77 | 78 | example {x₁ x₂ x₃ y₁ y₂ y₃ : Rat} (h₁ : x₁.abs > y₁.abs) (h₂ : x₂.abs > y₂.abs) (h₃ : x₃.abs > y₃.abs) : (x₁ * x₂ * x₃).abs > (y₁ * y₂ * y₃).abs := by 79 | smt [h₁, h₂, h₃] 80 | -------------------------------------------------------------------------------- /Test/Unit/Real.expected: -------------------------------------------------------------------------------- https://raw.githubusercontent.com/ufmg-smite/lean-smt/db6a7caf8685b33897ceebc9159bb1e180e2c568/Test/Unit/Real.expected -------------------------------------------------------------------------------- /Test/Unit/Real.lean: -------------------------------------------------------------------------------- 1 | import Smt 2 | import Smt.Real 3 | 4 | example : (1 : Real) < 2 := by 5 | smt 6 | 7 | example {x : Real} : x = 1 → x < 2 := by 8 | smt 9 | 10 | example {x : Real} : x * x ≥ 0 := by 11 | smt 12 | 13 | example {x y : Real} : x = y → x * y ≥ 0 := by 14 | smt 15 | 16 | example {x₁ x₂ y₁ y₂ : Real} (h₁ : |x₁| > |y₁|) (h₂ : |x₂| > |y₂|) : |x₁ * x₂| > |y₁ * y₂| := by 17 | smt [h₁, h₂] 18 | 19 | example {x₁ x₂ x₃ y₁ y₂ y₃ : Real} (h₁ : |x₁| > |y₁|) (h₂ : |x₂| > |y₂|) (h₃ : |x₃| > |y₃|) : |x₁ * x₂ * x₃| > |y₁ * y₂ * y₃| := by 20 | smt [h₁, h₂, h₃] 21 | -------------------------------------------------------------------------------- /Test/linarith.expected: -------------------------------------------------------------------------------- 1 | Test/linarith.lean:27:5: warning: unused variable `h₁` 2 | note: this linter can be disabled with `set_option linter.unusedVariables false` 3 | Test/linarith.lean:34:9: warning: unused variable `e` 4 | note: this linter can be disabled with `set_option linter.unusedVariables false` 5 | Test/linarith.lean:53:0: warning: declaration uses 'sorry' 6 | Test/linarith.lean:101:9: warning: unused variable `a` 7 | note: this linter can be disabled with `set_option linter.unusedVariables false` 8 | Test/linarith.lean:101:13: warning: unused variable `c` 9 | note: this linter can be disabled with `set_option linter.unusedVariables false` 10 | Test/linarith.lean:111:60: warning: unused variable `h3` 11 | note: this linter can be disabled with `set_option linter.unusedVariables false` 12 | Test/linarith.lean:118:9: warning: unused variable `a` 13 | note: this linter can be disabled with `set_option linter.unusedVariables false` 14 | Test/linarith.lean:118:13: warning: unused variable `c` 15 | note: this linter can be disabled with `set_option linter.unusedVariables false` 16 | Test/linarith.lean:161:34: warning: unused variable `z` 17 | note: this linter can be disabled with `set_option linter.unusedVariables false` 18 | Test/linarith.lean:162:5: warning: unused variable `h5` 19 | note: this linter can be disabled with `set_option linter.unusedVariables false` 20 | -------------------------------------------------------------------------------- /lean-toolchain: -------------------------------------------------------------------------------- 1 | leanprover/lean4:v4.20.0-rc5 2 | --------------------------------------------------------------------------------