├── .github ├── dependabot.yml └── workflows │ └── build.yml ├── .gitignore ├── Aesop.lean ├── Aesop ├── BaseM.lean ├── Builder.lean ├── Builder │ ├── Apply.lean │ ├── Basic.lean │ ├── Cases.lean │ ├── Constructors.lean │ ├── Default.lean │ ├── Forward.lean │ ├── NormSimp.lean │ ├── Tactic.lean │ └── Unfold.lean ├── BuiltinRules.lean ├── BuiltinRules │ ├── ApplyHyps.lean │ ├── Assumption.lean │ ├── DestructProducts.lean │ ├── Ext.lean │ ├── Intros.lean │ ├── Rfl.lean │ ├── Split.lean │ └── Subst.lean ├── Check.lean ├── Constants.lean ├── ElabM.lean ├── Exception.lean ├── Forward │ ├── CompleteMatchQueue.lean │ ├── LevelIndex.lean │ ├── Match.lean │ ├── Match │ │ └── Types.lean │ ├── PremiseIndex.lean │ ├── RuleInfo.lean │ ├── SlotIndex.lean │ ├── State.lean │ ├── State │ │ ├── ApplyGoalDiff.lean │ │ └── Initial.lean │ └── Substitution.lean ├── Frontend.lean ├── Frontend │ ├── Attribute.lean │ ├── Basic.lean │ ├── Command.lean │ ├── Extension.lean │ ├── Extension │ │ └── Init.lean │ ├── RuleExpr.lean │ ├── Saturate.lean │ └── Tactic.lean ├── Index.lean ├── Index │ ├── Basic.lean │ ├── DiscrKeyConfig.lean │ ├── DiscrTreeConfig.lean │ ├── Forward.lean │ └── RulePattern.lean ├── Main.lean ├── Nanos.lean ├── Options.lean ├── Options │ ├── Internal.lean │ └── Public.lean ├── Percent.lean ├── RPINF.lean ├── RPINF │ └── Basic.lean ├── Rule.lean ├── Rule │ ├── Basic.lean │ ├── Forward.lean │ └── Name.lean ├── RulePattern.lean ├── RulePattern │ └── Cache.lean ├── RuleSet.lean ├── RuleSet │ ├── Filter.lean │ ├── Member.lean │ └── Name.lean ├── RuleTac.lean ├── RuleTac │ ├── Apply.lean │ ├── Basic.lean │ ├── Cases.lean │ ├── Descr.lean │ ├── ElabRuleTerm.lean │ ├── FVarIdSubst.lean │ ├── Forward.lean │ ├── Forward │ │ └── Basic.lean │ ├── GoalDiff.lean │ ├── Preprocess.lean │ ├── RuleTerm.lean │ └── Tactic.lean ├── Saturate.lean ├── Script │ ├── Check.lean │ ├── CtorNames.lean │ ├── GoalWithMVars.lean │ ├── Main.lean │ ├── OptimizeSyntax.lean │ ├── SScript.lean │ ├── ScriptM.lean │ ├── SpecificTactics.lean │ ├── Step.lean │ ├── StructureDynamic.lean │ ├── StructureStatic.lean │ ├── Tactic.lean │ ├── TacticState.lean │ ├── UScript.lean │ ├── UScriptToSScript.lean │ └── Util.lean ├── Search │ ├── ExpandSafePrefix.lean │ ├── Expansion.lean │ ├── Expansion │ │ ├── Basic.lean │ │ ├── Norm.lean │ │ └── Simp.lean │ ├── Main.lean │ ├── Queue.lean │ ├── Queue │ │ └── Class.lean │ ├── RuleSelection.lean │ └── SearchM.lean ├── Stats │ ├── Basic.lean │ ├── Extension.lean │ └── Report.lean ├── Tracing.lean ├── Tree.lean ├── Tree │ ├── AddRapp.lean │ ├── Check.lean │ ├── Data.lean │ ├── Data │ │ └── ForwardRuleMatches.lean │ ├── ExtractProof.lean │ ├── ExtractScript.lean │ ├── Free.lean │ ├── RunMetaM.lean │ ├── State.lean │ ├── Tracing.lean │ ├── Traversal.lean │ ├── TreeM.lean │ └── UnsafeQueue.lean └── Util │ ├── Basic.lean │ ├── EqualUpToIds.lean │ ├── Tactic.lean │ ├── Tactic │ ├── Ext.lean │ └── Unfold.lean │ ├── UnionFind.lean │ └── UnorderedArraySet.lean ├── AesopTest ├── 10.lean ├── 12.lean ├── 125.lean ├── 126.lean ├── 13.lean ├── 13_2.lean ├── 18.lean ├── 2.lean ├── 20.lean ├── 203.lean ├── 205.lean ├── 207.lean ├── 23.lean ├── 26.lean ├── 27.lean ├── 41.lean ├── 43.lean ├── AddRulesCommand.lean ├── Aesop.lean ├── AllWeaken.lean ├── ApplyHypsTransparency.lean ├── ApplyTransparency.lean ├── AssumptionTransparency.lean ├── AuxDecl.lean ├── BigStep.lean ├── Cases.lean ├── CasesScript.lean ├── CasesTransparency.lean ├── Com.lean ├── CompositeLocalRuleTerm.lean ├── ConstructorEquations.lean ├── Constructors.lean ├── CustomIndexing.lean ├── CustomTactic.lean ├── DefaultRuleSets.lean ├── DefaultRuleSetsInit.lean ├── DestructProducts.lean ├── DestructProductsTransparency.lean ├── DocLists.lean ├── DroppedMVars.lean ├── ElabConfig.lean ├── EnableUnfold.lean ├── EqualUpToIds.lean ├── Erase.lean ├── EraseSimp.lean ├── EraseUnfold.lean ├── Ext.lean ├── ExtScript.lean ├── Filter.lean ├── Forward.lean ├── ForwardConstant.lean ├── ForwardRedundantHypsWithMVars.lean ├── ForwardTransparency.lean ├── ForwardUnknownFVar.lean ├── GlobalRuleIdentErrorChecking.lean ├── HeartbeatLimit.lean ├── IncompleteScript.lean ├── Intros.lean ├── IntrosAllTransparency.lean ├── Jesse.lean ├── LegacyForward.lean ├── List.lean ├── LocalRuleSet.lean ├── LocalTactic.lean ├── Logic.lean ├── MVarsInInitialGoal.lean ├── Metas.lean ├── NameResolution.lean ├── NoImportClash.lean ├── NoNormSimp.lean ├── NoProgress.lean ├── Nonterminal.lean ├── NormSimp.lean ├── Persistence0.lean ├── Persistence1.lean ├── Persistence2.lean ├── Persistence3.lean ├── PostponeSafeRules.lean ├── RecursiveUnfoldRule.lean ├── RulePattern.lean ├── RulePatternLooseBVar.lean ├── RulePatternUniverseBug.lean ├── RuleSetNameHygiene0.lean ├── RuleSetNameHygiene1.lean ├── RuleSets0.lean ├── RuleSets1.lean ├── Safe.lean ├── SafeExtractionCopyIntroducedMVars.lean ├── SafePrefixExpansionRappLimit.lean ├── SafePrefixInTerminalError.lean ├── SaturatePerformance.lean ├── ScriptWithOptions.lean ├── SeqCalcProver.lean ├── SimpLetHypotheses.lean ├── Simprocs.lean ├── Split.lean ├── SplitScript.lean ├── Stats.lean ├── Strategy.lean ├── Subst.lean ├── TacGen.lean ├── TacticConfig.lean ├── Tauto.lean ├── TerminalError.lean ├── TraceProof.lean ├── TryThisIndentation.lean ├── Unfold.lean ├── UnreachableTacticLinter.lean └── WarnApplyIff.lean ├── LICENSE ├── README.md ├── lake-manifest.json ├── lakefile.toml └── lean-toolchain /.github/dependabot.yml: -------------------------------------------------------------------------------- 1 | # See the documentation for all configuration options: 2 | # https://docs.github.com/code-security/dependabot/dependabot-version-updates/configuration-options-for-the-dependabot.yml-file 3 | 4 | version: 2 5 | updates: 6 | - package-ecosystem: "github-actions" # See documentation for possible values 7 | directory: "/" # Location of package manifests 8 | schedule: 9 | interval: "daily" 10 | -------------------------------------------------------------------------------- /.github/workflows/build.yml: -------------------------------------------------------------------------------- 1 | name: Build 2 | 3 | on: 4 | push: 5 | merge_group: 6 | 7 | jobs: 8 | build: 9 | runs-on: ${{ matrix.os }} 10 | strategy: 11 | matrix: 12 | os: [ubuntu-latest] 13 | steps: 14 | - uses: actions/checkout@v4 15 | 16 | - uses: leanprover/lean-action@v1 17 | with: 18 | check-reservoir-eligibility: true 19 | build-args: "--wfail" 20 | -------------------------------------------------------------------------------- /.gitignore: -------------------------------------------------------------------------------- 1 | # macOS leaves these files everywhere: 2 | .DS_Store 3 | # Prior to v4.3.0-rc2 lake stored files in these locations. 4 | # We'll leave them in the `.gitignore` for a while for users switching between toolchains. 5 | /build 6 | /lake-packages 7 | /lakefile.olean 8 | # After v4.3.0-rc2 lake stores its files here: 9 | /.lake 10 | -------------------------------------------------------------------------------- /Aesop.lean: -------------------------------------------------------------------------------- 1 | import Aesop.Main 2 | import Aesop.Frontend.Command 3 | import Aesop.Frontend.Saturate 4 | import Aesop.BuiltinRules 5 | -------------------------------------------------------------------------------- /Aesop/BaseM.lean: -------------------------------------------------------------------------------- 1 | /- 2 | Copyright (c) 2024 Jannis Limperg. All rights reserved. 3 | Released under Apache 2.0 license as described in the file LICENSE. 4 | Authors: Jannis Limperg 5 | -/ 6 | 7 | import Aesop.Stats.Basic 8 | import Aesop.RulePattern.Cache 9 | 10 | set_option linter.missingDocs true 11 | 12 | open Lean 13 | 14 | namespace Aesop 15 | 16 | /-- State of the `BaseM` monad. -/ 17 | structure BaseM.State where 18 | /-- The rule pattern cache. -/ 19 | rulePatternCache : RulePatternCache 20 | /-- The RPINF cache. -/ 21 | rpinfCache : RPINFCache 22 | /-- Stats collected during an Aesop call. -/ 23 | stats : Stats 24 | deriving Inhabited 25 | 26 | instance : EmptyCollection BaseM.State := 27 | ⟨by refine' { .. } <;> exact ∅⟩ 28 | 29 | /-- Aesop's base monad. Contains no interesting data, only various caches and 30 | stats. -/ 31 | abbrev BaseM := StateRefT BaseM.State MetaM 32 | 33 | namespace BaseM 34 | 35 | /-- Run a `BaseM` action. -/ 36 | protected def run (x : BaseM α) (stats : Stats := ∅) : MetaM (α × Stats) := do 37 | let (a, s) ← StateRefT'.run x { stats, rulePatternCache := ∅, rpinfCache := ∅ } 38 | return (a, s.stats) 39 | 40 | instance : MonadHashMapCacheAdapter Expr RulePatternCache.Entry BaseM where 41 | getCache := return (← get).rulePatternCache.map 42 | modifyCache f := modify λ s => 43 | { s with rulePatternCache.map := f s.rulePatternCache.map } 44 | 45 | instance : MonadHashMapCacheAdapter Expr RPINFRaw BaseM where 46 | getCache := return (← get).rpinfCache.map 47 | modifyCache f := modify λ s => { s with rpinfCache.map := f s.rpinfCache.map } 48 | 49 | instance : MonadStats BaseM where 50 | modifyGetStats f := modifyGet λ s => 51 | let (a, stats) := f s.stats 52 | (a, { s with stats }) 53 | getStats := return (← get).stats 54 | modifyStats f := modify λ s => { s with stats := f s.stats } 55 | 56 | instance : MonadBacktrack Meta.SavedState BaseM where 57 | saveState := Meta.saveState 58 | restoreState := (·.restore) 59 | 60 | end Aesop.BaseM 61 | -------------------------------------------------------------------------------- /Aesop/Builder.lean: -------------------------------------------------------------------------------- 1 | /- 2 | Copyright (c) 2022 Jannis Limperg. All rights reserved. 3 | Released under Apache 2.0 license as described in the file LICENSE. 4 | Authors: Jannis Limperg 5 | -/ 6 | 7 | import Aesop.Builder.Apply 8 | import Aesop.Builder.Basic 9 | import Aesop.Builder.Cases 10 | import Aesop.Builder.Constructors 11 | import Aesop.Builder.Default 12 | import Aesop.Builder.Forward 13 | import Aesop.Builder.NormSimp 14 | import Aesop.Builder.Tactic 15 | import Aesop.Builder.Unfold 16 | -------------------------------------------------------------------------------- /Aesop/Builder/Apply.lean: -------------------------------------------------------------------------------- 1 | /- 2 | Copyright (c) 2022 Jannis Limperg. All rights reserved. 3 | Released under Apache 2.0 license as described in the file LICENSE. 4 | Authors: Jannis Limperg 5 | -/ 6 | 7 | import Aesop.Builder.Basic 8 | 9 | open Lean 10 | open Lean.Meta 11 | 12 | namespace Aesop 13 | 14 | namespace RuleBuilderOptions 15 | 16 | def applyTransparency (opts : RuleBuilderOptions) : TransparencyMode := 17 | opts.transparency?.getD .default 18 | 19 | def applyIndexTransparency (opts : RuleBuilderOptions) : TransparencyMode := 20 | opts.indexTransparency?.getD .reducible 21 | 22 | end RuleBuilderOptions 23 | 24 | namespace RuleBuilder 25 | 26 | def getApplyIndexingMode (indexMd : TransparencyMode) (type : Expr) : 27 | MetaM IndexingMode := 28 | if indexMd != .reducible then 29 | return .unindexed 30 | else 31 | IndexingMode.targetMatchingConclusion type 32 | 33 | def checkNoIff (type : Expr) : MetaM Unit := do 34 | if aesop.warn.applyIff.get (← getOptions) then 35 | forallTelescope type λ _ conclusion => do 36 | if ← testHelper conclusion λ e => return e.isAppOf' ``Iff then 37 | logWarning m!"Apply builder was used for a theorem with conclusion A ↔ B.\nYou probably want to use the simp builder or create an alias that applies the theorem in one direction.\nUse `set_option aesop.warn.applyIff false` to disable this warning." 38 | 39 | def applyCore (t : ElabRuleTerm) (pat? : Option RulePattern) 40 | (imode? : Option IndexingMode) (md indexMd : TransparencyMode) 41 | (phase : PhaseSpec) : MetaM LocalRuleSetMember := do 42 | let e ← t.expr 43 | let type ← inferType e 44 | let imode ← imode?.getDM $ getApplyIndexingMode indexMd type 45 | let tac := .apply t.toRuleTerm md 46 | return .global $ .base $ phase.toRule (← t.name) .apply t.scope tac imode pat? 47 | 48 | def apply : RuleBuilder := λ input => do 49 | let opts := input.options 50 | let e ← elabRuleTermForApplyLike input.term 51 | let t := ElabRuleTerm.ofElaboratedTerm input.term e 52 | let type ← inferType e 53 | checkNoIff type 54 | let pat? ← opts.pattern?.mapM (RulePattern.elab · e) 55 | applyCore t pat? opts.indexingMode? opts.applyTransparency 56 | opts.applyIndexTransparency input.phase 57 | 58 | end Aesop.RuleBuilder 59 | -------------------------------------------------------------------------------- /Aesop/Builder/Cases.lean: -------------------------------------------------------------------------------- 1 | /- 2 | Copyright (c) 2022 Jannis Limperg. All rights reserved. 3 | Released under Apache 2.0 license as described in the file LICENSE. 4 | Authors: Jannis Limperg 5 | -/ 6 | 7 | import Aesop.Builder.Basic 8 | import Aesop.RuleTac.Cases 9 | 10 | open Lean 11 | open Lean.Meta 12 | 13 | namespace Aesop 14 | 15 | namespace CasesPattern 16 | 17 | def check (decl : Name) (p : CasesPattern) : MetaM Unit := 18 | withoutModifyingState do 19 | let p ← p.toExpr 20 | unless p.isAppOf' decl do 21 | throwError "expected pattern '{p}' ({toString p}) to be an application of '{decl}'" 22 | 23 | def toIndexingMode (p : CasesPattern) : MetaM IndexingMode := 24 | withoutModifyingState do .hyps <$> mkDiscrTreePath (← p.toExpr) 25 | 26 | end CasesPattern 27 | 28 | 29 | namespace RuleBuilderOptions 30 | 31 | def casesTransparency (opts : RuleBuilderOptions) : TransparencyMode := 32 | opts.transparency?.getD .reducible 33 | 34 | def casesIndexTransparency (opts : RuleBuilderOptions) : TransparencyMode := 35 | opts.indexTransparency?.getD .reducible 36 | 37 | def casesPatterns (opts : RuleBuilderOptions) : Array CasesPattern := 38 | opts.casesPatterns?.getD #[] 39 | 40 | end RuleBuilderOptions 41 | 42 | namespace RuleBuilder 43 | 44 | def mkCasesTarget (decl : Name) (casesPatterns : Array CasesPattern) : 45 | CasesTarget := 46 | if casesPatterns.isEmpty then 47 | .decl decl 48 | else 49 | .patterns casesPatterns 50 | 51 | def getCasesIndexingMode (decl : Name) (indexMd : TransparencyMode) 52 | (casesPatterns : Array CasesPattern) : MetaM IndexingMode := do 53 | if indexMd != .reducible then 54 | return .unindexed 55 | if casesPatterns.isEmpty then 56 | IndexingMode.hypsMatchingConst decl 57 | else 58 | .or <$> casesPatterns.mapM (·.toIndexingMode) 59 | 60 | def casesCore (info : InductiveVal) (pats : Array CasesPattern) 61 | (imode? : Option IndexingMode) (md indexMd : TransparencyMode) 62 | (phase : PhaseSpec) : MetaM LocalRuleSetMember := do 63 | let decl := info.name 64 | pats.forM (·.check decl) 65 | let imode ← imode?.getDM $ getCasesIndexingMode decl indexMd pats 66 | let target := mkCasesTarget decl pats 67 | let ctorNames ← mkCtorNames info 68 | let tac := .cases target md info.isRec ctorNames 69 | return .global $ .base $ phase.toRule decl .cases .global tac imode none 70 | 71 | def cases : RuleBuilder := λ input => do 72 | let opts := input.options 73 | if input.phase.phase == .norm then throwError 74 | "aesop: cases builder cannot currently be used for norm rules." 75 | -- TODO `Meta.cases` may assign and introduce metavariables. 76 | -- (Specifically, it can *replace* existing metavariables, which Aesop 77 | -- counts as an assignment and an introduction.) 78 | let info ← elabInductiveRuleIdent .cases input.term 79 | casesCore info opts.casesPatterns opts.indexingMode? opts.casesTransparency 80 | opts.casesIndexTransparency input.phase 81 | 82 | end Aesop.RuleBuilder 83 | -------------------------------------------------------------------------------- /Aesop/Builder/Constructors.lean: -------------------------------------------------------------------------------- 1 | /- 2 | Copyright (c) 2022 Jannis Limperg. All rights reserved. 3 | Released under Apache 2.0 license as described in the file LICENSE. 4 | Authors: Jannis Limperg 5 | -/ 6 | 7 | import Aesop.Builder.Basic 8 | 9 | open Lean 10 | open Lean.Meta 11 | 12 | namespace Aesop.RuleBuilderOptions 13 | 14 | def constructorsTransparency (opts : RuleBuilderOptions) : TransparencyMode := 15 | opts.transparency?.getD .default 16 | 17 | def constructorsIndexTransparency (opts : RuleBuilderOptions) : TransparencyMode := 18 | opts.indexTransparency?.getD .reducible 19 | 20 | end RuleBuilderOptions 21 | 22 | namespace RuleBuilder 23 | 24 | def getConstructorsIndexingMode (indexMd : TransparencyMode) 25 | (info : InductiveVal) : MetaM IndexingMode := do 26 | if indexMd != .reducible then 27 | return .unindexed 28 | else 29 | let mut imodes := Array.mkEmpty info.numCtors 30 | for ctor in info.ctors do 31 | let ctorInfo ← getConstInfo ctor 32 | let imode ← IndexingMode.targetMatchingConclusion ctorInfo.type 33 | imodes := imodes.push imode 34 | return .or imodes 35 | 36 | def constructorsCore (info : InductiveVal) (imode? : Option IndexingMode) 37 | (md indexMd : TransparencyMode) (phase : PhaseSpec) : 38 | MetaM LocalRuleSetMember := do 39 | let tac := .constructors info.ctors.toArray md 40 | let imode ← imode?.getDM $ getConstructorsIndexingMode indexMd info 41 | return .global $ .base $ 42 | phase.toRule info.name .constructors .global tac imode none 43 | 44 | def constructors : RuleBuilder := λ input => do 45 | let info ← elabInductiveRuleIdent .constructors input.term 46 | let opts := input.options 47 | constructorsCore info opts.indexingMode? opts.constructorsTransparency 48 | opts.constructorsIndexTransparency input.phase 49 | 50 | end Aesop.RuleBuilder 51 | -------------------------------------------------------------------------------- /Aesop/Builder/Default.lean: -------------------------------------------------------------------------------- 1 | /- 2 | Copyright (c) 2022 Jannis Limperg. All rights reserved. 3 | Released under Apache 2.0 license as described in the file LICENSE. 4 | Authors: Jannis Limperg 5 | -/ 6 | 7 | import Aesop.Builder.Constructors 8 | import Aesop.Builder.NormSimp 9 | import Aesop.Builder.Tactic 10 | import Aesop.Builder.Apply 11 | 12 | open Lean 13 | open Lean.Meta 14 | 15 | namespace Aesop 16 | 17 | -- TODO In the default builders below, we should distinguish between fatal and 18 | -- nonfatal errors. E.g. if the `tactic` builder finds a declaration that is not 19 | -- of tactic type, this is a nonfatal error and we should continue with the next 20 | -- builder. But if the simp builder finds an equation that cannot be interpreted 21 | -- as a simp lemma for some reason, this is a fatal error. Continuing with the 22 | -- next builder is more confusing than anything because the user probably 23 | -- intended to add a simp lemma. 24 | 25 | def RuleBuilder.default : RuleBuilder := λ input => 26 | match input.phase.phase with 27 | | .safe => 28 | constructors input <|> 29 | tactic input <|> 30 | apply input <|> 31 | err "a safe" input 32 | | .unsafe => 33 | constructors input <|> 34 | tactic input <|> 35 | apply input <|> 36 | err "an unsafe" input 37 | | .norm => 38 | constructors input <|> 39 | tactic input <|> 40 | simp input <|> 41 | apply input <|> 42 | err "a norm" input 43 | where 44 | err (ruleType : String) : RuleBuilder := λ input => 45 | throwError m!"aesop: Unable to interpret '{input.term}' as {ruleType} rule. Try specifying a builder." 46 | 47 | end Aesop 48 | -------------------------------------------------------------------------------- /Aesop/Builder/NormSimp.lean: -------------------------------------------------------------------------------- 1 | /- 2 | Copyright (c) 2022 Jannis Limperg. All rights reserved. 3 | Released under Apache 2.0 license as described in the file LICENSE. 4 | Authors: Jannis Limperg 5 | -/ 6 | 7 | import Aesop.Builder.Basic 8 | 9 | open Lean 10 | open Lean.Meta 11 | 12 | namespace Aesop 13 | 14 | private def getSimpEntriesFromPropConst (decl : Name) : 15 | MetaM (Array SimpEntry) := do 16 | let thms ← ({} : SimpTheorems).addConst decl 17 | return SimpTheorems.simpEntries thms 18 | 19 | private def getSimpEntriesForConst (decl : Name) : MetaM (Array SimpEntry) := do 20 | let info ← getConstInfo decl 21 | let mut thms : SimpTheorems := {} 22 | if (← isProp info.type) then 23 | thms ← thms.addConst decl 24 | else if info.hasValue then 25 | thms ← thms.addDeclToUnfold decl 26 | return SimpTheorems.simpEntries thms 27 | 28 | def PhaseSpec.getSimpPrio [Monad m] [MonadError m] : PhaseSpec → m Nat 29 | | .norm info => 30 | if info.penalty ≥ 0 then 31 | return info.penalty.toNat 32 | else 33 | throwError "aesop: simp rules must be given a non-negative integer priority" 34 | | _ => throwError "aesop: simp builder can only construct 'norm' rules" 35 | 36 | namespace RuleBuilder 37 | 38 | def simpCore (decl : Name) (phase : PhaseSpec) : MetaM LocalRuleSetMember := 39 | withExceptionTransform (λ msg => m!"aesop: simp builder: exception while trying to add {decl} as a simp theorem:{indentD msg}") do 40 | let entries ← getSimpEntriesForConst decl 41 | let prio ← phase.getSimpPrio 42 | let entries := entries.map (updateSimpEntryPriority prio) 43 | let name := 44 | { name := decl, scope := .global, builder := .simp, phase := .norm } 45 | return .global $ .normSimpRule { name, entries } 46 | 47 | def simp : RuleBuilder := λ input => do 48 | if let some decl ← elabGlobalRuleIdent? input.term then 49 | simpCore decl input.phase 50 | else 51 | checkElabRuleTermForSimp input.term (isSimpAll := true) -- TODO (isSimpAll := true) correct? 52 | return .localNormSimpRule { 53 | id := ← mkFreshId 54 | simpTheorem := input.term 55 | } 56 | 57 | end Aesop.RuleBuilder 58 | -------------------------------------------------------------------------------- /Aesop/Builder/Tactic.lean: -------------------------------------------------------------------------------- 1 | /- 2 | Copyright (c) 2022 Jannis Limperg. All rights reserved. 3 | Released under Apache 2.0 license as described in the file LICENSE. 4 | Authors: Jannis Limperg 5 | -/ 6 | 7 | import Aesop.Builder.Basic 8 | 9 | open Lean Lean.Meta 10 | open Lean.Elab.Tactic (TacticM) 11 | open Lean.Parser.Tactic (tacticSeq) 12 | 13 | namespace Aesop 14 | 15 | def matchByTactic? : Term → Option (TSyntax ``tacticSeq) 16 | | `(by $ts:tacticSeq) => some ts 17 | | _ => none 18 | 19 | namespace RuleBuilder 20 | 21 | def tacticIMode (imode? : Option IndexingMode) : IndexingMode := 22 | imode?.getD .unindexed 23 | 24 | def tacticCore (t : Sum Name (TSyntax ``tacticSeq)) (imode? : Option IndexingMode) (phase : PhaseSpec) : 25 | MetaM LocalRuleSetMember := do 26 | let imode := imode?.getD .unindexed 27 | match t with 28 | | .inl decl => 29 | let type := (← getConstInfo decl).type 30 | let tac ← 31 | if ← isDefEq (mkApp (mkConst ``TacticM) (mkConst ``Unit)) type then 32 | pure $ .tacticM decl 33 | else if ← isDefEq (mkConst ``SingleRuleTac) type then 34 | pure $ .singleRuleTac decl 35 | else if ← isDefEq (mkConst ``RuleTac) type then 36 | pure $ .ruleTac decl 37 | else if ← isDefEq (mkConst ``TacGen) type then 38 | pure $ .tacGen decl 39 | else 40 | throwError "aesop: tactic builder: expected {decl} to be a tactic, i.e. to have one of these types:\n TacticM Unit\n SimpleRuleTac\n RuleTac\n TacGen\nHowever, it has type{indentExpr type}" 41 | return .global $ .base $ phase.toRule decl .tactic .global tac imode none 42 | | .inr tacticSeq => 43 | let name ← mkFreshId 44 | let tac := .tacticStx tacticSeq 45 | return .global $ .base $ phase.toRule name .tactic .global tac imode none 46 | 47 | def tactic : RuleBuilder := λ input => do 48 | let opts := input.options 49 | let t ← 50 | if input.term.raw.isIdent then 51 | .inl <$> elabGlobalRuleIdent .tactic input.term 52 | else if let some stx := matchByTactic? input.term then 53 | pure $ .inr stx 54 | else 55 | throwError "aesop: tactic builder: expected '{input.term}' to be a tactic" 56 | tacticCore t opts.indexingMode? input.phase 57 | 58 | end Aesop.RuleBuilder 59 | -------------------------------------------------------------------------------- /Aesop/Builder/Unfold.lean: -------------------------------------------------------------------------------- 1 | /- 2 | Copyright (c) 2022 Jannis Limperg. All rights reserved. 3 | Released under Apache 2.0 license as described in the file LICENSE. 4 | Authors: Jannis Limperg 5 | -/ 6 | 7 | import Aesop.Builder.Basic 8 | import Aesop.Util.Tactic.Unfold 9 | 10 | open Lean 11 | open Lean.Meta 12 | 13 | namespace Aesop.RuleBuilder 14 | 15 | -- Somewhat inefficient since `foldConsts` doesn't short-circuit. 16 | def hasConst (c : Name) (e : Expr) : Bool := 17 | e.foldConsts (init := false) λ c' acc => acc || c' == c 18 | 19 | def checkUnfoldableConst (decl : Name) : MetaM (Option Name) := 20 | withoutModifyingState do 21 | let e ← mkConstWithFreshMVarLevels decl 22 | let t := (← getConstInfo decl).type 23 | let unfoldThm? ← getUnfoldEqnFor? decl 24 | forallTelescope t λ args _ => do 25 | let testExpr := mkAppN e args 26 | let unfoldResult ← 27 | unfoldMany (if · == decl then some unfoldThm? else none) testExpr 28 | match unfoldResult with 29 | | none => 30 | throwError "Declaration '{decl}' cannot be unfolded." 31 | | some (e', _) => 32 | if hasConst decl e' then 33 | throwError "Recursive definition '{decl}' cannot be used as an unfold rule (it would be unfolded infinitely often). Try adding a simp rule for it." 34 | return unfoldThm? 35 | 36 | def unfoldCore (decl : Name) : MetaM LocalRuleSetMember := do 37 | let unfoldThm? ← checkUnfoldableConst decl 38 | return .global $ .base $ .unfoldRule { decl, unfoldThm? } 39 | 40 | -- TODO support local unfold rules 41 | def unfold : RuleBuilder := λ input => do 42 | let decl ← elabGlobalRuleIdent .unfold input.term 43 | unfoldCore decl 44 | 45 | end Aesop.RuleBuilder 46 | -------------------------------------------------------------------------------- /Aesop/BuiltinRules.lean: -------------------------------------------------------------------------------- 1 | /- 2 | Copyright (c) 2021 Jannis Limperg. All rights reserved. 3 | Released under Apache 2.0 license as described in the file LICENSE. 4 | Authors: Jannis Limperg 5 | -/ 6 | 7 | -- The Aesop.BuiltinRules.* imports are needed to ensure that the tactics from 8 | -- these files are registered. 9 | import Aesop.BuiltinRules.Assumption 10 | import Aesop.BuiltinRules.ApplyHyps 11 | import Aesop.BuiltinRules.DestructProducts 12 | import Aesop.BuiltinRules.Ext 13 | import Aesop.BuiltinRules.Intros 14 | import Aesop.BuiltinRules.Rfl 15 | import Aesop.BuiltinRules.Split 16 | import Aesop.BuiltinRules.Subst 17 | import Aesop.Frontend.Attribute 18 | 19 | namespace Aesop.BuiltinRules 20 | 21 | attribute [aesop (rule_sets := [builtin]) safe 0 apply] PUnit.unit 22 | 23 | -- Hypotheses of product type are split by a separate builtin rule because the 24 | -- `cases` builder currently cannot be used for norm rules. 25 | attribute [aesop (rule_sets := [builtin]) safe 101 constructors] 26 | And Prod PProd MProd 27 | 28 | attribute [aesop (rule_sets := [builtin]) unsafe 30% constructors] 29 | Exists Subtype Sigma PSigma 30 | 31 | -- Sums are split and introduced lazily. 32 | attribute [aesop (rule_sets := [builtin]) [safe 100 cases, 50% constructors]] 33 | Or Sum PSum 34 | 35 | -- A goal ⊢ P ↔ Q is split into ⊢ P → Q and ⊢ Q → P. Hypotheses of type `P ↔ Q` 36 | -- are treated as equations `P = Q` by the simplifier and by our builtin subst 37 | -- rule. 38 | attribute [aesop (rule_sets := [builtin]) safe 100 constructors] Iff 39 | 40 | -- A negated goal Γ ⊢ ¬ P is transformed into Γ, P ⊢ ⊥. A goal with a 41 | -- negated hypothesis Γ, h : ¬ P ⊢ Q is transformed into Γ[P := ⊥] ⊢ Q[P := ⊥] 42 | -- by the simplifier. Quantified negated hypotheses h : ∀ x : T, ¬ P x are also 43 | -- supported by the simplifier if the premises x can be discharged. 44 | @[aesop (rule_sets := [builtin]) safe 0] 45 | theorem not_intro (h : P → False) : ¬ P := h 46 | 47 | @[aesop (rule_sets := [builtin]) norm destruct] 48 | theorem empty_false (h : Empty) : False := nomatch h 49 | 50 | @[aesop (rule_sets := [builtin]) norm destruct] 51 | theorem pEmpty_false (h : PEmpty) : False := nomatch h 52 | 53 | attribute [aesop (rule_sets := [builtin]) norm constructors] ULift 54 | 55 | attribute [aesop (rule_sets := [builtin]) norm 0 destruct] ULift.down 56 | 57 | @[aesop (rule_sets := [builtin]) norm simp] 58 | theorem heq_iff_eq (x y : α) : HEq x y ↔ x = y := 59 | ⟨eq_of_heq, heq_of_eq⟩ 60 | 61 | end Aesop.BuiltinRules 62 | -------------------------------------------------------------------------------- /Aesop/BuiltinRules/ApplyHyps.lean: -------------------------------------------------------------------------------- 1 | /- 2 | Copyright (c) 2021 Jannis Limperg. All rights reserved. 3 | Released under Apache 2.0 license as described in the file LICENSE. 4 | Authors: Jannis Limperg 5 | -/ 6 | 7 | import Aesop.Frontend.Attribute 8 | 9 | namespace Aesop.BuiltinRules 10 | 11 | open Lean 12 | open Lean.Meta 13 | 14 | def applyHyp (hyp : FVarId) (goal : MVarId) (md : TransparencyMode) : 15 | BaseM RuleApplication := do 16 | let (goals, #[step]) ← applyS goal (.fvar hyp) none md |>.run 17 | | throwError "aesop: internal error in applyHyps: multiple steps" 18 | return { 19 | goals := goals.map λ mvarId => { diff := .empty goal mvarId } 20 | postState := step.postState 21 | scriptSteps? := #[step] 22 | successProbability? := none 23 | } 24 | 25 | @[aesop unsafe 75% tactic (rule_sets := [builtin])] 26 | def applyHyps : RuleTac := λ input => 27 | input.goal.withContext do 28 | let lctx ← getLCtx 29 | let md := input.options.applyHypsTransparency 30 | let mut rapps := Array.mkEmpty lctx.decls.size 31 | for localDecl in lctx do 32 | if localDecl.isImplementationDetail then continue 33 | let initialState ← saveState 34 | try 35 | let rapp ← applyHyp localDecl.fvarId input.goal md 36 | rapps := rapps.push rapp 37 | catch _ => continue 38 | finally restoreState initialState 39 | return ⟨rapps⟩ 40 | 41 | end Aesop.BuiltinRules 42 | -------------------------------------------------------------------------------- /Aesop/BuiltinRules/Assumption.lean: -------------------------------------------------------------------------------- 1 | /- 2 | Copyright (c) 2021 Jannis Limperg. All rights reserved. 3 | Released under Apache 2.0 license as described in the file LICENSE. 4 | Authors: Jannis Limperg 5 | -/ 6 | 7 | import Aesop.Frontend.Attribute 8 | 9 | open Lean 10 | open Lean.Meta 11 | 12 | namespace Aesop.BuiltinRules 13 | 14 | @[aesop safe -50 (rule_sets := [builtin])] 15 | def assumption : RuleTac := λ input => do 16 | let goal := input.goal 17 | let md := input.options.assumptionTransparency 18 | goal.withContext do 19 | goal.checkNotAssigned `Aesop.BuiltinRules.assumption 20 | goal.instantiateMVars 21 | let tgt ← goal.getType 22 | let tgtHasMVar := tgt.hasMVar 23 | let initialState ← saveState 24 | let mut applications := #[] 25 | for ldecl in ← getLCtx do 26 | if ldecl.isImplementationDetail then 27 | continue 28 | restoreState initialState 29 | let (some (application, proofHasMVar)) ← tryHyp goal ldecl.fvarId md 30 | | continue 31 | if ! tgtHasMVar && ! proofHasMVar then 32 | applications := #[application] 33 | break 34 | else 35 | applications := applications.push application 36 | if applications.isEmpty then 37 | throwTacticEx `Aesop.BuiltinRules.assumption goal "no matching assumption found" 38 | return ⟨applications⟩ 39 | where 40 | tryHyp (goal : MVarId) (fvarId : FVarId) (md : TransparencyMode) : 41 | BaseM (Option (RuleApplication × Bool)) := do 42 | let (true, steps) ← tryExactFVarS goal fvarId md |>.run 43 | | return none 44 | let #[step] := steps 45 | | throwError "aesop: internal error in assumption: multiple steps" 46 | let proofHasMVar := (← fvarId.getType).hasMVar 47 | let app := { 48 | goals := #[] 49 | postState := step.postState 50 | scriptSteps? := #[step] 51 | successProbability? := none 52 | } 53 | return some (app, proofHasMVar) 54 | 55 | end Aesop.BuiltinRules 56 | -------------------------------------------------------------------------------- /Aesop/BuiltinRules/Ext.lean: -------------------------------------------------------------------------------- 1 | /- 2 | Copyright (c) 2023 Jannis Limperg. All rights reserved. 3 | Released under Apache 2.0 license as described in the file LICENSE. 4 | Authors: Jannis Limperg 5 | -/ 6 | 7 | import Aesop.Frontend.Attribute 8 | 9 | namespace Aesop.BuiltinRules 10 | 11 | open Lean Lean.Meta 12 | 13 | def extCore (goal : MVarId) : ScriptM (Option (Array MVarId)) := 14 | saturate1 goal λ goal => do 15 | let r ← straightLineExtS goal 16 | if r.depth == 0 then 17 | return none 18 | else 19 | return r.goals.map (·.1) 20 | 21 | @[aesop 80% tactic (index := [target _ = _]) (rule_sets := [builtin])] 22 | def ext : RuleTac := RuleTac.ofSingleRuleTac λ input => do 23 | let (some goals, steps) ← extCore input.goal |>.run 24 | | throwError "found no applicable ext lemma" 25 | let goals ← goals.mapM (mvarIdToSubgoal (parentMVarId := input.goal) ·) 26 | return (goals, steps, none) 27 | 28 | end Aesop.BuiltinRules 29 | -------------------------------------------------------------------------------- /Aesop/BuiltinRules/Intros.lean: -------------------------------------------------------------------------------- 1 | /- 2 | Copyright (c) 2022 Jannis Limperg. All rights reserved. 3 | Released under Apache 2.0 license as described in the file LICENSE. 4 | Authors: Jannis Limperg, Kyle Miller 5 | -/ 6 | 7 | import Aesop.Frontend.Attribute 8 | 9 | open Lean 10 | open Lean.Meta 11 | 12 | namespace Aesop.BuiltinRules 13 | 14 | @[aesop norm -100 (rule_sets := [builtin])] 15 | def intros : RuleTac := RuleTac.ofSingleRuleTac λ input => do 16 | let md? := input.options.introsTransparency? 17 | let ((goal, newFVarIds), steps) ← 18 | match md? with 19 | | none => introsS input.goal |>.run 20 | | some md => introsUnfoldingS input.goal md |>.run 21 | if newFVarIds.size == 0 then 22 | throwError "nothing to introduce" 23 | let addedFVars := newFVarIds.foldl (init := ∅) λ set fvarId => 24 | set.insert fvarId 25 | let diff := { 26 | oldGoal := input.goal 27 | newGoal := goal 28 | addedFVars 29 | removedFVars := ∅ 30 | targetChanged := .true 31 | } 32 | return (#[{ diff }], steps, none) 33 | 34 | end Aesop.BuiltinRules 35 | -------------------------------------------------------------------------------- /Aesop/BuiltinRules/Rfl.lean: -------------------------------------------------------------------------------- 1 | /- 2 | Copyright (c) 2024 Jannis Limperg. All rights reserved. 3 | Released under Apache 2.0 license as described in the file LICENSE. 4 | Authors: Jannis Limperg 5 | -/ 6 | 7 | import Aesop.Frontend.Attribute 8 | 9 | open Lean Lean.Elab.Tactic 10 | 11 | namespace Aesop.BuiltinRules 12 | 13 | @[aesop safe 0 (rule_sets := [builtin])] 14 | def rfl : RuleTac := 15 | RuleTac.ofTacticSyntax λ _ => `(tactic| rfl) 16 | 17 | end Aesop.BuiltinRules 18 | -------------------------------------------------------------------------------- /Aesop/BuiltinRules/Split.lean: -------------------------------------------------------------------------------- 1 | /- 2 | Copyright (c) 2022 Jannis Limperg. All rights reserved. 3 | Released under Apache 2.0 license as described in the file LICENSE. 4 | Authors: Sebastian Ullrich, Jannis Limperg 5 | -/ 6 | 7 | import Aesop.Frontend.Attribute 8 | 9 | open Lean 10 | open Lean.Meta 11 | 12 | namespace Aesop.BuiltinRules 13 | 14 | @[aesop (rule_sets := [builtin]) safe 100] 15 | def splitTarget : RuleTac := RuleTac.ofSingleRuleTac λ input => do 16 | let (some goals, steps) ← splitTargetS? input.goal |>.run | throwError 17 | "nothing to split in target" 18 | let goals ← goals.mapM (mvarIdToSubgoal input.goal ·) 19 | return (goals, steps, none) 20 | 21 | partial def splitHypothesesCore (goal : MVarId) : 22 | ScriptM (Option (Array MVarId)) := 23 | withIncRecDepth do 24 | let some goals ← splitFirstHypothesisS? goal 25 | | return none 26 | let mut subgoals := #[] 27 | for g in goals do 28 | if let some subgoals' ← splitHypothesesCore g then 29 | subgoals := subgoals ++ subgoals' 30 | else 31 | subgoals := subgoals.push g 32 | return subgoals 33 | 34 | @[aesop (rule_sets := [builtin]) safe 1000] 35 | def splitHypotheses : RuleTac := RuleTac.ofSingleRuleTac λ input => do 36 | let (some goals, steps) ← splitHypothesesCore input.goal |>.run 37 | | throwError "no splittable hypothesis found" 38 | let goals ← goals.mapM (mvarIdToSubgoal input.goal ·) 39 | return (goals, steps, none) 40 | 41 | end Aesop.BuiltinRules 42 | -------------------------------------------------------------------------------- /Aesop/Check.lean: -------------------------------------------------------------------------------- 1 | /- 2 | Copyright (c) 2021 Jannis Limperg. All rights reserved. 3 | Released under Apache 2.0 license as described in the file LICENSE. 4 | Authors: Jannis Limperg 5 | -/ 6 | 7 | import Lean.Data.Options 8 | 9 | open Lean 10 | 11 | namespace Aesop 12 | 13 | initialize Check.all : Lean.Option Bool ← 14 | Option.register `aesop.check.all 15 | { defValue := false 16 | descr := "(aesop) Enable all runtime checks. Individual checks can still be disabled." 17 | group := "tactic" } 18 | 19 | structure Check where 20 | toOption : Lean.Option Bool 21 | 22 | namespace Check 23 | 24 | def get (opts : Options) (opt : Check) : Bool := 25 | if let some val := opt.toOption.get? opts then 26 | val 27 | else 28 | Check.all.get opts 29 | 30 | def isEnabled [Monad m] [MonadOptions m] (opt : Check) : m Bool := 31 | return opt.get (← getOptions) 32 | 33 | def name (opt : Check) : Name := 34 | opt.toOption.name 35 | 36 | end Check 37 | 38 | -- Inspired by `Lean.Data.Options`. 39 | local macro "register_aesop_check_option" optName:ident descr:str : command => 40 | `(initialize option : Lean.Option Bool ← 41 | Lean.Option.register $(quote $ `aesop.check ++ optName.getId) 42 | { defValue := false, descr := $descr, group := "tactic" } 43 | 44 | def $(mkIdent $ `Check ++ optName.getId) : Aesop.Check := ⟨option⟩) 45 | 46 | register_aesop_check_option proofReconstruction 47 | "(aesop) Typecheck partial proof terms during proof reconstruction." 48 | 49 | register_aesop_check_option tree 50 | "(aesop) Check search tree invariants after every iteration of the search loop. Quite expensive." 51 | 52 | register_aesop_check_option rules 53 | "(aesop) Check that information reported by rules is correct." 54 | 55 | register_aesop_check_option script 56 | "(aesop) Check that the tactic script generated by Aesop proves the goal. When this check is active, Aesop generates a tactic script even if the user did not request one." 57 | 58 | register_aesop_check_option script.steps 59 | "(aesop) Check each step of the tactic script generated by Aesop. When this check is active, Aesop generates a tactic script even if the user did not request one." 60 | 61 | end Aesop 62 | -------------------------------------------------------------------------------- /Aesop/Constants.lean: -------------------------------------------------------------------------------- 1 | /- 2 | Copyright (c) 2021 Jannis Limperg. All rights reserved. 3 | Released under Apache 2.0 license as described in the file LICENSE. 4 | Authors: Jannis Limperg 5 | -/ 6 | 7 | import Aesop.Percent 8 | 9 | namespace Aesop 10 | 11 | def unificationGoalPenalty : Percent := 12 | ⟨0.8⟩ 13 | 14 | def postponedSafeRuleSuccessProbability : Percent := 15 | ⟨0.9⟩ 16 | 17 | end Aesop 18 | -------------------------------------------------------------------------------- /Aesop/ElabM.lean: -------------------------------------------------------------------------------- 1 | /- 2 | Copyright (c) 2022 Jannis Limperg. All rights reserved. 3 | Released under Apache 2.0 license as described in the file LICENSE. 4 | Authors: Jannis Limperg 5 | -/ 6 | 7 | import Lean 8 | 9 | open Lean Lean.Meta Lean.Elab 10 | 11 | namespace Aesop.ElabM 12 | 13 | structure Context where 14 | parsePriorities : Bool 15 | goal : MVarId 16 | 17 | namespace Context 18 | 19 | def forAdditionalRules (goal : MVarId) : Context where 20 | parsePriorities := true 21 | goal := goal 22 | 23 | -- HACK: Some of the elaboration functions require that we pass in the current 24 | -- goal. The goal is used exclusively to look up fvars in the lctx, so when 25 | -- we operate outside a goal, we pass in a dummy mvar with empty lctx. 26 | def forAdditionalGlobalRules : MetaM Context := do 27 | let mvarId := (← mkFreshExprMVarAt {} {} (.const ``True [])).mvarId! 28 | return .forAdditionalRules mvarId 29 | 30 | def forErasing (goal : MVarId) : Context where 31 | parsePriorities := false 32 | goal := goal 33 | 34 | -- HACK: See `forAdditionalGlobalRules` 35 | def forGlobalErasing : MetaM Context := do 36 | let mvarId := (← mkFreshExprMVarAt {} {} (.const ``True [])).mvarId! 37 | return .forErasing mvarId 38 | 39 | end ElabM.Context 40 | 41 | 42 | abbrev ElabM := ReaderT ElabM.Context $ TermElabM 43 | 44 | -- Generate specialized pure/bind implementations so we don't need to optimise 45 | -- them on the fly at each use site. 46 | instance : Monad ElabM := 47 | { inferInstanceAs (Monad ElabM) with } 48 | 49 | protected def ElabM.run (ctx : Context) (x : ElabM α) : TermElabM α := do 50 | ReaderT.run x ctx 51 | 52 | def shouldParsePriorities : ElabM Bool := 53 | return (← read).parsePriorities 54 | 55 | def getGoal : ElabM MVarId := 56 | return (← read).goal 57 | 58 | end Aesop 59 | -------------------------------------------------------------------------------- /Aesop/Exception.lean: -------------------------------------------------------------------------------- 1 | /- 2 | Copyright (c) 2024 Jannis Limperg. All rights reserved. 3 | Released under Apache 2.0 license as described in the file LICENSE. 4 | Authors: Jannis Limperg 5 | -/ 6 | 7 | import Lean 8 | 9 | open Lean 10 | 11 | namespace Aesop 12 | 13 | scoped macro "declare_aesop_exception" 14 | excName:ident idName:ident testName:ident : command => 15 | `(initialize $idName : InternalExceptionId ← 16 | Lean.registerInternalExceptionId $(quote $ `Aesop ++ excName.getId) 17 | 18 | def $excName : Exception := 19 | .internal $idName 20 | 21 | def $testName : Exception → Bool 22 | | .internal id _ => id == $idName 23 | | _ => false) 24 | 25 | end Aesop 26 | -------------------------------------------------------------------------------- /Aesop/Forward/CompleteMatchQueue.lean: -------------------------------------------------------------------------------- 1 | import Batteries.Data.BinomialHeap.Basic 2 | import Aesop.Forward.Match 3 | 4 | set_option linter.missingDocs true 5 | 6 | namespace Aesop 7 | 8 | open Batteries (BinomialHeap) 9 | 10 | /-- A complete match queue. -/ 11 | abbrev CompleteMatchQueue := BinomialHeap ForwardRuleMatch ForwardRuleMatch.le 12 | 13 | namespace CompleteMatchQueue 14 | 15 | /-- Drop elements satisfying `f` from the front of `queue` until we reach 16 | an element that does not satisfy `f` (or until the queue is empty). -/ 17 | partial def dropInitial (queue : CompleteMatchQueue) 18 | (f : ForwardRuleMatch → Bool) : CompleteMatchQueue := 19 | match queue.deleteMin with 20 | | none => queue 21 | | some (m, queue') => 22 | if f m then 23 | dropInitial queue' f 24 | else 25 | queue 26 | 27 | end Aesop.CompleteMatchQueue 28 | -------------------------------------------------------------------------------- /Aesop/Forward/LevelIndex.lean: -------------------------------------------------------------------------------- 1 | /- 2 | Copyright (c) 2024 Jannis Limperg. All rights reserved. 3 | Released under Apache 2.0 license as described in the file LICENSE. 4 | Authors: Jannis Limperg 5 | -/ 6 | 7 | namespace Aesop 8 | 9 | structure LevelIndex where 10 | toNat : Nat 11 | deriving Inhabited, BEq, Hashable, DecidableEq, Ord 12 | 13 | instance : LT LevelIndex where 14 | lt i j := i.toNat < j.toNat 15 | 16 | instance : DecidableRel (α := LevelIndex) (· < ·) := 17 | λ i j => inferInstanceAs $ Decidable (i.toNat < j.toNat) 18 | 19 | instance : LE LevelIndex where 20 | le i j := i.toNat ≤ j.toNat 21 | 22 | instance : DecidableRel (α := LevelIndex) (· ≤ ·) := 23 | λ i j => inferInstanceAs $ Decidable (i.toNat ≤ j.toNat) 24 | 25 | instance : ToString LevelIndex where 26 | toString i := toString i.toNat 27 | 28 | end Aesop 29 | -------------------------------------------------------------------------------- /Aesop/Forward/PremiseIndex.lean: -------------------------------------------------------------------------------- 1 | /- 2 | Copyright (c) 2024 Jannis Limperg. All rights reserved. 3 | Released under Apache 2.0 license as described in the file LICENSE. 4 | Authors: Jannis Limperg 5 | -/ 6 | 7 | namespace Aesop 8 | 9 | structure PremiseIndex where 10 | toNat : Nat 11 | deriving Inhabited, BEq, Hashable, DecidableEq, Ord 12 | 13 | instance : LT PremiseIndex where 14 | lt i j := i.toNat < j.toNat 15 | 16 | instance : DecidableRel (α := PremiseIndex) (· < ·) := 17 | λ i j => inferInstanceAs $ Decidable (i.toNat < j.toNat) 18 | 19 | instance : LE PremiseIndex where 20 | le i j := i.toNat ≤ j.toNat 21 | 22 | instance : DecidableRel (α := PremiseIndex) (· ≤ ·) := 23 | λ i j => inferInstanceAs $ Decidable (i.toNat ≤ j.toNat) 24 | 25 | instance : ToString PremiseIndex where 26 | toString i := toString i.toNat 27 | 28 | end Aesop 29 | -------------------------------------------------------------------------------- /Aesop/Forward/SlotIndex.lean: -------------------------------------------------------------------------------- 1 | /- 2 | Copyright (c) 2024 Jannis Limperg. All rights reserved. 3 | Released under Apache 2.0 license as described in the file LICENSE. 4 | Authors: Jannis Limperg 5 | -/ 6 | 7 | namespace Aesop 8 | 9 | structure SlotIndex where 10 | toNat : Nat 11 | deriving Inhabited, BEq, Hashable, DecidableEq, Ord 12 | 13 | instance : LT SlotIndex where 14 | lt i j := i.toNat < j.toNat 15 | 16 | instance : DecidableRel (α := SlotIndex) (· < ·) := 17 | λ i j => inferInstanceAs $ Decidable (i.toNat < j.toNat) 18 | 19 | instance : LE SlotIndex where 20 | le i j := i.toNat ≤ j.toNat 21 | 22 | instance : DecidableRel (α := SlotIndex) (· ≤ ·) := 23 | λ i j => inferInstanceAs $ Decidable (i.toNat ≤ j.toNat) 24 | 25 | instance : HAdd SlotIndex Nat SlotIndex where 26 | hAdd i j := ⟨i.toNat + j⟩ 27 | 28 | instance : HSub SlotIndex Nat SlotIndex where 29 | hSub i j := ⟨i.toNat - j⟩ 30 | 31 | instance : ToString SlotIndex where 32 | toString i := toString i.toNat 33 | 34 | end Aesop 35 | -------------------------------------------------------------------------------- /Aesop/Forward/State/ApplyGoalDiff.lean: -------------------------------------------------------------------------------- 1 | /- 2 | Copyright (c) 2024 Jannis Limperg. All rights reserved. 3 | Released under Apache 2.0 license as described in the file LICENSE. 4 | Authors: Jannis Limperg 5 | -/ 6 | 7 | import Aesop.Forward.State 8 | import Aesop.RuleSet 9 | 10 | namespace Aesop 11 | 12 | open Lean Lean.Meta 13 | 14 | /-- Apply a goal diff to the state, adding and removing hypotheses as indicated 15 | by the diff. -/ 16 | def ForwardState.applyGoalDiff (rs : LocalRuleSet) (diff : GoalDiff) 17 | (fs : ForwardState) : BaseM (ForwardState × Array ForwardRuleMatch) := do 18 | if ! aesop.dev.statefulForward.get (← getOptions) then 19 | -- We still update the hyp types since these are also used by stateless 20 | -- forward reasoning. 21 | return ({ fs with hypTypes := ← updateHypTypes fs.hypTypes } , #[]) 22 | let fs ← diff.oldGoal.withContext do 23 | diff.removedFVars.foldM (init := fs) λ fs h => do eraseHyp h fs 24 | diff.newGoal.withContext do 25 | let (fs, ruleMatches) ← 26 | diff.addedFVars.foldM (init := (fs, ∅)) λ (fs, ruleMatches) h => do 27 | addHyp h fs ruleMatches 28 | if ← diff.targetChanged' then 29 | updateTarget fs ruleMatches 30 | else 31 | return (fs, ruleMatches) 32 | where 33 | eraseHyp (h : FVarId) (fs : ForwardState) : BaseM ForwardState := 34 | withConstAesopTraceNode .forward (return m!"erase hyp {Expr.fvar h} ({h.name})") do 35 | return fs.eraseHyp h (← rpinf (← h.getType)) 36 | 37 | addHyp (h : FVarId) (fs : ForwardState) 38 | (ruleMatches : Array ForwardRuleMatch) : 39 | BaseM (ForwardState × Array ForwardRuleMatch) := do 40 | let rules ← rs.applicableForwardRules (← h.getType) 41 | let patInsts ← rs.forwardRulePatternSubstsInLocalDecl (← h.getDecl) 42 | fs.addHypWithPatSubstsCore ruleMatches diff.newGoal h rules patInsts 43 | 44 | updateTarget (fs : ForwardState) (ruleMatches : Array ForwardRuleMatch) : 45 | BaseM (ForwardState × Array ForwardRuleMatch) := do 46 | let patInsts ← 47 | rs.forwardRulePatternSubstsInExpr (← diff.newGoal.getType) 48 | fs.updateTargetPatSubstsCore ruleMatches diff.newGoal patInsts 49 | 50 | updateHypTypes (hypTypes : PHashSet RPINF) : BaseM (PHashSet RPINF) := do 51 | let mut hypTypes := hypTypes 52 | for fvarId in diff.removedFVars do 53 | let type ← diff.oldGoal.withContext do rpinf (← fvarId.getType) 54 | hypTypes := hypTypes.erase type 55 | for fvarId in diff.addedFVars do 56 | let type ← diff.newGoal.withContext do rpinf (← fvarId.getType) 57 | hypTypes := hypTypes.insert type 58 | return hypTypes 59 | 60 | end Aesop 61 | -------------------------------------------------------------------------------- /Aesop/Forward/State/Initial.lean: -------------------------------------------------------------------------------- 1 | /- 2 | Copyright (c) 2024 Jannis Limperg. All rights reserved. 3 | Released under Apache 2.0 license as described in the file LICENSE. 4 | Authors: Jannis Limperg 5 | -/ 6 | 7 | import Aesop.Forward.State 8 | import Aesop.RuleSet 9 | 10 | open Lean Lean.Meta 11 | 12 | namespace Aesop.LocalRuleSet 13 | 14 | def mkInitialForwardState (goal : MVarId) (rs : LocalRuleSet) : 15 | BaseM (ForwardState × Array ForwardRuleMatch) := 16 | goal.withContext do 17 | if ! aesop.dev.statefulForward.get (← getOptions) then 18 | -- We still initialise the hyp types since these are also used by 19 | -- stateless forward reasoning. 20 | let mut hypTypes := ∅ 21 | for ldecl in ← getLCtx do 22 | if ! ldecl.isImplementationDetail then 23 | hypTypes := hypTypes.insert (← rpinf ldecl.type) 24 | return ({ (∅ : ForwardState) with hypTypes }, #[]) 25 | let mut fs : ForwardState := ∅ 26 | let mut ruleMatches := rs.constForwardRuleMatches 27 | aesop_trace[forward] do 28 | for m in ruleMatches do 29 | aesop_trace![forward] "match for constant rule {m.rule.name}" 30 | for ldecl in ← show MetaM _ from getLCtx do 31 | if ldecl.isImplementationDetail then 32 | continue 33 | let rules ← rs.applicableForwardRules ldecl.type 34 | let patInsts ← rs.forwardRulePatternSubstsInLocalDecl ldecl 35 | let (fs', ruleMatches') ← 36 | fs.addHypWithPatSubstsCore ruleMatches goal ldecl.fvarId rules patInsts 37 | fs := fs' 38 | ruleMatches := ruleMatches' 39 | let patInsts ← rs.forwardRulePatternSubstsInExpr (← goal.getType) 40 | fs.addPatSubstsCore ruleMatches goal patInsts 41 | 42 | end Aesop.LocalRuleSet 43 | -------------------------------------------------------------------------------- /Aesop/Frontend.lean: -------------------------------------------------------------------------------- 1 | import Aesop.Frontend.Attribute 2 | import Aesop.Frontend.Command 3 | import Aesop.Frontend.RuleExpr 4 | import Aesop.Frontend.Tactic 5 | -------------------------------------------------------------------------------- /Aesop/Frontend/Attribute.lean: -------------------------------------------------------------------------------- 1 | /- 2 | Copyright (c) 2022 Jannis Limperg. All rights reserved. 3 | Released under Apache 2.0 license as described in the file LICENSE. 4 | Authors: Jannis Limperg 5 | -/ 6 | 7 | import Aesop.Frontend.Extension 8 | import Aesop.Frontend.RuleExpr 9 | 10 | open Lean 11 | open Lean.Elab 12 | 13 | namespace Aesop.Frontend 14 | 15 | namespace Parser 16 | 17 | declare_syntax_cat Aesop.attr_rules 18 | 19 | syntax Aesop.rule_expr : Aesop.attr_rules 20 | syntax "[" Aesop.rule_expr,+,? "]" : Aesop.attr_rules 21 | 22 | syntax (name := aesop) "aesop " Aesop.attr_rules : attr 23 | 24 | end Parser 25 | 26 | structure AttrConfig where 27 | rules : Array RuleExpr 28 | deriving Inhabited 29 | 30 | namespace AttrConfig 31 | 32 | def «elab» (stx : Syntax) : TermElabM AttrConfig := 33 | withRef stx do 34 | match stx with 35 | | `(attr| aesop $e:Aesop.rule_expr) => do 36 | let r ← RuleExpr.elab e |>.run $ ← ElabM.Context.forAdditionalGlobalRules 37 | return { rules := #[r] } 38 | | `(attr| aesop [ $es:Aesop.rule_expr,* ]) => do 39 | let ctx ← ElabM.Context.forAdditionalGlobalRules 40 | let rs ← (es : Array Syntax).mapM λ e => RuleExpr.elab e |>.run ctx 41 | return { rules := rs } 42 | | _ => throwUnsupportedSyntax 43 | 44 | end AttrConfig 45 | 46 | 47 | initialize registerBuiltinAttribute { 48 | name := `aesop 49 | descr := "Register a declaration as an Aesop rule." 50 | applicationTime := .afterCompilation 51 | add := λ decl stx attrKind => withRef stx do 52 | let rules ← runTermElabMAsCoreM do 53 | let config ← AttrConfig.elab stx 54 | config.rules.flatMapM (·.buildAdditionalGlobalRules decl) 55 | for (rule, rsNames) in rules do 56 | for rsName in rsNames do 57 | addGlobalRule rsName rule attrKind (checkNotExists := true) 58 | erase := λ decl => 59 | let ruleFilter := 60 | { name := decl, scope := .global, builders := #[], phases := #[] } 61 | eraseGlobalRules RuleSetNameFilter.all ruleFilter (checkExists := true) 62 | } 63 | 64 | end Aesop.Frontend 65 | -------------------------------------------------------------------------------- /Aesop/Frontend/Basic.lean: -------------------------------------------------------------------------------- 1 | /- 2 | Copyright (c) 2023 Jannis Limperg. All rights reserved. 3 | Released under Apache 2.0 license as described in the file LICENSE. 4 | Authors: Jannis Limperg 5 | -/ 6 | 7 | import Lean 8 | 9 | open Lean Lean.Elab Lean.Elab.Term 10 | 11 | namespace Aesop.Frontend.Parser 12 | 13 | declare_syntax_cat Aesop.bool_lit (behavior := symbol) 14 | 15 | syntax "true" : Aesop.bool_lit 16 | syntax "false" : Aesop.bool_lit 17 | 18 | end Parser 19 | 20 | def elabBoolLit [Monad m] [MonadRef m] [MonadExceptOf Exception m] 21 | (stx : TSyntax `Aesop.bool_lit) : m Bool := 22 | withRef stx do 23 | match stx with 24 | | `(bool_lit| true) => return true 25 | | `(bool_lit| false) => return false 26 | | _ => throwUnsupportedSyntax 27 | 28 | end Aesop.Frontend 29 | -------------------------------------------------------------------------------- /Aesop/Frontend/Extension/Init.lean: -------------------------------------------------------------------------------- 1 | /- 2 | Copyright (c) 2023 Jannis Limperg. All rights reserved. 3 | Released under Apache 2.0 license as described in the file LICENSE. 4 | Authors: Jannis Limperg 5 | -/ 6 | 7 | import Aesop.RuleSet 8 | 9 | open Lean 10 | 11 | namespace Aesop 12 | 13 | /-- 14 | An environment extension containing an Aesop rule set. Each rule set has its 15 | own extension. 16 | -/ 17 | abbrev RuleSetExtension := 18 | SimpleScopedEnvExtension BaseRuleSetMember BaseRuleSet 19 | 20 | /-- 21 | Structure containing information about all declared Aesop rule sets. 22 | -/ 23 | structure DeclaredRuleSets where 24 | /-- 25 | The collection of declared rule sets. Each rule set has an extension, the 26 | name of the associated `SimpExtension` and the name of the associated 27 | `SimprocExtension`. The two simp extensions are expected to be declared. 28 | -/ 29 | ruleSets : Std.HashMap RuleSetName (RuleSetExtension × Name × Name) 30 | /-- 31 | The set of Aesop rule sets that are enabled by default. 32 | -/ 33 | defaultRuleSets : Std.HashSet RuleSetName 34 | deriving Inhabited 35 | 36 | instance : EmptyCollection DeclaredRuleSets := 37 | ⟨∅, ∅⟩ 38 | 39 | initialize declaredRuleSetsRef : IO.Ref DeclaredRuleSets ← 40 | IO.mkRef ∅ 41 | 42 | def getDeclaredRuleSets : IO (Std.HashMap RuleSetName (RuleSetExtension × Name × Name)) := 43 | return (← declaredRuleSetsRef.get).ruleSets 44 | 45 | def getDefaultRuleSetNames : IO (Std.HashSet Name) := 46 | return (← declaredRuleSetsRef.get).defaultRuleSets 47 | 48 | end Aesop 49 | -------------------------------------------------------------------------------- /Aesop/Index/DiscrKeyConfig.lean: -------------------------------------------------------------------------------- 1 | /- 2 | Copyright (c) 2024 Jannis Limperg. All rights reserved. 3 | Released under Apache 2.0 license as described in the file LICENSE. 4 | Authors: Jannis Limperg 5 | -/ 6 | 7 | import Lean 8 | 9 | open Lean Lean.Meta 10 | 11 | namespace Aesop 12 | 13 | /-- The configuration used by all Aesop indices. -/ 14 | -- I don't really know what I'm doing here. I'm just copying the config used 15 | -- by `simp`; see `Meta/Simp/Types.lean:mkIndexConfig`. 16 | def indexConfig : ConfigWithKey := 17 | ({ proj := .no, transparency := .reducible : Config }).toConfigWithKey 18 | 19 | def mkDiscrTreePath (e : Expr) : MetaM (Array DiscrTree.Key) := 20 | withConfigWithKey indexConfig $ DiscrTree.mkPath e 21 | 22 | def getUnify (t : DiscrTree α) (e : Expr) : MetaM (Array α) := 23 | withConfigWithKey indexConfig $ t.getUnify e 24 | 25 | def getMatch (t : DiscrTree α) (e : Expr) : MetaM (Array α) := 26 | withConfigWithKey indexConfig $ t.getMatch e 27 | 28 | end Aesop 29 | -------------------------------------------------------------------------------- /Aesop/Index/DiscrTreeConfig.lean: -------------------------------------------------------------------------------- 1 | /- 2 | Copyright (c) 2021-2024 Jannis Limperg. All rights reserved. 3 | Released under Apache 2.0 license as described in the file LICENSE. 4 | Authors: Jannis Limperg 5 | -/ 6 | 7 | import Lean 8 | 9 | namespace Aesop 10 | 11 | open Lean Lean.Meta 12 | 13 | /-- The configuration used by all Aesop indices. -/ 14 | -- I don't really know what I'm doing here. I'm just copying the config used 15 | -- by `simp`; see `Meta/Simp/Types.lean:mkIndexConfig`. 16 | def indexConfig : ConfigWithKey := 17 | ({ proj := .no, transparency := .reducible : Config }).toConfigWithKey 18 | 19 | def mkDiscrTreePath (e : Expr) : MetaM (Array DiscrTree.Key) := 20 | withConfigWithKey indexConfig $ DiscrTree.mkPath e 21 | 22 | def getUnify (t : DiscrTree α) (e : Expr) : MetaM (Array α) := 23 | withConfigWithKey indexConfig $ t.getUnify e 24 | 25 | def getMatch (t : DiscrTree α) (e : Expr) : MetaM (Array α) := 26 | withConfigWithKey indexConfig $ t.getMatch e 27 | 28 | end Aesop 29 | -------------------------------------------------------------------------------- /Aesop/Main.lean: -------------------------------------------------------------------------------- 1 | /- 2 | Copyright (c) 2021 Jannis Limperg. All rights reserved. 3 | Released under Apache 2.0 license as described in the file LICENSE. 4 | Authors: Jannis Limperg 5 | -/ 6 | 7 | import Aesop.Search.Main 8 | import Aesop.Frontend.Tactic 9 | import Aesop.Stats.Extension 10 | 11 | open Lean 12 | open Lean.Elab.Tactic 13 | 14 | namespace Aesop 15 | 16 | @[tactic Frontend.Parser.aesopTactic, tactic Frontend.Parser.aesopTactic?] 17 | def evalAesop : Tactic := λ stx => do 18 | profileitM Exception "aesop" (← getOptions) do 19 | let goal ← getMainGoal 20 | goal.withContext do 21 | let (_, stats) ← go stx goal |>.run ∅ 22 | recordStatsForCurrentFileIfEnabled stx stats 23 | stats.trace .stats 24 | where 25 | go (stx : Syntax) (goal : MVarId) : StateRefT Stats TacticM Unit := 26 | profiling (λ s _ t => { s with total := t }) do 27 | let config ← profiling (λ s _ t => { s with configParsing := t }) do 28 | Frontend.TacticConfig.parse stx goal 29 | let ruleSet ← 30 | profiling (λ s _ t => { s with ruleSetConstruction := t }) do 31 | config.getRuleSet goal 32 | withConstAesopTraceNode .ruleSet (return "Rule set") do 33 | ruleSet.trace .ruleSet 34 | profiling (λ s _ t => { s with search := t }) do 35 | let (goals, stats) ← 36 | search goal ruleSet config.options config.simpConfig 37 | config.simpConfigSyntax? (← getStats) 38 | replaceMainGoal goals.toList 39 | modifyStats λ _ => stats 40 | 41 | end Aesop 42 | -------------------------------------------------------------------------------- /Aesop/Nanos.lean: -------------------------------------------------------------------------------- 1 | /- 2 | Copyright (c) 2022 Jannis Limperg. All rights reserved. 3 | Released under Apache 2.0 license as described in the file LICENSE. 4 | Authors: Jannis Limperg 5 | -/ 6 | 7 | namespace Aesop 8 | 9 | structure Nanos where 10 | nanos : Nat 11 | deriving Inhabited, BEq, Ord 12 | 13 | namespace Nanos 14 | 15 | instance : OfNat Nanos n where 16 | ofNat := ⟨n⟩ 17 | 18 | instance : LT Nanos where 19 | lt | ⟨n₁⟩, ⟨n₂⟩ => n₁ < n₂ 20 | 21 | instance : DecidableRel (α := Nanos) (· < ·) := 22 | λ ⟨n₁⟩ ⟨n₂⟩ => inferInstanceAs (Decidable (n₁ < n₂)) 23 | 24 | instance : LE Nanos where 25 | le | ⟨n₁⟩, ⟨n₂⟩ => n₁ ≤ n₂ 26 | 27 | instance : DecidableRel (α := Nanos) (· ≤ ·) := 28 | λ ⟨n₁⟩ ⟨n₂⟩ => inferInstanceAs (Decidable (n₁ ≤ n₂)) 29 | 30 | instance : Add Nanos where 31 | add n m := ⟨n.nanos + m.nanos⟩ 32 | 33 | instance : HDiv Nanos Nat Nanos where 34 | hDiv n m := ⟨n.nanos / m⟩ 35 | 36 | def printAsMillis (n : Nanos) : String := 37 | let str := toString (n.nanos.toFloat / 1000000) 38 | match str.split λ c => c == '.' with 39 | | [beforePoint] => beforePoint ++ "ms" 40 | | [beforePoint, afterPoint] => beforePoint ++ "." ++ afterPoint.take 1 ++ "ms" 41 | | _ => unreachable! 42 | 43 | end Aesop.Nanos 44 | -------------------------------------------------------------------------------- /Aesop/Options.lean: -------------------------------------------------------------------------------- 1 | /- 2 | Copyright (c) 2023 Jannis Limperg. All rights reserved. 3 | Released under Apache 2.0 license as described in the file LICENSE. 4 | Authors: Jannis Limperg 5 | -/ 6 | 7 | import Aesop.Options.Internal 8 | import Aesop.Options.Public 9 | -------------------------------------------------------------------------------- /Aesop/Options/Internal.lean: -------------------------------------------------------------------------------- 1 | /- 2 | Copyright (c) 2023 Jannis Limperg. All rights reserved. 3 | Released under Apache 2.0 license as described in the file LICENSE. 4 | Authors: Jannis Limperg 5 | -/ 6 | 7 | import Aesop.Check 8 | import Aesop.Options.Public 9 | 10 | open Lean 11 | open Lean.Meta 12 | 13 | namespace Aesop 14 | 15 | structure Options' extends Options where 16 | generateScript : Bool 17 | forwardMaxDepth? : Option Nat 18 | deriving Inhabited 19 | 20 | def Options.toOptions' [Monad m] [MonadOptions m] (opts : Options) 21 | (forwardMaxDepth? : Option Nat := none) : m Options' := do 22 | let generateScript ← 23 | pure (aesop.dev.generateScript.get (← getOptions)) <||> 24 | pure opts.traceScript <||> 25 | Check.script.isEnabled <||> 26 | Check.script.steps.isEnabled 27 | return { opts with generateScript, forwardMaxDepth? } 28 | 29 | end Aesop 30 | -------------------------------------------------------------------------------- /Aesop/Percent.lean: -------------------------------------------------------------------------------- 1 | /- 2 | Copyright (c) 2021 Jannis Limperg. All rights reserved. 3 | Released under Apache 2.0 license as described in the file LICENSE. 4 | Authors: Jannis Limperg, Asta Halkjær From 5 | -/ 6 | 7 | namespace Aesop 8 | 9 | open Lean 10 | open Std 11 | 12 | /- 13 | Invariant: between 0 and 1.0 14 | -/ 15 | structure Percent where 16 | toFloat : Float 17 | deriving Inhabited 18 | 19 | namespace Percent 20 | 21 | protected def ofFloat (f : Float) : Option Percent := 22 | if 0 <= f && f <= 1.0 then some ⟨f⟩ else none 23 | 24 | instance : Mul Percent where 25 | mul p q := ⟨p.toFloat * q.toFloat⟩ 26 | 27 | @[inline] 28 | def δ : Percent := 29 | ⟨0.00001⟩ 30 | 31 | instance : BEq Percent where 32 | beq | ⟨p⟩, ⟨q⟩ => if p > q then p - q < δ.toFloat else q - p < δ.toFloat 33 | 34 | instance : Ord Percent where 35 | compare p q := 36 | if p == q then Ordering.eq 37 | else if p.toFloat < q.toFloat then Ordering.lt 38 | else Ordering.gt 39 | 40 | -- NOTE: This is not the same as 41 | -- 42 | -- p < q := p.toFloat < q.toFloat 43 | -- 44 | -- That definition would not agree with the Ord instance. 45 | instance : LT Percent := 46 | ltOfOrd 47 | 48 | instance : LE Percent := 49 | leOfOrd 50 | 51 | instance : ToString Percent where 52 | toString p := toString p.toFloat 53 | 54 | instance : HPow Percent Nat Percent where 55 | hPow | ⟨p⟩, n => ⟨p ^ n.toFloat⟩ 56 | 57 | def hundred : Percent := 58 | ⟨1⟩ 59 | 60 | def fifty : Percent := 61 | ⟨0.5⟩ 62 | 63 | def toHumanString (p : Percent) : String := 64 | let str := toString (p.toFloat * 100) 65 | match str.split λ c => c == '.' with 66 | | [beforePoint] => beforePoint ++ "%" 67 | | [beforePoint, afterPoint] => 68 | beforePoint ++ "." ++ afterPoint.take 4 ++ "%" 69 | | _ => unreachable! 70 | 71 | protected def ofNat (n : Nat) : Option Percent := 72 | Percent.ofFloat $ n.toFloat / 100 73 | 74 | end Aesop.Percent 75 | -------------------------------------------------------------------------------- /Aesop/RPINF.lean: -------------------------------------------------------------------------------- 1 | import Aesop.Util.Basic 2 | import Aesop.BaseM 3 | 4 | open Lean Lean.Meta 5 | 6 | namespace Aesop 7 | 8 | local instance : MonadCache Expr Expr BaseM where 9 | findCached? e := 10 | return (← MonadCache.findCached? e : Option RPINFRaw).map (·.toExpr) 11 | cache k v := MonadCache.cache k (⟨v⟩ : RPINFRaw) 12 | 13 | @[specialize] 14 | partial def rpinfRaw (e : Expr) : BaseM RPINFRaw := 15 | withReducible do return ⟨← go e⟩ 16 | where 17 | go (e : Expr) : BaseM Expr := 18 | withIncRecDepth do 19 | checkCache e λ _ => do 20 | if ← isProof e then 21 | return .mdata (mdataSetIsProof {}) e 22 | let e ← whnf e 23 | match e with 24 | | .app .. => 25 | let f ← go e.getAppFn' 26 | let mut args := e.getAppArgs' 27 | for i in [:args.size] do 28 | let arg := args[i]! 29 | args := args.set! i default -- prevent nonlinear access to args[i] 30 | let arg ← go arg 31 | args := args.set! i arg 32 | if f.isConstOf ``Nat.succ && args.size == 1 && args[0]!.isRawNatLit then 33 | return mkRawNatLit (args[0]!.rawNatLit?.get! + 1) 34 | else 35 | return mkAppN f args 36 | | .lam .. => 37 | -- TODO disable cache? 38 | lambdaTelescope e λ xs e => withNewFVars xs do 39 | mkLambdaFVars xs (← go e) 40 | | .forallE .. => 41 | -- TODO disable cache? 42 | forallTelescope e λ xs e => withNewFVars xs do 43 | mkForallFVars xs (← go e) 44 | | .proj t i e => 45 | return .proj t i (← go e) 46 | | .sort .. | .mvar .. | .lit .. | .const .. | .fvar .. => 47 | return e 48 | | .letE .. | .mdata .. | .bvar .. => unreachable! 49 | 50 | withNewFVars {α} (fvars : Array Expr) (k : BaseM α) : BaseM α := do 51 | let mut lctx ← (getLCtx : MetaM _) 52 | for fvar in fvars do 53 | let fvarId := fvar.fvarId! 54 | let ldecl ← fvarId.getDecl 55 | let ldecl := ldecl.setType $ ← go ldecl.type 56 | lctx := lctx.modifyLocalDecl fvarId λ _ => ldecl 57 | withLCtx lctx (← getLocalInstances) k 58 | 59 | def rpinf (e : Expr) : BaseM RPINF := 60 | withConstAesopTraceNode .rpinf (return m!"rpinf") do 61 | aesop_trace[rpinf] "input:{indentExpr e}" 62 | let e ← rpinfRaw e 63 | let hash := pinfHash e.toExpr 64 | aesop_trace[rpinf] "result hash: {hash}" 65 | aesop_trace[rpinf] "resut:{indentExpr e.toExpr}" 66 | return { e with hash } 67 | 68 | end Aesop 69 | -------------------------------------------------------------------------------- /Aesop/Rule/Basic.lean: -------------------------------------------------------------------------------- 1 | /- 2 | Copyright (c) 2021 Jannis Limperg. All rights reserved. 3 | Released under Apache 2.0 license as described in the file LICENSE. 4 | Authors: Jannis Limperg 5 | -/ 6 | 7 | import Aesop.Index.Basic 8 | import Aesop.Rule.Name 9 | import Aesop.RuleTac.Basic 10 | import Aesop.RuleTac.Descr 11 | 12 | open Lean 13 | 14 | namespace Aesop 15 | 16 | structure Rule (α : Type) where 17 | name : RuleName 18 | indexingMode : IndexingMode 19 | pattern? : Option RulePattern 20 | extra : α 21 | tac : RuleTacDescr 22 | deriving Inhabited 23 | 24 | namespace Rule 25 | 26 | instance : BEq (Rule α) where 27 | beq r s := r.name == s.name 28 | 29 | instance : Ord (Rule α) where 30 | compare r s := compare r.name s.name 31 | 32 | instance : Hashable (Rule α) where 33 | hash r := hash r.name 34 | 35 | def compareByPriority [Ord α] (r s : Rule α) : Ordering := 36 | compare r.extra s.extra 37 | 38 | def compareByName (r s : Rule α) : Ordering := 39 | r.name.compare s.name 40 | 41 | def compareByPriorityThenName [Ord α] (r s : Rule α) : Ordering := 42 | compareByPriority r s |>.then $ compareByName r s 43 | 44 | @[inline] 45 | protected def map (f : α → β) (r : Rule α) : Rule β := 46 | { r with extra := f r.extra } 47 | 48 | @[inline] 49 | protected def mapM [Monad m] (f : α → m β) (r : Rule α) : m (Rule β) := 50 | return { r with extra := ← f r.extra } 51 | 52 | end Aesop.Rule 53 | -------------------------------------------------------------------------------- /Aesop/RulePattern/Cache.lean: -------------------------------------------------------------------------------- 1 | /- 2 | Copyright (c) 2024 Jannis Limperg. All rights reserved. 3 | Released under Apache 2.0 license as described in the file LICENSE. 4 | Authors: Jannis Limperg 5 | -/ 6 | 7 | import Aesop.Forward.Substitution 8 | import Aesop.Rule.Name 9 | 10 | open Lean 11 | 12 | set_option linter.missingDocs true 13 | 14 | namespace Aesop 15 | 16 | /-- Entry of the rule pattern cache. -/ 17 | def RulePatternCache.Entry := Array (RuleName × Substitution) 18 | 19 | set_option linter.missingDocs false in 20 | /-- A cache for the rule pattern index. -/ 21 | structure RulePatternCache where 22 | map : Std.HashMap Expr RulePatternCache.Entry 23 | deriving Inhabited 24 | 25 | instance : EmptyCollection RulePatternCache := 26 | ⟨⟨∅⟩⟩ 27 | 28 | end Aesop 29 | -------------------------------------------------------------------------------- /Aesop/RuleSet/Filter.lean: -------------------------------------------------------------------------------- 1 | /- 2 | Copyright (c) 2021-2024 Jannis Limperg. All rights reserved. 3 | Released under Apache 2.0 license as described in the file LICENSE. 4 | Authors: Jannis Limperg 5 | -/ 6 | 7 | import Aesop.RuleSet.Name 8 | import Aesop.Rule.Name 9 | 10 | open Lean 11 | 12 | namespace Aesop 13 | 14 | structure RuleFilter where 15 | name : Name 16 | scope : ScopeName 17 | /-- 18 | `#[]` means 'match any builder' 19 | -/ 20 | builders : Array BuilderName 21 | /-- 22 | `#[]` means 'match any phase' 23 | -/ 24 | phases : Array PhaseName 25 | 26 | namespace RuleFilter 27 | 28 | def matchesPhase (f : RuleFilter) (p : PhaseName) : Bool := 29 | f.phases.isEmpty || f.phases.contains p 30 | 31 | def matchesBuilder (f : RuleFilter) (b : BuilderName) : Bool := 32 | f.builders.isEmpty || f.builders.contains b 33 | 34 | def «matches» (f : RuleFilter) (n : RuleName) : Bool := 35 | f.name == n.name && 36 | f.scope == n.scope && 37 | f.matchesPhase n.phase && 38 | f.matchesBuilder n.builder 39 | 40 | def matchesSimpTheorem? (f : RuleFilter) : Option Name := 41 | if f.scope == .global && f.matchesBuilder .simp then 42 | some f.name 43 | else 44 | none 45 | 46 | /-- 47 | Returns the identifier of the local norm simp rule matched by `f`, if any. 48 | -/ 49 | def matchesLocalNormSimpRule? (f : RuleFilter) : Option Name := Id.run do 50 | if f.scope == .local && f.matchesBuilder .simp then 51 | return some f.name 52 | return none 53 | 54 | end RuleFilter 55 | 56 | 57 | structure RuleSetNameFilter where 58 | ns : Array RuleSetName -- #[] means 'match any rule set' 59 | 60 | namespace RuleSetNameFilter 61 | 62 | protected def all : RuleSetNameFilter := 63 | ⟨#[]⟩ 64 | 65 | def matchesAll (f : RuleSetNameFilter) : Bool := 66 | f.ns.isEmpty 67 | 68 | def «matches» (f : RuleSetNameFilter) (n : RuleSetName) : Bool := 69 | f.matchesAll || f.ns.contains n 70 | 71 | def matchedRuleSetNames (f : RuleSetNameFilter) : Option (Array RuleSetName) := 72 | if f.matchesAll then none else some f.ns 73 | 74 | end Aesop.RuleSetNameFilter 75 | -------------------------------------------------------------------------------- /Aesop/RuleSet/Member.lean: -------------------------------------------------------------------------------- 1 | /- 2 | Copyright (c) 2021-2024 Jannis Limperg. All rights reserved. 3 | Released under Apache 2.0 license as described in the file LICENSE. 4 | Authors: Jannis Limperg 5 | -/ 6 | 7 | import Aesop.Rule 8 | import Aesop.Rule.Forward 9 | 10 | namespace Aesop 11 | 12 | inductive BaseRuleSetMember 13 | | normRule (r : NormRule) 14 | | unsafeRule (r : UnsafeRule) 15 | | safeRule (r : SafeRule) 16 | | unfoldRule (r : UnfoldRule) 17 | | normForwardRule (r₁ : ForwardRule) (r₂ : NormRule) 18 | | unsafeForwardRule (r₁ : ForwardRule) (r₂ : UnsafeRule) 19 | | safeForwardRule (r₁ : ForwardRule) (r₂ : SafeRule) 20 | deriving Inhabited 21 | 22 | def BaseRuleSetMember.name : BaseRuleSetMember → RuleName 23 | | normRule r => r.name 24 | | unsafeRule r => r.name 25 | | safeRule r => r.name 26 | | unfoldRule r => r.name 27 | | normForwardRule r _ => r.name 28 | | unsafeForwardRule r _ => r.name 29 | | safeForwardRule r _ => r.name 30 | 31 | inductive GlobalRuleSetMember 32 | | base (m : BaseRuleSetMember) 33 | | normSimpRule (e : NormSimpRule) 34 | deriving Inhabited 35 | 36 | def GlobalRuleSetMember.name : GlobalRuleSetMember → RuleName 37 | | base m => m.name 38 | | normSimpRule r => r.name 39 | 40 | inductive LocalRuleSetMember 41 | | global (m : GlobalRuleSetMember) 42 | | localNormSimpRule (r : LocalNormSimpRule) 43 | deriving Inhabited 44 | 45 | def LocalRuleSetMember.name : LocalRuleSetMember → RuleName 46 | | global m => m.name 47 | | localNormSimpRule r => r.name 48 | 49 | def LocalRuleSetMember.toGlobalRuleSetMember? : 50 | LocalRuleSetMember → Option GlobalRuleSetMember 51 | | global m => some m 52 | | _ => none 53 | 54 | end Aesop 55 | -------------------------------------------------------------------------------- /Aesop/RuleSet/Name.lean: -------------------------------------------------------------------------------- 1 | /- 2 | Copyright (c) 2021-2024 Jannis Limperg. All rights reserved. 3 | Released under Apache 2.0 license as described in the file LICENSE. 4 | Authors: Jannis Limperg 5 | -/ 6 | 7 | import Lean 8 | 9 | open Lean 10 | 11 | namespace Aesop 12 | 13 | abbrev RuleSetName := Name -- Not really an abbreviation is it? 14 | 15 | def defaultRuleSetName : RuleSetName := `default 16 | 17 | def builtinRuleSetName : RuleSetName := `builtin 18 | 19 | def localRuleSetName : RuleSetName := `local 20 | 21 | def builtinRuleSetNames : Array RuleSetName := 22 | #[defaultRuleSetName, builtinRuleSetName] 23 | 24 | def RuleSetName.isReserved (n : RuleSetName) : Bool := 25 | n == localRuleSetName || builtinRuleSetNames.contains n 26 | 27 | end Aesop 28 | -------------------------------------------------------------------------------- /Aesop/RuleTac.lean: -------------------------------------------------------------------------------- 1 | /- 2 | Copyright (c) 2022 Jannis Limperg. All rights reserved. 3 | Released under Apache 2.0 license as described in the file LICENSE. 4 | Authors: Jannis Limperg 5 | -/ 6 | 7 | import Aesop.RuleTac.Apply 8 | import Aesop.RuleTac.Basic 9 | import Aesop.RuleTac.Cases 10 | import Aesop.RuleTac.Forward 11 | import Aesop.RuleTac.Preprocess 12 | import Aesop.RuleTac.Tactic 13 | 14 | open Lean 15 | 16 | namespace Aesop.RuleTacDescr 17 | 18 | protected def run : RuleTacDescr → RuleTac 19 | | apply t md => RuleTac.apply t md 20 | | constructors cs md => RuleTac.applyConsts cs md 21 | | forward t immediate clear => RuleTac.forward t immediate clear 22 | | cases target md isRecursiveType ctorNames => 23 | RuleTac.cases target md isRecursiveType ctorNames 24 | | tacticM decl => RuleTac.tacticM decl 25 | | singleRuleTac decl => RuleTac.singleRuleTac decl 26 | | ruleTac decl => RuleTac.ruleTac decl 27 | | tacticStx stx => RuleTac.tacticStx stx 28 | | tacGen decl => RuleTac.tacGen decl 29 | | preprocess => RuleTac.preprocess 30 | | forwardMatches m => RuleTac.forwardMatches m 31 | 32 | end RuleTacDescr 33 | -------------------------------------------------------------------------------- /Aesop/RuleTac/Descr.lean: -------------------------------------------------------------------------------- 1 | import Aesop.RuleTac.Basic 2 | import Aesop.Forward.Match.Types 3 | 4 | open Lean Lean.Meta 5 | 6 | namespace Aesop 7 | 8 | inductive RuleTacDescr 9 | | apply (term : RuleTerm) (md : TransparencyMode) 10 | | constructors (constructorNames : Array Name) (md : TransparencyMode) 11 | | forward (term : RuleTerm) (immediate : UnorderedArraySet PremiseIndex) 12 | (isDestruct : Bool) 13 | | cases (target : CasesTarget) (md : TransparencyMode) 14 | (isRecursiveType : Bool) (ctorNames : Array CtorNames) 15 | | tacticM (decl : Name) 16 | | ruleTac (decl : Name) 17 | | tacGen (decl : Name) 18 | | singleRuleTac (decl : Name) 19 | | tacticStx (stx : Syntax) 20 | | preprocess 21 | | forwardMatches (ms : Array ForwardRuleMatch) 22 | deriving Inhabited 23 | 24 | namespace RuleTacDescr 25 | 26 | def forwardRuleMatches? : RuleTacDescr → Option (Array ForwardRuleMatch) 27 | | forwardMatches ms => ms 28 | | _ => none 29 | 30 | end RuleTacDescr 31 | 32 | end Aesop 33 | -------------------------------------------------------------------------------- /Aesop/RuleTac/FVarIdSubst.lean: -------------------------------------------------------------------------------- 1 | /- 2 | Copyright (c) 2024 Jannis Limperg. All rights reserved. 3 | Released under Apache 2.0 license as described in the file LICENSE. 4 | Authors: Jannis Limperg 5 | -/ 6 | 7 | import Lean 8 | 9 | open Lean Lean.Meta 10 | 11 | namespace Aesop 12 | 13 | structure FVarIdSubst where 14 | map : Std.HashMap FVarId FVarId 15 | deriving Inhabited 16 | 17 | namespace FVarIdSubst 18 | 19 | def isEmpty (s : FVarIdSubst) : Bool := 20 | s.map.isEmpty 21 | 22 | def contains (s : FVarIdSubst) (fvarId : FVarId) : Bool := 23 | s.map.contains fvarId 24 | 25 | def find? (s : FVarIdSubst) (fvarId : FVarId) : Option FVarId := 26 | s.map[fvarId]? 27 | 28 | def get (s : FVarIdSubst) (fvarId : FVarId) : FVarId := 29 | s.find? fvarId |>.getD fvarId 30 | 31 | def apply (s : FVarIdSubst) (e : Expr) : Expr := 32 | if s.isEmpty || ! e.hasFVar then 33 | e 34 | else 35 | e.replace λ 36 | | .fvar fvarId => s.find? fvarId |>.map mkFVar 37 | | _ => none 38 | 39 | def applyToLocalDecl (s : FVarIdSubst) : LocalDecl → LocalDecl 40 | | .cdecl i id n t bi k => .cdecl i id n (s.apply t) bi k 41 | | .ldecl i id n t v nd k => .ldecl i id n (s.apply t) (s.apply v) nd k 42 | 43 | def domain (s : FVarIdSubst) : Std.HashSet FVarId := 44 | s.map.fold (init := ∅) λ r k _ => r.insert k 45 | 46 | def codomain (s : FVarIdSubst) : Std.HashSet FVarId := 47 | s.map.fold (init := ∅) λ r _ v => r.insert v 48 | 49 | protected def empty : FVarIdSubst := 50 | ⟨∅⟩ 51 | 52 | instance : EmptyCollection FVarIdSubst := 53 | ⟨.empty⟩ 54 | 55 | def insert (s : FVarIdSubst) (old new : FVarId) : FVarIdSubst := 56 | let map : Std.HashMap _ _ := s.map.fold (init := ∅) λ map k v => 57 | map.insert k $ if v == old then new else v 58 | ⟨map.insert old new⟩ 59 | 60 | def erase (s : FVarIdSubst) (fvarId : FVarId) : FVarIdSubst := 61 | ⟨s.map.erase fvarId⟩ 62 | 63 | def append (s t : FVarIdSubst) : FVarIdSubst := 64 | let map : Std.HashMap _ _ := s.map.fold (init := ∅) λ map k v => 65 | map.insert k $ t.get v 66 | ⟨t.map.fold (init := map) λ s k v => s.insert k v⟩ 67 | 68 | def ofFVarSubstIgnoringNonFVarIds (s : FVarSubst) : FVarIdSubst := .mk $ 69 | s.map.foldl (init := ∅) λ map k v => 70 | if let .fvar fvarId := v then map.insert k fvarId else map 71 | 72 | def ofFVarSubst? (s : FVarSubst) : Option FVarIdSubst := Id.run do 73 | let mut result := ∅ 74 | for (k, v) in s.map do 75 | if let .fvar fvarId := v then 76 | result := result.insert k fvarId 77 | else 78 | return none 79 | return some result 80 | 81 | end Aesop.FVarIdSubst 82 | -------------------------------------------------------------------------------- /Aesop/RuleTac/Preprocess.lean: -------------------------------------------------------------------------------- 1 | /- 2 | Copyright (c) 2023 Jannis Limperg. All rights reserved. 3 | Released under Apache 2.0 license as described in the file LICENSE. 4 | Authors: Jannis Limperg 5 | -/ 6 | 7 | import Aesop.RuleTac.Basic 8 | import Aesop.Script.SpecificTactics 9 | 10 | open Lean Lean.Meta 11 | 12 | namespace Aesop.RuleTac 13 | 14 | /-- 15 | This `RuleTac` is applied once to the root goal, before any other rules are 16 | tried. 17 | -/ 18 | def preprocess : RuleTac := RuleTac.ofSingleRuleTac λ input => do 19 | let ((mvarId, _), steps) ← renameInaccessibleFVarsS input.goal |>.run 20 | return (#[{ diff := .empty input.goal mvarId }], steps, none) 21 | 22 | end Aesop.RuleTac 23 | -------------------------------------------------------------------------------- /Aesop/RuleTac/RuleTerm.lean: -------------------------------------------------------------------------------- 1 | import Aesop.Rule.Name 2 | 3 | open Lean Lean.Meta 4 | 5 | namespace Aesop 6 | 7 | inductive RuleTerm 8 | | const (decl : Name) 9 | | term (term : Term) 10 | deriving Inhabited 11 | 12 | instance : ToMessageData RuleTerm where 13 | toMessageData 14 | | .const decl => m!"{decl}" 15 | | .term tm => m!"{tm}" 16 | 17 | inductive ElabRuleTerm 18 | | const (decl : Name) 19 | | term (term : Term) (expr : Expr) 20 | deriving Inhabited 21 | 22 | namespace ElabRuleTerm 23 | 24 | instance : ToMessageData ElabRuleTerm where 25 | toMessageData 26 | | .const decl => m!"{decl}" 27 | | .term tm _ => m!"{tm}" 28 | 29 | def expr : ElabRuleTerm → MetaM Expr 30 | | const decl => mkConstWithFreshMVarLevels decl 31 | | term _ e => return e 32 | 33 | def scope : ElabRuleTerm → ScopeName 34 | | const .. => .global 35 | | term .. => .local 36 | 37 | def name : ElabRuleTerm → MetaM Name 38 | | const decl => return decl 39 | | term _ e => getRuleNameForExpr e 40 | 41 | def toRuleTerm : ElabRuleTerm → RuleTerm 42 | | const decl => .const decl 43 | | term t _ => .term t 44 | 45 | def ofElaboratedTerm (tm : Term) (expr : Expr) : ElabRuleTerm := 46 | if let some decl := expr.constName? then 47 | .const decl 48 | else 49 | .term tm expr 50 | 51 | end Aesop.ElabRuleTerm 52 | -------------------------------------------------------------------------------- /Aesop/Script/Check.lean: -------------------------------------------------------------------------------- 1 | /- 2 | Copyright (c) 2024 Jannis Limperg. All rights reserved. 3 | Released under Apache 2.0 license as described in the file LICENSE. 4 | Authors: Jannis Limperg 5 | -/ 6 | 7 | import Aesop.Check 8 | import Aesop.Script.UScript 9 | 10 | open Lean Lean.Elab.Tactic 11 | open Lean.Parser.Tactic (tacticSeq) 12 | 13 | namespace Aesop.Script 14 | 15 | def UScript.checkIfEnabled (uscript : UScript) : MetaM Unit := do 16 | unless ← Check.script.steps.isEnabled do 17 | return 18 | try 19 | uscript.validate 20 | catch e => 21 | throwError "{Check.script.steps.name}: {e.toMessageData}" 22 | 23 | end Script 24 | 25 | def checkRenderedScriptIfEnabled (script : TSyntax ``tacticSeq) 26 | (preState : Meta.SavedState) (goal : MVarId) (expectCompleteProof : Bool) : 27 | MetaM Unit := do 28 | if ! (← Check.script.isEnabled) then 29 | return 30 | let go : TacticM Unit := do 31 | setGoals [goal] 32 | evalTactic script 33 | if expectCompleteProof && ! (← getUnsolvedGoals).isEmpty then 34 | throwError "script executed successfully but did not solve the main goal" 35 | try 36 | show MetaM Unit from withoutModifyingState do 37 | preState.restore 38 | go.run { elaborator := .anonymous, recover := false } 39 | |>.run' { goals := [goal] } 40 | |>.run' 41 | catch e => throwError 42 | "{Check.script.name}: error while executing generated script:{indentD e.toMessageData}" 43 | 44 | end Aesop 45 | -------------------------------------------------------------------------------- /Aesop/Script/CtorNames.lean: -------------------------------------------------------------------------------- 1 | import Aesop.Util.Basic 2 | 3 | open Lean Lean.Meta 4 | 5 | namespace Aesop 6 | 7 | structure CtorNames where 8 | ctor : Name 9 | args : Array Name 10 | /-- Whether the constructor has at least one implicit argument. -/ 11 | hasImplicitArg : Bool 12 | 13 | namespace CtorNames 14 | 15 | def toRCasesPat (cn : CtorNames) : TSyntax `rcasesPat := Unhygienic.run do 16 | let ns ← cn.args.mapM λ n => 17 | `(Lean.Parser.Tactic.rcasesPatLo| $(mkIdent n):ident) 18 | if cn.hasImplicitArg then 19 | `(rcasesPat| @⟨ $ns,* ⟩) 20 | else 21 | `(rcasesPat| ⟨ $ns,* ⟩) 22 | 23 | private def nameBase : Name → Name 24 | | .anonymous => .anonymous 25 | | .str _ s => .str .anonymous s 26 | | .num _ n => .num .anonymous n 27 | 28 | open Lean.Parser.Tactic in 29 | def toInductionAltLHS (cn : CtorNames) : 30 | TSyntax ``inductionAltLHS := Unhygienic.run do 31 | let ns := cn.args.map mkIdent 32 | let ctor := mkIdent $ nameBase cn.ctor 33 | if cn.hasImplicitArg then 34 | `(inductionAltLHS| | @$ctor $ns:ident*) 35 | else 36 | `(inductionAltLHS| | $ctor $ns:ident*) 37 | 38 | open Lean.Parser.Tactic in 39 | def toInductionAlt (cn : CtorNames) (tacticSeq : Array Syntax.Tactic) : 40 | TSyntax ``inductionAlt := Unhygienic.run do 41 | `(inductionAlt| $(cn.toInductionAltLHS):inductionAltLHS => $tacticSeq:tactic*) 42 | 43 | def toAltVarNames (cn : CtorNames) : AltVarNames where 44 | explicit := true 45 | varNames := cn.args.toList 46 | 47 | def mkFreshArgNames (lctx : LocalContext) (cn : CtorNames) : 48 | CtorNames × LocalContext := 49 | let (args, lctx) := getUnusedNames lctx cn.args 50 | ({ cn with args }, lctx) 51 | 52 | end CtorNames 53 | 54 | open Lean.Parser.Tactic in 55 | def ctorNamesToRCasesPats (cns : Array CtorNames) : TSyntax ``rcasesPatMed := 56 | Unhygienic.run do `(rcasesPatMed| $(cns.map (·.toRCasesPat)):rcasesPat|*) 57 | 58 | open Lean.Parser.Tactic in 59 | def ctorNamesToInductionAlts (cns : Array (CtorNames × Array Syntax.Tactic)) : 60 | TSyntax ``inductionAlts := Unhygienic.run do 61 | let alts := cns.map λ (cn, tacticSeq) => cn.toInductionAlt tacticSeq 62 | `(inductionAlts| with $alts:inductionAlt*) 63 | 64 | def mkCtorNames (iv : InductiveVal) : CoreM (Array CtorNames) := MetaM.run' do 65 | iv.ctors.toArray.mapM λ ctor => do 66 | let cv ← getConstInfoCtor ctor 67 | let arity := cv.numParams + cv.numFields 68 | forallBoundedTelescope cv.type arity λ args _ => do 69 | let mut result := Array.mkEmpty cv.numFields 70 | let mut hasImplicitArg := false 71 | for arg in args[cv.numParams:] do 72 | let ldecl ← arg.fvarId!.getDecl 73 | result := result.push ldecl.userName 74 | hasImplicitArg := hasImplicitArg || ! ldecl.binderInfo.isExplicit 75 | return { args := result, ctor, hasImplicitArg } 76 | 77 | end Aesop 78 | -------------------------------------------------------------------------------- /Aesop/Script/GoalWithMVars.lean: -------------------------------------------------------------------------------- 1 | import Lean 2 | 3 | open Lean Std Lean.Meta 4 | 5 | namespace Aesop 6 | 7 | structure GoalWithMVars where 8 | goal : MVarId 9 | mvars : Std.HashSet MVarId 10 | deriving Inhabited 11 | 12 | instance : Repr GoalWithMVars where 13 | reprPrec 14 | | g, _ => s!"\{ goal := {repr g.goal}, mvars := {repr g.mvars.toArray} }" 15 | 16 | instance : BEq GoalWithMVars := 17 | ⟨λ g₁ g₂ => g₁.goal == g₂.goal⟩ 18 | 19 | namespace GoalWithMVars 20 | 21 | def ofMVarId (goal : MVarId) : MetaM GoalWithMVars := do 22 | return { goal, mvars := ← goal.getMVarDependencies } 23 | 24 | end Aesop.GoalWithMVars 25 | -------------------------------------------------------------------------------- /Aesop/Script/OptimizeSyntax.lean: -------------------------------------------------------------------------------- 1 | /- 2 | Copyright (c) 2024 Jannis Limperg. All rights reserved. 3 | Released under Apache 2.0 license as described in the file LICENSE. 4 | Authors: Jannis Limperg 5 | -/ 6 | 7 | import Lean 8 | 9 | namespace Aesop 10 | 11 | open Lean Lean.Meta Lean.Parser.Tactic 12 | 13 | variable [Monad m] [MonadQuotation m] 14 | 15 | partial def optimizeFocusRenameI : Syntax → m Syntax 16 | | stx@.missing | stx@(.atom ..) | stx@(.ident ..) => return stx 17 | | .node info kind args => do 18 | let stx := .node info kind (← args.mapM optimizeFocusRenameI) 19 | match stx with 20 | | `(tactic| · $tacs:tactic*) 21 | | `(tactic| · { $tacs:tactic* }) => do 22 | let tacs := tacs.getElems 23 | let some (tac : TSyntax `tactic) := tacs[0]? 24 | | return stx 25 | match tac with 26 | | `(tactic| rename_i $[$ns:ident]*) => 27 | `(tactic| next $[$ns:ident]* => $(tacs[1:]):tactic*) 28 | | _ => return stx 29 | | _ => return stx 30 | 31 | private partial def addIdents (acc : Std.HashSet Name) : Syntax → Std.HashSet Name 32 | | .missing | .atom .. => acc 33 | | .ident (val := val) .. => acc.insert val 34 | | .node _ _ args => 35 | args.foldl (init := acc) λ acc stx => addIdents acc stx 36 | 37 | def optimizeInitialRenameI : Syntax → m Syntax 38 | | stx@`(tacticSeq| $tacs:tactic*) => do 39 | let tacs := tacs.getElems 40 | let some (tac : TSyntax `tactic) := tacs[0]? 41 | | return stx 42 | match tac with 43 | | `(tactic| rename_i $[$ns:ident]*) => 44 | let usedNames := tacs[1:].foldl (init := ∅) λ usedNames stx => 45 | addIdents usedNames stx.raw 46 | let mut dropUntil := 0 47 | for n in ns do 48 | if usedNames.contains n.getId then 49 | break 50 | else 51 | dropUntil := dropUntil + 1 52 | if dropUntil == 0 then 53 | return stx 54 | else if dropUntil == ns.size then 55 | tacsToTacticSeq tacs[1:] 56 | else 57 | let ns : TSyntaxArray `ident := ns[dropUntil:].toArray 58 | let tac ← `(tactic| rename_i $[$ns:ident]*) 59 | let mut result : Array (TSyntax `tactic) := Array.mkEmpty tacs.size 60 | result := result.push tac 61 | result := result ++ tacs[1:] 62 | tacsToTacticSeq result 63 | | _ => return stx 64 | | stx => return stx 65 | where 66 | -- Inlining this helper function triggers a code gen issue: 67 | -- https://github.com/leanprover/lean4/issues/4548 68 | tacsToTacticSeq (tacs : Array (TSyntax `tactic)) : m (TSyntax ``tacticSeq) := 69 | `(tacticSeq| $tacs:tactic*) 70 | 71 | def optimizeSyntax (stx : TSyntax kind) : m (TSyntax kind) := do 72 | let mut stx := stx.raw 73 | stx ← optimizeFocusRenameI stx 74 | stx ← optimizeInitialRenameI stx 75 | return ⟨stx⟩ 76 | 77 | end Aesop 78 | -------------------------------------------------------------------------------- /Aesop/Script/SScript.lean: -------------------------------------------------------------------------------- 1 | import Aesop.Script.Step 2 | 3 | open Lean Lean.Meta 4 | 5 | variable [Monad m] [MonadError m] [MonadQuotation m] 6 | 7 | namespace Aesop.Script 8 | 9 | inductive SScript 10 | | empty 11 | | onGoal (goalPos : Nat) (step : Step) (tail : SScript) 12 | | focusAndSolve (goalPos : Nat) (here tail : SScript) 13 | deriving Inhabited 14 | 15 | namespace SScript 16 | 17 | def takeNConsecutiveFocusAndSolve? (acc : Array SScript) : 18 | Nat → SScript → Option (Array SScript × SScript) 19 | | 0, tail => some (acc, tail) 20 | | _ + 1, empty => none 21 | | _ + 1, onGoal .. => none 22 | | n + 1, focusAndSolve 0 here tail => 23 | takeNConsecutiveFocusAndSolve? (acc.push here) n tail 24 | | _ + 1, focusAndSolve (_ + 1) .. => none 25 | 26 | partial def render (script : SScript) : m (Array Syntax.Tactic) := do 27 | go #[] script 28 | where 29 | go (acc : Array Syntax.Tactic) : 30 | SScript → m (Array Syntax.Tactic) 31 | | empty => return acc 32 | | onGoal goalPos step tail => do 33 | if let some (tactic, tail) ← renderSTactic? goalPos step tail then 34 | let script := acc.push tactic 35 | go script tail 36 | else 37 | let script := acc.push $ mkOnGoal goalPos step.uTactic 38 | go script tail 39 | | focusAndSolve goalPos here tail => do 40 | let nestedScript ← go #[] here 41 | let t ← 42 | if goalPos == 0 then 43 | `(tactic| · $[$nestedScript:tactic]*) 44 | else 45 | let posLit := mkOneBasedNumLit goalPos 46 | `(tactic| on_goal $posLit:num => { $nestedScript:tactic* }) 47 | go (acc.push t) tail 48 | 49 | renderSTactic? (goalPos : Nat) (step : Step) (tail : SScript) : 50 | m (Option (Syntax.Tactic × SScript)) := do 51 | let some sTactic := step.sTactic? 52 | | return none 53 | let some (nested, tail) := 54 | takeNConsecutiveFocusAndSolve? #[] sTactic.numSubgoals tail 55 | | return none 56 | let nestedScripts ← nested.mapM (go #[]) 57 | let tactic := mkOnGoal goalPos $ sTactic.run nestedScripts 58 | return (tactic, tail) 59 | 60 | end Aesop.Script.SScript 61 | -------------------------------------------------------------------------------- /Aesop/Script/ScriptM.lean: -------------------------------------------------------------------------------- 1 | /- 2 | Copyright (c) 2024 Jannis Limperg. All rights reserved. 3 | Released under Apache 2.0 license as described in the file LICENSE. 4 | Authors: Jannis Limperg 5 | -/ 6 | 7 | import Aesop.BaseM 8 | import Aesop.Script.Step 9 | import Aesop.Script.Tactic 10 | 11 | open Lean Aesop.Script 12 | 13 | namespace Aesop 14 | 15 | abbrev ScriptT m := StateRefT' IO.RealWorld (Array LazyStep) m 16 | 17 | namespace ScriptT 18 | 19 | protected def run [Monad m] [MonadLiftT (ST IO.RealWorld) m] (x : ScriptT m α) : 20 | m (α × Array LazyStep) := 21 | StateRefT'.run x #[] 22 | 23 | end ScriptT 24 | 25 | abbrev ScriptM := ScriptT BaseM 26 | 27 | variable [MonadStateOf (Array LazyStep) m] 28 | 29 | def recordScriptStep (step : LazyStep) : m Unit := 30 | modify (·.push step) 31 | 32 | def recordScriptSteps (steps : Array LazyStep) : m Unit := 33 | modify (· ++ steps) 34 | 35 | def withScriptStep (preGoal : MVarId) (postGoals : α → Array MVarId) 36 | (success : α → Bool) (tacticBuilder : α → TacticBuilder) (x : MetaM α) : 37 | ScriptM α := do 38 | let preState ← show MetaM _ from saveState 39 | let a ← x 40 | if success a then 41 | let postState ← show MetaM _ from saveState 42 | recordScriptStep { 43 | tacticBuilders := #[tacticBuilder a] 44 | postGoals := postGoals a 45 | preGoal, preState, postState 46 | } 47 | return a 48 | 49 | def withOptScriptStep (preGoal : MVarId) (postGoals : α → Array MVarId) 50 | (tacticBuilder : α → TacticBuilder) (x : MetaM (Option α)) : 51 | ScriptM (Option α) := do 52 | let preState ← show MetaM _ from saveState 53 | let some a ← x 54 | | return none 55 | let postState ← show MetaM _ from saveState 56 | recordScriptStep { 57 | tacticBuilders := #[tacticBuilder a] 58 | postGoals := postGoals a 59 | preGoal, preState, postState 60 | } 61 | return some a 62 | 63 | end Aesop 64 | -------------------------------------------------------------------------------- /Aesop/Script/Tactic.lean: -------------------------------------------------------------------------------- 1 | /- 2 | Copyright (c) 2024 Jannis Limperg. All rights reserved. 3 | Released under Apache 2.0 license as described in the file LICENSE. 4 | Authors: Jannis Limperg 5 | -/ 6 | 7 | import Lean 8 | 9 | open Lean 10 | 11 | namespace Aesop.Script 12 | 13 | abbrev UTactic := Syntax.Tactic 14 | 15 | structure STactic where 16 | numSubgoals : Nat 17 | run : Array (Array Syntax.Tactic) → Syntax.Tactic 18 | 19 | structure Tactic where 20 | uTactic : UTactic 21 | sTactic? : Option STactic 22 | 23 | namespace Tactic 24 | 25 | instance : ToMessageData Tactic where 26 | toMessageData t := toMessageData t.uTactic 27 | 28 | def unstructured (uTactic : UTactic) : Tactic where 29 | uTactic := uTactic 30 | sTactic? := none 31 | 32 | def structured (uTactic : UTactic) (sTactic : STactic) : Tactic where 33 | uTactic := uTactic 34 | sTactic? := some sTactic 35 | 36 | end Tactic 37 | 38 | abbrev TacticBuilder := MetaM Tactic 39 | 40 | end Aesop.Script 41 | -------------------------------------------------------------------------------- /Aesop/Script/UScript.lean: -------------------------------------------------------------------------------- 1 | /- 2 | Copyright (c) 2024 Jannis Limperg. All rights reserved. 3 | Released under Apache 2.0 license as described in the file LICENSE. 4 | Authors: Jannis Limperg 5 | -/ 6 | 7 | import Aesop.Script.Step 8 | 9 | open Lean Lean.Meta 10 | open Lean.Parser.Tactic (tacticSeq) 11 | 12 | namespace Aesop.Script 13 | 14 | abbrev UScript := Array Step 15 | 16 | namespace UScript 17 | 18 | variable [Monad m] [MonadError m] [MonadQuotation m] in 19 | def render (tacticState : TacticState) (s : UScript) : 20 | m (Array Syntax.Tactic) := do 21 | let mut script := Array.mkEmpty s.size 22 | let mut tacticState := tacticState 23 | for step in s do 24 | let (script', tacticState') ← step.render script tacticState 25 | script := script' 26 | tacticState := tacticState' 27 | return script 28 | 29 | def renderTacticSeq (uscript : UScript) (preState : Meta.SavedState) 30 | (goal : MVarId) : MetaM (TSyntax ``tacticSeq) := do 31 | let tacticState ← preState.runMetaM' $ Script.TacticState.mkInitial goal 32 | `(tacticSeq| $(← uscript.render tacticState):tactic*) 33 | 34 | def validate (s : UScript) : MetaM Unit := 35 | s.forM (·.validate) 36 | 37 | end Aesop.Script.UScript 38 | -------------------------------------------------------------------------------- /Aesop/Script/Util.lean: -------------------------------------------------------------------------------- 1 | /- 2 | Copyright (c) 2024 Jannis Limperg. All rights reserved. 3 | Released under Apache 2.0 license as described in the file LICENSE. 4 | Authors: Jannis Limperg 5 | -/ 6 | 7 | import Aesop.Util.Basic 8 | import Aesop.Util.EqualUpToIds 9 | import Batteries.Lean.Meta.SavedState 10 | 11 | open Lean Std Lean.Meta 12 | 13 | namespace Aesop.Script 14 | 15 | -- Without {α β : Type} universe inference goes haywire. 16 | @[specialize] 17 | def findFirstStep? {α β : Type} (goals : Array α) (step? : α → Option β) 18 | (stepOrder : β → Nat) : Option (Nat × α × β) := Id.run do 19 | let mut firstStep? := none 20 | for h : pos in [:goals.size] do 21 | let g := goals[pos] 22 | if let some step := step? g then 23 | if let some (_, _, currentFirstStep) := firstStep? then 24 | if stepOrder step < stepOrder currentFirstStep then 25 | firstStep? := some (pos, g, step) 26 | else 27 | firstStep? := some (pos, g, step) 28 | else 29 | continue 30 | return firstStep? 31 | 32 | def matchGoals (postState₁ postState₂ : Meta.SavedState) 33 | (goals₁ goals₂ : Array MVarId) : MetaM (Option (Std.HashMap MVarId MVarId)) := do 34 | let goals₁ ← getProperGoals postState₁ goals₁ 35 | let goals₂ ← getProperGoals postState₂ goals₂ 36 | let (equal, s) ← 37 | tacticStatesEqualUpToIds' none postState₁.meta.mctx 38 | postState₂.meta.mctx goals₁ goals₂ (allowAssignmentDiff := true) 39 | if ! equal then 40 | return none 41 | else 42 | return s.equalMVarIds 43 | where 44 | getProperGoals (state : Meta.SavedState) (goals : Array MVarId) : 45 | MetaM (Array MVarId) := 46 | state.runMetaM' do 47 | let (properGoals, _) ← partitionGoalsAndMVars id goals 48 | return properGoals.map (·.fst) 49 | 50 | end Aesop.Script 51 | -------------------------------------------------------------------------------- /Aesop/Search/Expansion/Basic.lean: -------------------------------------------------------------------------------- 1 | /- 2 | Copyright (c) 2023 Jannis Limperg. All rights reserved. 3 | Released under Apache 2.0 license as described in the file LICENSE. 4 | Authors: Jannis Limperg 5 | -/ 6 | 7 | import Aesop.RuleTac.Basic 8 | 9 | open Lean 10 | open Lean.Meta 11 | 12 | namespace Aesop 13 | 14 | def runRuleTac (tac : RuleTac) (ruleName : RuleName) 15 | (preState : Meta.SavedState) (input : RuleTacInput) : 16 | BaseM (Except Exception RuleTacOutput) := do 17 | let result ← 18 | tryCatchRuntimeEx 19 | (Except.ok <$> runInMetaState preState do 20 | withAtMostMaxHeartbeats input.options.maxRuleHeartbeats do 21 | tac input) 22 | (λ e => return .error e) 23 | if ← Check.rules.isEnabled then 24 | if let .ok ruleOutput := result then 25 | ruleOutput.applications.forM λ rapp => do 26 | if let (some err) ← rapp.check input then 27 | throwError "{Check.rules.name}: while applying rule {ruleName}: {err}" 28 | return result 29 | 30 | end Aesop 31 | -------------------------------------------------------------------------------- /Aesop/Search/Queue/Class.lean: -------------------------------------------------------------------------------- 1 | /- 2 | Copyright (c) 2022 Jannis Limperg. All rights reserved. 3 | Released under Apache 2.0 license as described in the file LICENSE. 4 | Authors: Jannis Limperg 5 | -/ 6 | 7 | import Aesop.Tree.Data 8 | 9 | open Lean 10 | 11 | namespace Aesop 12 | 13 | class Queue (Q : Type) where 14 | init : BaseIO Q 15 | addGoals : Q → Array GoalRef → BaseIO Q 16 | popGoal : Q → BaseIO (Option GoalRef × Q) 17 | 18 | namespace Queue 19 | 20 | def init' [Queue Q] (grefs : Array GoalRef) : BaseIO Q := do 21 | addGoals (← init) grefs 22 | 23 | end Aesop.Queue 24 | -------------------------------------------------------------------------------- /Aesop/Search/RuleSelection.lean: -------------------------------------------------------------------------------- 1 | /- 2 | Copyright (c) 2022 Jannis Limperg. All rights reserved. 3 | Released under Apache 2.0 license as described in the file LICENSE. 4 | Authors: Jannis Limperg 5 | -/ 6 | import Aesop.Tree.RunMetaM 7 | import Aesop.Search.SearchM 8 | 9 | open Lean 10 | 11 | namespace Aesop 12 | 13 | variable [Aesop.Queue Q] 14 | 15 | def selectNormRules (rs : LocalRuleSet) (fms : ForwardRuleMatches) 16 | (goal : MVarId) : BaseM (Array (IndexMatchResult NormRule)) := 17 | profilingRuleSelection do rs.applicableNormalizationRules fms goal 18 | 19 | def preprocessRule : SafeRule where 20 | name := { 21 | name := `Aesop.BuiltinRule.preprocess 22 | builder := .tactic 23 | phase := .safe 24 | scope := .global 25 | } 26 | indexingMode := .unindexed 27 | pattern? := none 28 | extra := { penalty := 0, safety := .safe } 29 | tac := .preprocess 30 | 31 | def selectSafeRules (g : Goal) : 32 | SearchM Q (Array (IndexMatchResult SafeRule)) := do 33 | profilingRuleSelection do 34 | if ← g.isRoot then 35 | return #[{ 36 | rule := preprocessRule 37 | locations := ∅ 38 | patternSubsts? := none 39 | }] 40 | let ruleSet := (← read).ruleSet 41 | g.runMetaMInPostNormState' λ postNormGoal => 42 | ruleSet.applicableSafeRules g.forwardRuleMatches postNormGoal 43 | 44 | def selectUnsafeRules (postponedSafeRules : Array PostponedSafeRule) 45 | (gref : GoalRef) : SearchM Q UnsafeQueue := do 46 | profilingRuleSelection do 47 | let g ← gref.get 48 | match g.unsafeQueue? with 49 | | some rules => return rules 50 | | none => do 51 | let ruleSet := (← read).ruleSet 52 | let unsafeRules ← 53 | g.runMetaMInPostNormState' λ postNormGoal => 54 | ruleSet.applicableUnsafeRules g.forwardRuleMatches postNormGoal 55 | let unsafeQueue := UnsafeQueue.initial postponedSafeRules unsafeRules 56 | gref.set $ g.setUnsafeRulesSelected true |>.setUnsafeQueue unsafeQueue 57 | return unsafeQueue 58 | 59 | end Aesop 60 | -------------------------------------------------------------------------------- /Aesop/Stats/Extension.lean: -------------------------------------------------------------------------------- 1 | /- 2 | Copyright (c) 2024 Jannis Limperg. All rights reserved. 3 | Released under Apache 2.0 license as described in the file LICENSE. 4 | Authors: Jannis Limperg 5 | -/ 6 | 7 | import Aesop.Stats.Basic 8 | 9 | open Lean 10 | 11 | namespace Aesop 12 | 13 | structure StatsExtensionEntry where 14 | /-- 15 | The Aesop call for which stats were collected. 16 | -/ 17 | aesopStx : Syntax 18 | /- 19 | The file in which Aesop was called. 20 | -/ 21 | fileName : String 22 | /- 23 | The position in the file where Aesop was called. 24 | -/ 25 | position? : Option Position 26 | /-- 27 | The collected stats. 28 | -/ 29 | stats : Stats 30 | 31 | namespace StatsExtensionEntry 32 | 33 | def forCurrentFile [Monad m] [MonadLog m] (stx : Syntax) (stats : Stats) : 34 | m StatsExtensionEntry := do 35 | let fileName ← getFileName 36 | let fileMap ← getFileMap 37 | let position? := stx.getPos?.map fileMap.toPosition 38 | return { aesopStx := stx, stats, fileName, position? } 39 | 40 | end StatsExtensionEntry 41 | 42 | abbrev StatsExtension := SimplePersistentEnvExtension StatsExtensionEntry Unit 43 | 44 | def StatsExtension.importedEntries (env : Environment) (ext : StatsExtension) : 45 | Array (Array StatsExtensionEntry) := 46 | ext.toEnvExtension.getState env |>.importedEntries 47 | 48 | initialize statsExtension : StatsExtension ← 49 | registerSimplePersistentEnvExtension { 50 | addEntryFn := λ _ _ => () 51 | addImportedFn := λ _ => () 52 | asyncMode := .sync -- Maybe we could use a less restrictive mode, but (a) it's not clear to me and (b) this extension is unused by default. 53 | } 54 | 55 | def recordStatsIfEnabled [Monad m] [MonadEnv m] [MonadOptions m] 56 | (s : StatsExtensionEntry) : m Unit := do 57 | if ← isStatsCollectionEnabled then 58 | modifyEnv λ env => statsExtension.addEntry env s 59 | 60 | def recordStatsForCurrentFileIfEnabled [Monad m] [MonadEnv m] [MonadOptions m] 61 | [MonadLog m] (aesopStx : Syntax) (stats : Stats) : m Unit := do 62 | if ← isStatsCollectionEnabled then 63 | let entry ← StatsExtensionEntry.forCurrentFile aesopStx stats 64 | modifyEnv λ env => statsExtension.addEntry env entry 65 | 66 | abbrev StatsArray := Array StatsExtensionEntry 67 | 68 | def mkStatsArray (localEntries : List StatsExtensionEntry) 69 | (importedEntries : Array (Array StatsExtensionEntry)) : 70 | StatsArray := Id.run do 71 | let mut result := #[] 72 | for entry in localEntries do 73 | result := result.push entry 74 | for entries in importedEntries do 75 | result := result ++ entries 76 | return result 77 | 78 | def getStatsArray [Monad m] [MonadEnv m] : m StatsArray := do 79 | let env ← getEnv 80 | let current := statsExtension.getEntries env 81 | let imported := statsExtension.importedEntries env 82 | return mkStatsArray current imported 83 | 84 | end Aesop 85 | -------------------------------------------------------------------------------- /Aesop/Tree.lean: -------------------------------------------------------------------------------- 1 | /- 2 | Copyright (c) 2022 Jannis Limperg. All rights reserved. 3 | Released under Apache 2.0 license as described in the file LICENSE. 4 | Authors: Jannis Limperg 5 | -/ 6 | 7 | import Aesop.Tree.AddRapp 8 | import Aesop.Tree.Check 9 | import Aesop.Tree.Data 10 | import Aesop.Tree.ExtractProof 11 | import Aesop.Tree.ExtractScript 12 | import Aesop.Tree.Free 13 | import Aesop.Tree.RunMetaM 14 | import Aesop.Tree.State 15 | import Aesop.Tree.Tracing 16 | import Aesop.Tree.Traversal 17 | import Aesop.Tree.TreeM 18 | import Aesop.Tree.UnsafeQueue 19 | -------------------------------------------------------------------------------- /Aesop/Tree/Free.lean: -------------------------------------------------------------------------------- 1 | /- 2 | Copyright (c) 2022 Jannis Limperg. All rights reserved. 3 | Released under Apache 2.0 license as described in the file LICENSE. 4 | Authors: Jannis Limperg 5 | -/ 6 | 7 | import Aesop.Tree.TreeM 8 | 9 | namespace Aesop 10 | 11 | -- In Lean 4, cylic structures -- such as our trees with their parent pointers 12 | -- -- are not freed automatically. This is because the runtime uses reference 13 | -- counting and a parent node and its child, holding references to each other, 14 | -- will always have a reference count ≥ 1. So in order to free a tree, we must 15 | -- break the cycles by removing the parent pointers. 16 | 17 | mutual 18 | variable 19 | (dummyGoalRef : GoalRef) 20 | (dummyMVarClusterRef : MVarClusterRef) 21 | 22 | private partial def freeGoalRef (gref : GoalRef) : BaseIO Unit := do 23 | gref.modify λ g => g.setParent dummyMVarClusterRef 24 | (← gref.get).children.forM freeRappRef 25 | 26 | private partial def freeRappRef (rref : RappRef) : BaseIO Unit := do 27 | rref.modify λ r => r.setParent dummyGoalRef 28 | (← rref.get).children.forM freeMVarClusterRef 29 | 30 | private partial def freeMVarClusterRef (cref : MVarClusterRef) : 31 | BaseIO Unit := do 32 | cref.modify λ c => c.setParent none 33 | (← cref.get).goals.forM freeGoalRef 34 | end 35 | 36 | private def mkDummyRefs : BaseIO (GoalRef × MVarClusterRef) := do 37 | let cref ← IO.mkRef $ by refine' MVarCluster.mk {..} <;> exact default 38 | let gref ← IO.mkRef $ by refine' Goal.mk { parent := cref, .. } <;> exact default 39 | return (gref, cref) 40 | 41 | def GoalRef.free (gref : GoalRef) : BaseIO Unit := do 42 | let (dgref, dcref) ← mkDummyRefs 43 | freeGoalRef dgref dcref gref 44 | 45 | def RappRef.free (rref : RappRef) : BaseIO Unit := do 46 | let (dgref, dcref) ← mkDummyRefs 47 | freeRappRef dgref dcref rref 48 | 49 | def MVarClusterRef.free (cref : MVarClusterRef) : BaseIO Unit := do 50 | let (dgref, dcref) ← mkDummyRefs 51 | freeMVarClusterRef dgref dcref cref 52 | 53 | def freeTree : TreeM Unit := do 54 | (← getThe Tree).root.free 55 | 56 | end Aesop 57 | -------------------------------------------------------------------------------- /Aesop/Tree/UnsafeQueue.lean: -------------------------------------------------------------------------------- 1 | /- 2 | Copyright (c) 2022 Jannis Limperg. All rights reserved. 3 | Released under Apache 2.0 license as described in the file LICENSE. 4 | Authors: Jannis Limperg 5 | -/ 6 | 7 | import Aesop.Constants 8 | import Aesop.Rule 9 | 10 | open Lean 11 | 12 | namespace Aesop 13 | 14 | structure PostponedSafeRule where 15 | rule : SafeRule 16 | output : RuleTacOutput 17 | deriving Inhabited 18 | 19 | namespace PostponedSafeRule 20 | 21 | def toUnsafeRule (r : PostponedSafeRule) : UnsafeRule := 22 | { r.rule with extra := ⟨postponedSafeRuleSuccessProbability⟩ } 23 | 24 | end PostponedSafeRule 25 | 26 | 27 | inductive UnsafeQueueEntry 28 | | unsafeRule (r : IndexMatchResult UnsafeRule) 29 | | postponedSafeRule (r : PostponedSafeRule) 30 | deriving Inhabited 31 | 32 | namespace UnsafeQueueEntry 33 | 34 | instance : ToString UnsafeQueueEntry where 35 | toString 36 | | unsafeRule r => toString r.rule.name 37 | | postponedSafeRule r => toString r.rule.name 38 | 39 | def successProbability : UnsafeQueueEntry → Percent 40 | | unsafeRule r => r.rule.extra.successProbability 41 | | postponedSafeRule .. => postponedSafeRuleSuccessProbability 42 | 43 | def name : UnsafeQueueEntry → RuleName 44 | | unsafeRule r => r.rule.name 45 | | postponedSafeRule r => r.rule.name 46 | 47 | instance : Ord UnsafeQueueEntry := 48 | ⟨ compareLex 49 | (λ x y => compareOn successProbability x y |>.swap) 50 | (compareOn name) ⟩ 51 | 52 | end UnsafeQueueEntry 53 | 54 | 55 | def UnsafeQueue := Subarray UnsafeQueueEntry 56 | 57 | namespace UnsafeQueue 58 | 59 | instance : EmptyCollection UnsafeQueue := 60 | inferInstanceAs $ EmptyCollection (Subarray _) 61 | 62 | instance : Inhabited UnsafeQueue := 63 | inferInstanceAs $ Inhabited (Subarray _) 64 | 65 | -- Precondition: `unsafeRules` is ordered lexicographically by descending 66 | -- success probability, then by rule name. 67 | def initial (postponedSafeRules : Array PostponedSafeRule) 68 | (unsafeRules : Array (IndexMatchResult UnsafeRule)) : UnsafeQueue := 69 | let unsafeRules := unsafeRules.map .unsafeRule 70 | let postponedSafeRules := 71 | postponedSafeRules.map .postponedSafeRule 72 | |>.qsort (λ x y => compare x y |>.isLT) 73 | postponedSafeRules.mergeDedup unsafeRules |>.toSubarray 74 | 75 | def entriesToMessageData (q : UnsafeQueue) : Array MessageData := 76 | q.toArray.map toMessageData 77 | 78 | end UnsafeQueue 79 | 80 | end Aesop 81 | -------------------------------------------------------------------------------- /Aesop/Util/Tactic.lean: -------------------------------------------------------------------------------- 1 | /- 2 | Copyright (c) 2022 Jannis Limperg. All rights reserved. 3 | Released under Apache 2.0 license as described in the file LICENSE. 4 | Authors: Jannis Limperg 5 | -/ 6 | 7 | import Lean 8 | import Batteries.Tactic.OpenPrivate 9 | 10 | open Lean Lean.Meta Lean.Elab.Tactic 11 | 12 | namespace Aesop 13 | 14 | /-- 15 | A `MetaM` version of the `replace` tactic. If `fvarId` refers to the 16 | hypothesis `h`, this tactic asserts a new hypothesis `h : type` with proof 17 | `proof : type` and then tries to clear `fvarId`. Unlike `replaceLocalDecl`, 18 | `replaceFVar` always adds the new hypothesis at the end of the local context. 19 | 20 | `replaceFVar` returns the new goal, the `FVarId` of the newly asserted 21 | hypothesis and whether the old hypothesis was cleared. 22 | -/ 23 | def replaceFVar (goal : MVarId) (fvarId : FVarId) (type : Expr) (proof : Expr) : 24 | MetaM (MVarId × FVarId × Bool) := do 25 | let userName ← goal.withContext $ fvarId.getUserName 26 | let preClearGoal ← goal.assert userName type proof 27 | let goal ← preClearGoal.tryClear fvarId 28 | let clearSuccess := preClearGoal != goal 29 | let (newFVarId, goal) ← intro1Core goal (preserveBinderNames := true) 30 | return (goal, newFVarId, clearSuccess) 31 | 32 | /-- Introduce as many binders as possible while unfolding definitions with the 33 | ambient transparency. -/ 34 | partial def introsUnfolding (mvarId : MVarId) : MetaM (Array FVarId × MVarId) := 35 | run mvarId #[] 36 | where 37 | run (mvarId : MVarId) (fvars : Array FVarId) : MetaM (Array FVarId × MVarId) := 38 | mvarId.withContext do 39 | let type ← whnf (← mvarId.getType) 40 | let size := getIntrosSize type 41 | if 0 < size then 42 | let (fvars', mvarId') ← mvarId.introN size 43 | run mvarId' (fvars ++ fvars') 44 | else 45 | return (fvars, mvarId) 46 | 47 | end Aesop 48 | -------------------------------------------------------------------------------- /Aesop/Util/Tactic/Ext.lean: -------------------------------------------------------------------------------- 1 | /- 2 | Copyright (c) 2024 Jannis Limperg. All rights reserved. 3 | Released under Apache 2.0 license as described in the file LICENSE. 4 | Authors: Jannis Limperg 5 | -/ 6 | 7 | import Aesop.Tracing 8 | 9 | open Lean Lean.Meta Lean.Elab.Tactic.Ext 10 | 11 | namespace Aesop 12 | 13 | structure ExtResult where 14 | depth : Nat 15 | commonFVarIds : Array FVarId 16 | goals : Array (MVarId × Array FVarId) 17 | 18 | partial def straightLineExt (goal : MVarId) : MetaM ExtResult := 19 | go goal 0 #[] 20 | where 21 | go (goal : MVarId) (depth : Nat) (commonFVarIds : Array FVarId) : 22 | MetaM ExtResult:= do 23 | withConstAesopTraceNode .debug (return m!"straightLineExt") do 24 | aesop_trace[debug] "goal:{indentD goal}" 25 | let goals? ← show MetaM _ from observing? $ applyExtTheoremAt goal 26 | if let some goals := goals? then 27 | aesop_trace[debug] "ext lemma applied; new goals:{indentD $ Elab.goalsToMessageData goals}" 28 | let goals := goals.toArray 29 | if goals.isEmpty then 30 | return { depth := depth + 1, goals := #[], commonFVarIds := #[] } 31 | let goals ← goals.mapM λ goal => do 32 | let (fvarIds, goal) ← goal.intros 33 | return (goal, fvarIds) 34 | if h : goals.size = 1 then 35 | let (goal, fvarIds) := goals[0] 36 | go goal (depth + 1) (commonFVarIds ++ fvarIds) 37 | else 38 | return { depth := depth + 1, goals, commonFVarIds } 39 | else 40 | aesop_trace[debug] "no applicable ext lemma" 41 | return { depth, goals := #[(goal, #[])], commonFVarIds } 42 | 43 | def straightLineExtProgress (goal : MVarId) : MetaM ExtResult := do 44 | let result ← straightLineExt goal 45 | if result.depth == 0 then 46 | throwError "no applicable extensionality theorem found" 47 | return result 48 | 49 | end Aesop 50 | -------------------------------------------------------------------------------- /AesopTest/10.lean: -------------------------------------------------------------------------------- 1 | /- 2 | Copyright (c) 2022 Jannis Limperg. All rights reserved. 3 | Released under Apache 2.0 license as described in the file LICENSE. 4 | Authors: Asta H. From, Jannis Limperg 5 | -/ 6 | import Aesop 7 | 8 | set_option aesop.check.all true 9 | 10 | attribute [aesop unsafe [50% constructors, 50% cases]] List.Mem 11 | 12 | theorem Mem.map (f : α → β) (x : α) (xs : List α) (h : x ∈ xs) : 13 | f x ∈ xs.map f := by 14 | induction h <;> aesop 15 | -------------------------------------------------------------------------------- /AesopTest/12.lean: -------------------------------------------------------------------------------- 1 | /- 2 | Copyright (c) 2022 Jannis Limperg. All rights reserved. 3 | Released under Apache 2.0 license as described in the file LICENSE. 4 | Authors: Jannis Limperg 5 | -/ 6 | import Aesop 7 | 8 | set_option aesop.check.all true 9 | 10 | attribute [aesop unsafe 50% constructors] List.Mem 11 | 12 | @[aesop safe [constructors, cases (cases_patterns := [All _ [], All _ (_ :: _)])]] 13 | inductive All (P : α → Prop) : List α → Prop where 14 | | none : All P [] 15 | | more (x xs) : P x → All P xs → All P (x :: xs) 16 | 17 | theorem weaken (P Q : α → Prop) (wk : ∀ x, P x → Q x) (xs : List α) (h : All P xs) 18 | : All Q xs := by 19 | induction h <;> aesop 20 | 21 | theorem in_self (xs : List α) : All (· ∈ xs) xs := by 22 | induction xs 23 | case nil => 24 | aesop 25 | case cons x xs ih => 26 | have wk : ∀ a, a ∈ xs → a ∈ x :: xs := by aesop 27 | have ih' : All (fun a => a ∈ x :: xs) xs := by aesop (add unsafe 1% weaken) 28 | aesop 29 | -------------------------------------------------------------------------------- /AesopTest/125.lean: -------------------------------------------------------------------------------- 1 | import Aesop 2 | 3 | set_option aesop.check.all true 4 | set_option aesop.smallErrorMessages true 5 | 6 | @[aesop 90%] 7 | def myTacGen : Aesop.TacGen := fun _ => do 8 | return #[("exact ⟨val - f { val := val, property := property }, fun a ha => by simpa⟩", 9 | 0.9)] 10 | 11 | /-- error: tactic 'aesop' failed, made no progress -/ 12 | #guard_msgs in 13 | theorem foo (f : { x // 0 < x } → { x // 0 < x }) (val : Nat) 14 | (property : 0 < val) : 15 | ∃ w x, ∀ (a : Nat) (b : 0 < a), ↑(f { val := a, property := b }) = w * a + x := by 16 | constructor 17 | aesop 18 | -------------------------------------------------------------------------------- /AesopTest/126.lean: -------------------------------------------------------------------------------- 1 | import Aesop 2 | 3 | set_option aesop.check.all true 4 | 5 | -- TODO simp_all introduces spurious metavariables in this example 6 | set_option aesop.check.rules false 7 | set_option aesop.check.tree false 8 | 9 | theorem foo {a : Nat → Nat} (ha : a 0 = 37) : 10 | (match h : a 0 with | 42 => by simp_all | n => n) = 37 := by 11 | aesop 12 | -------------------------------------------------------------------------------- /AesopTest/13.lean: -------------------------------------------------------------------------------- 1 | /- 2 | Copyright (c) 2022 Asta H. From. All rights reserved. 3 | Released under Apache 2.0 license as described in the file LICENSE. 4 | Authors: Asta H. From, Jannis Limperg 5 | -/ 6 | import Aesop 7 | 8 | set_option aesop.check.all true 9 | set_option aesop.smallErrorMessages true 10 | 11 | inductive Any (P : α → Prop) : List α → Prop where 12 | | here (x xs) : P x → Any P (x :: xs) 13 | 14 | inductive Perm : (xs ys : List α) → Type where 15 | | refl xs : Perm xs xs 16 | | prep (x xs ys) : Perm xs ys → Perm (x :: xs) (x :: ys) 17 | 18 | theorem Perm.any {xs ys : List α} (perm : Perm xs ys) (P : α → Prop) 19 | : Any P xs → Any P ys := by 20 | induction perm <;> aesop (add safe [constructors Any, cases Any]) 21 | 22 | /-- 23 | error: tactic 'aesop' failed, maximum number of rule applications (100) reached. Set the 'maxRuleApplications' option to increase the limit. 24 | -/ 25 | #guard_msgs in 26 | theorem error (P : Nat → Prop) (Δ : List Nat) : Any P Δ := by 27 | aesop (add 50% [constructors Perm, constructors Any, Perm.any]) 28 | (config := { maxRuleApplications := 100, terminal := true }) 29 | 30 | /-- 31 | error: tactic 'aesop' failed, failed to prove the goal after exhaustive search. 32 | -/ 33 | #guard_msgs in 34 | theorem fine (P : α → Prop) (Δ : List α) : Any P Δ := by 35 | aesop (add unsafe [50% constructors Perm, 50% constructors Any, apply 50% Perm.any]) 36 | (config := { maxRuleApplications := 10, terminal := true }) 37 | -------------------------------------------------------------------------------- /AesopTest/13_2.lean: -------------------------------------------------------------------------------- 1 | /- 2 | Copyright (c) 2022 Asta H. From. All rights reserved. 3 | Released under Apache 2.0 license as described in the file LICENSE. 4 | Authors: Asta H. From, Jannis Limperg 5 | -/ 6 | import Aesop 7 | 8 | set_option aesop.check.all true 9 | set_option aesop.smallErrorMessages true 10 | 11 | inductive Perm : (xs ys : List α) → Type where 12 | | prep {xs} x : Perm (x :: xs) (x :: xs) 13 | 14 | inductive Proof : (Γ Δ : List Φ) → Prop where 15 | | basic (Γ Δ n) : Proof (n :: Γ) (n :: Δ) 16 | | per_l (Γ Γ' Δ) : Proof Γ Δ → Perm Γ' Γ → Proof Γ' Δ 17 | 18 | /-- 19 | error: tactic 'aesop' failed, maximum number of rule applications (50) reached. Set the 'maxRuleApplications' option to increase the limit. 20 | -/ 21 | #guard_msgs in 22 | theorem weaken (Γ Δ : List Φ) (prf : Proof Γ Δ) (δ : Φ) : Proof Γ (δ :: Δ) := by 23 | induction prf 24 | case basic Γ Δ n => 25 | aesop (add unsafe [constructors Proof, constructors Perm]) 26 | (config := { maxRuleApplications := 50, terminal := true }) 27 | case per_l Γ Γ' Δ _ perm ih => 28 | apply Proof.per_l Γ Γ' (δ :: Δ) ih perm 29 | -------------------------------------------------------------------------------- /AesopTest/18.lean: -------------------------------------------------------------------------------- 1 | /- 2 | Copyright (c) 2022 Asta H. From. All rights reserved. 3 | Released under Apache 2.0 license as described in the file LICENSE. 4 | Authors: Asta H. From, Jannis Limperg 5 | -/ 6 | import Aesop 7 | 8 | set_option aesop.check.all true 9 | set_option aesop.smallErrorMessages true 10 | 11 | attribute [aesop safe cases (cases_patterns := [List.Mem _ []])] List.Mem 12 | attribute [aesop unsafe 50% cases (cases_patterns := [List.Mem _ (_ :: _)])] List.Mem 13 | 14 | /-- 15 | error: tactic 'aesop' failed, failed to prove the goal after exhaustive search. 16 | -/ 17 | #guard_msgs in 18 | theorem Mem.split [DecidableEq α] (xs : List α) (v : α) (h : v ∈ xs) 19 | : ∃ l r, xs = l ++ v :: r := by 20 | induction xs 21 | case nil => 22 | aesop 23 | case cons x xs ih => 24 | have dec : Decidable (x = v) := inferInstance 25 | cases dec 26 | case isFalse no => 27 | aesop (config := { terminal := true }) (erase Aesop.BuiltinRules.ext) 28 | case isTrue yes => 29 | apply Exists.intro [] 30 | apply Exists.intro xs 31 | rw [yes] 32 | rfl 33 | -------------------------------------------------------------------------------- /AesopTest/2.lean: -------------------------------------------------------------------------------- 1 | /- 2 | Copyright (c) 2022 Jannis Limperg. All rights reserved. 3 | Released under Apache 2.0 license as described in the file LICENSE. 4 | Authors: Jannis Limperg 5 | -/ 6 | 7 | import Aesop 8 | 9 | -- This option would make Aesop time out. 10 | set_option aesop.check.all false 11 | 12 | inductive Even : Nat → Prop 13 | | zero : Even Nat.zero 14 | | plus_two {n} : Even n → Even (n + 2) 15 | 16 | attribute [aesop safe] Even.zero Even.plus_two 17 | 18 | example : Even 500 := by 19 | aesop (config := { maxRuleApplications := 0, maxRuleApplicationDepth := 0 }) 20 | -------------------------------------------------------------------------------- /AesopTest/20.lean: -------------------------------------------------------------------------------- 1 | /- 2 | Copyright (c) 2022 Jannis Limperg. All rights reserved. 3 | Released under Apache 2.0 license as described in the file LICENSE. 4 | Authors: Asta H. From, Jannis Limperg 5 | -/ 6 | import Aesop 7 | 8 | set_option aesop.check.all true 9 | 10 | attribute [aesop safe cases (cases_patterns := [List.Mem _ []])] List.Mem 11 | attribute [aesop unsafe 50% constructors] List.Mem 12 | attribute [aesop unsafe 50% cases (cases_patterns := [List.Mem _ (_ :: _)])] List.Mem 13 | 14 | @[aesop safe [constructors, cases (cases_patterns := [All _ [], All _ (_ :: _)])]] 15 | inductive All (P : α → Prop) : List α → Prop where 16 | | none : All P [] 17 | | more {x xs} : P x → All P xs → All P (x :: xs) 18 | 19 | @[simp] 20 | theorem All.cons (P : α → Prop) (x : α) (xs : List α) 21 | : All P (x :: xs) ↔ (P x ∧ All P xs) := by 22 | aesop 23 | 24 | theorem mem (P : α → Prop) (xs : List α) 25 | : All P xs ↔ ∀ a : α, a ∈ xs → P a := by 26 | induction xs 27 | case nil => aesop 28 | case cons x xs ih => aesop (config := { useSimpAll := false }) 29 | 30 | theorem mem' (P : α → Prop) (xs : List α) 31 | : All P xs ↔ ∀ a : α, a ∈ xs → P a := by 32 | induction xs <;> aesop 33 | -------------------------------------------------------------------------------- /AesopTest/203.lean: -------------------------------------------------------------------------------- 1 | /- 2 | Copyright (c) 2025 Jannis Limperg. All rights reserved. 3 | Released under Apache 2.0 license as described in the file LICENSE. 4 | Authors: Jannis Limperg 5 | -/ 6 | 7 | import Aesop.Frontend 8 | 9 | -- Thank to Damiano Testa for this bug report. 10 | 11 | /-- 12 | info: @[aesop (rule_sets := [builtin✝]) safe✝ apply✝] 13 | example : True✝ := 14 | trivial✝ 15 | -/ 16 | #guard_msgs in 17 | #eval do 18 | let stx ← `(@[aesop (rule_sets := [builtin]) safe apply] example : True := trivial) 19 | Lean.logInfo stx 20 | -------------------------------------------------------------------------------- /AesopTest/205.lean: -------------------------------------------------------------------------------- 1 | /- 2 | Copyright (c) 2025 Jannis Limperg. All rights reserved. 3 | Released under Apache 2.0 license as described in the file LICENSE. 4 | Authors: Jannis Limperg 5 | -/ 6 | 7 | -- Thanks to Bruno Dutertre for reporting this bug. 8 | 9 | import Aesop 10 | 11 | axiom R {α} : α → α → Prop 12 | 13 | @[aesop safe forward] axiom sym : R x y → R y x 14 | @[aesop safe forward] axiom tran : R x y → R y z → R x z 15 | 16 | /-- 17 | info: Try this: 18 | have fwd : R b y := sym h₃ 19 | have fwd_1 : R a x := sym h₂ 20 | have fwd_2 : R b b := tran (y := y) fwd h₃ 21 | have fwd_3 : R y y := tran (y := b) h₃ fwd 22 | have fwd_4 : R a a := tran (y := x) fwd_1 h₂ 23 | have fwd_5 : R x x := tran (y := a) h₂ fwd_1 24 | sorry 25 | --- 26 | warning: declaration uses 'sorry' 27 | -/ 28 | #guard_msgs in 29 | example (α : Type u_1) (x y a b : α) (h₂ : R x a) (h₃ : R y b) : False := by 30 | aesop? (rule_sets := [-builtin]) (config := { warnOnNonterminal := false }) 31 | sorry 32 | -------------------------------------------------------------------------------- /AesopTest/207.lean: -------------------------------------------------------------------------------- 1 | /- 2 | Copyright (c) 2025 Jannis Limperg. All rights reserved. 3 | Released under Apache 2.0 license as described in the file LICENSE. 4 | Authors: Jannis Limperg 5 | -/ 6 | 7 | -- Thanks to Jireh Loreaux for reporting this bug. 8 | 9 | import Aesop 10 | 11 | def Set α := α → Prop 12 | 13 | instance : Membership α (Set α) := 14 | ⟨λ s a => s a⟩ 15 | 16 | instance : Singleton α (Set α) := 17 | ⟨λ a b => b = a⟩ 18 | 19 | axiom EqOn : (f₁ f₂ : α → α) → Set α → Prop 20 | 21 | @[aesop safe forward] 22 | axiom EqOn.eq_of_mem (h : EqOn f₁ f₂ s) (ha : a ∈ s) : f₁ a = f₂ a 23 | 24 | @[simp] 25 | axiom eqOn_singleton : EqOn f₁ f₂ {x} ↔ f₁ x = f₂ x 26 | 27 | example (s : Set Nat) (x : Nat) (hx : x ∈ s) (f : Nat → Nat) 28 | (h_eqOn_x : EqOn f (λ _ => 1) {x}) (this : EqOn f (λ _ => 0) s) : 29 | False := by 30 | aesop 31 | 32 | example (s : Set Nat) (x : Nat) (hx : x ∈ s) (f : Nat → Nat) 33 | (this : EqOn f (λ _ => 0) s) (h_eqOn_x : EqOn f (λ _ => 1) {x}) : 34 | False := by 35 | aesop 36 | -------------------------------------------------------------------------------- /AesopTest/23.lean: -------------------------------------------------------------------------------- 1 | /- 2 | Copyright (c) 2022 Jannis Limperg. All rights reserved. 3 | Released under Apache 2.0 license as described in the file LICENSE. 4 | Authors: Jannis Limperg 5 | -/ 6 | 7 | import Aesop 8 | 9 | set_option aesop.check.all true 10 | 11 | def Involutive (f : α → α) : Prop := 12 | ∀ x, f (f x) = x 13 | 14 | example : Involutive not := by 15 | aesop (add norm simp Involutive) 16 | 17 | example : Involutive not := by 18 | aesop (add norm unfold Involutive) 19 | -------------------------------------------------------------------------------- /AesopTest/26.lean: -------------------------------------------------------------------------------- 1 | /- 2 | Copyright (c) 2022 Jannis Limperg. All rights reserved. 3 | Released under Apache 2.0 license as described in the file LICENSE. 4 | Authors: Asta H. From, Jannis Limperg 5 | -/ 6 | 7 | import Aesop 8 | 9 | set_option aesop.check.all true 10 | 11 | attribute [-simp] List.all_cons List.all_nil List.all_eq_true 12 | 13 | /-- 14 | warning: aesop: failed to prove the goal after exhaustive search. 15 | --- 16 | error: unsolved goals 17 | case left 18 | α : Type u_1 19 | P : α → Bool 20 | x : α 21 | xs : List α 22 | h : (x :: xs).all P = true 23 | ⊢ P x = true 24 | 25 | case right 26 | α : Type u_1 27 | P : α → Bool 28 | x : α 29 | xs : List α 30 | h : (x :: xs).all P = true 31 | ⊢ xs.all P = true 32 | -/ 33 | #guard_msgs in 34 | theorem all_cons (P : α → Bool) (x : α) (xs : List α) (h : (x :: xs).all P) 35 | : P x ∧ xs.all P := by 36 | aesop 37 | -------------------------------------------------------------------------------- /AesopTest/41.lean: -------------------------------------------------------------------------------- 1 | import Aesop 2 | 3 | set_option aesop.check.all true 4 | 5 | /-- 6 | error: no such rule set: 'Nonexistent' 7 | (Use 'declare_aesop_rule_set' to declare rule sets. 8 | Declared rule sets are not visible in the current file; they only become visible once you import the declaring file.) 9 | -/ 10 | #guard_msgs in 11 | example : True := by 12 | aesop (rule_sets := [Nonexistent]) 13 | -------------------------------------------------------------------------------- /AesopTest/43.lean: -------------------------------------------------------------------------------- 1 | /- 2 | Copyright (c) 2023 Jannis Limperg. All rights reserved. 3 | Released under Apache 2.0 license as described in the file LICENSE. 4 | Authors: Jannis Limperg 5 | -/ 6 | 7 | import Aesop 8 | 9 | set_option aesop.check.all true 10 | 11 | structure A 12 | 13 | open Lean.Elab.Tactic in 14 | @[aesop norm] 15 | def tac : TacticM Unit := do 16 | evalTactic $ ← `(tactic| exact A.mk) 17 | 18 | example : A := by 19 | set_option aesop.check.script false in 20 | set_option aesop.check.script.steps false in 21 | aesop 22 | -------------------------------------------------------------------------------- /AesopTest/AddRulesCommand.lean: -------------------------------------------------------------------------------- 1 | /- 2 | Copyright (c) 2024 Jannis Limperg. All rights reserved. 3 | Released under Apache 2.0 license as described in the file LICENSE. 4 | Authors: Jannis Limperg 5 | -/ 6 | 7 | import Aesop 8 | 9 | set_option aesop.check.all true 10 | set_option aesop.smallErrorMessages true 11 | 12 | -- Basic examples 13 | 14 | structure TT₁ where 15 | 16 | /-- error: tactic 'aesop' failed, made no progress -/ 17 | #guard_msgs in 18 | example : TT₁ := by 19 | aesop 20 | 21 | add_aesop_rules safe TT₁ 22 | 23 | example : TT₁ := by 24 | aesop 25 | 26 | -- Local rules 27 | 28 | structure TT₂ where 29 | 30 | namespace Test 31 | 32 | local add_aesop_rules safe TT₂ 33 | 34 | example : TT₂ := by 35 | aesop 36 | 37 | end Test 38 | 39 | /-- error: tactic 'aesop' failed, made no progress -/ 40 | #guard_msgs in 41 | example : TT₂ := by 42 | aesop 43 | 44 | -- Scoped rules 45 | 46 | structure TT₃ where 47 | 48 | namespace Test 49 | 50 | scoped add_aesop_rules safe TT₃ 51 | 52 | example : TT₃ := by 53 | aesop 54 | 55 | end Test 56 | 57 | /-- error: tactic 'aesop' failed, made no progress -/ 58 | #guard_msgs in 59 | example : TT₃ := by 60 | aesop 61 | 62 | def Test.example : TT₃ := by 63 | aesop 64 | 65 | -- Tactics 66 | 67 | structure TT₄ where 68 | 69 | /-- error: tactic 'aesop' failed, made no progress -/ 70 | #guard_msgs in 71 | example : TT₄ := by 72 | aesop 73 | 74 | add_aesop_rules safe (by exact TT₄.mk) 75 | 76 | example : TT₄ := by 77 | aesop 78 | 79 | -- Multiple rules 80 | 81 | axiom T : Type 82 | axiom U : Type 83 | axiom f : T → U 84 | axiom t : T 85 | 86 | /-- error: tactic 'aesop' failed, made no progress -/ 87 | #guard_msgs in 88 | example : U := by 89 | aesop 90 | 91 | add_aesop_rules safe [(by apply f), t] 92 | 93 | noncomputable example : U := by 94 | aesop 95 | -------------------------------------------------------------------------------- /AesopTest/Aesop.lean: -------------------------------------------------------------------------------- 1 | /- 2 | Copyright (c) 2021 Jannis Limperg. All rights reserved. 3 | Released under Apache 2.0 license as described in the file LICENSE. 4 | Authors: Jannis Limperg 5 | -/ 6 | 7 | import Aesop 8 | 9 | set_option aesop.check.all true 10 | set_option aesop.smallErrorMessages true 11 | 12 | open Lean 13 | open Lean.Meta 14 | open Lean.Elab.Tactic 15 | 16 | section EvenOdd 17 | 18 | inductive Even : Nat → Prop 19 | | zero : Even 0 20 | | plus_two {n} : Even n → Even (n + 2) 21 | 22 | inductive Odd : Nat → Prop 23 | | one : Odd 1 24 | | plus_two {n} : Odd n → Odd (n + 2) 25 | 26 | inductive EvenOrOdd : Nat → Prop 27 | | even {n} : Even n → EvenOrOdd n 28 | | odd {n} : Odd n → EvenOrOdd n 29 | 30 | attribute [aesop unsafe] EvenOrOdd.even EvenOrOdd.odd 31 | attribute [aesop safe] Even.zero Even.plus_two 32 | attribute [aesop 100%] Odd.one Odd.plus_two 33 | 34 | @[aesop norm unfold] 35 | def EvenOrOdd' (n : Nat) : Prop := EvenOrOdd n 36 | 37 | example : EvenOrOdd' 3 := by 38 | aesop 39 | 40 | end EvenOdd 41 | 42 | -- In this example, the goal is solved already during normalisation. 43 | example : 0 = 0 := by aesop 44 | 45 | 46 | -- An intentionally looping Aesop call, to test the limiting options 47 | section Loop 48 | 49 | structure Wrap (α) where 50 | unwrap : α 51 | 52 | /-- 53 | error: tactic 'aesop' failed, maximum number of rule applications (20) reached. Set the 'maxRuleApplications' option to increase the limit. 54 | -/ 55 | #guard_msgs in 56 | example (h : α → α) (h' : Wrap α) : α := by 57 | aesop (add safe h) 58 | (config := { maxRuleApplications := 20, maxGoals := 0, maxRuleApplicationDepth := 0, terminal := true }) 59 | 60 | /-- 61 | error: tactic 'aesop' failed, maximum number of goals (20) reached. Set the 'maxGoals' option to increase the limit. 62 | -/ 63 | #guard_msgs in 64 | example (h : α → α) (h' : Wrap α) : α := by 65 | aesop (add safe h) 66 | (config := { maxGoals := 20, maxRuleApplications := 0, maxRuleApplicationDepth := 0, terminal := true }) 67 | 68 | /-- 69 | error: tactic 'aesop' failed, failed to prove the goal. Some goals were not explored because the maximum rule application depth (20) was reached. Set option 'maxRuleApplicationDepth' to increase the limit. 70 | -/ 71 | #guard_msgs in 72 | example (h : α → α) (h' : Wrap α) : α := by 73 | aesop (add safe h) 74 | (config := { maxRuleApplicationDepth := 20, maxGoals := 0, maxRuleApplications := 0, terminal := true }) 75 | 76 | end Loop 77 | 78 | 79 | -- This example tests the builtin rule that applies local hypotheses. 80 | example (Even : Nat → Prop) (zero : Even 0) 81 | (plusTwo : ∀ n, Even n → Even (n + 2)) : Even 20 := by 82 | aesop 83 | -------------------------------------------------------------------------------- /AesopTest/AllWeaken.lean: -------------------------------------------------------------------------------- 1 | /- 2 | Copyright (c) 2022 Jannis Limperg. All rights reserved. 3 | Released under Apache 2.0 license as described in the file LICENSE. 4 | Authors: Jannis Limperg, Asta H. From 5 | -/ 6 | 7 | import Aesop 8 | 9 | set_option aesop.check.all true 10 | set_option aesop.smallErrorMessages true 11 | 12 | inductive All (P : α → Prop) : List α → Prop where 13 | | none : All P [] 14 | | more {x xs} : P x → All P xs → All P (x :: xs) 15 | 16 | @[aesop unsafe] 17 | axiom weaken {α} (P Q : α → Prop) (wk : ∀ x, P x → Q x) (xs : List α) 18 | (h : All P xs) : All Q xs 19 | 20 | /-- 21 | error: tactic 'aesop' failed, maximum number of rule applications (50) reached. Set the 'maxRuleApplications' option to increase the limit. 22 | -/ 23 | #guard_msgs in 24 | example : All (· ∈ []) (@List.nil α) := by 25 | aesop (config := { maxRuleApplications := 50, terminal := true }) 26 | -------------------------------------------------------------------------------- /AesopTest/ApplyHypsTransparency.lean: -------------------------------------------------------------------------------- 1 | /- 2 | Copyright (c) 2023 Jannis Limperg. All rights reserved. 3 | Released under Apache 2.0 license as described in the file LICENSE. 4 | Authors: Jannis Limperg 5 | -/ 6 | 7 | import Aesop 8 | 9 | set_option aesop.check.all true 10 | set_option aesop.smallErrorMessages true 11 | 12 | def T := Unit → Nat 13 | 14 | /-- 15 | error: tactic 'aesop' failed, failed to prove the goal after exhaustive search. 16 | -/ 17 | #guard_msgs in 18 | example (h : T) : Nat := by 19 | aesop (config := { applyHypsTransparency := .reducible, terminal := true }) 20 | 21 | example (h : T) : Nat := by 22 | aesop 23 | 24 | @[irreducible] def U := Empty 25 | 26 | /-- 27 | error: tactic 'aesop' failed, failed to prove the goal after exhaustive search. 28 | -/ 29 | #guard_msgs in 30 | example (h : Unit → Empty) : U := by 31 | aesop (config := { applyHypsTransparency := .reducible, terminal := true }) 32 | 33 | /-- 34 | error: tactic 'aesop' failed, failed to prove the goal after exhaustive search. 35 | -/ 36 | #guard_msgs in 37 | example (h : Unit → Empty) : U := by 38 | aesop (config := { terminal := true }) 39 | 40 | example (h : Unit → Empty) : U := by 41 | aesop (config := { applyHypsTransparency := .all }) 42 | -------------------------------------------------------------------------------- /AesopTest/ApplyTransparency.lean: -------------------------------------------------------------------------------- 1 | /- 2 | Copyright (c) 2023 Jannis Limperg. All rights reserved. 3 | Released under Apache 2.0 license as described in the file LICENSE. 4 | Authors: Jannis Limperg 5 | -/ 6 | 7 | import Aesop 8 | 9 | set_option aesop.check.all true 10 | set_option aesop.smallErrorMessages true 11 | 12 | def T := True 13 | 14 | /-- 15 | error: tactic 'aesop' failed, failed to prove the goal after exhaustive search. 16 | -/ 17 | #guard_msgs in 18 | example : T := by 19 | aesop (add safe apply True.intro) (config := { terminal := true }) 20 | 21 | /-- 22 | error: tactic 'aesop' failed, failed to prove the goal after exhaustive search. 23 | -/ 24 | #guard_msgs in 25 | example : T := by 26 | aesop (add safe apply (transparency := default) True.intro) 27 | (config := { terminal := true }) 28 | 29 | example : T := by 30 | aesop (add safe apply (transparency! := default) True.intro) 31 | 32 | @[irreducible] def U := T 33 | 34 | /-- 35 | error: tactic 'aesop' failed, failed to prove the goal after exhaustive search. 36 | -/ 37 | #guard_msgs in 38 | example : U := by 39 | aesop (add safe apply True.intro) (config := { terminal := true }) 40 | 41 | /-- 42 | error: tactic 'aesop' failed, failed to prove the goal after exhaustive search. 43 | -/ 44 | #guard_msgs in 45 | example : U := by 46 | aesop (add safe apply (transparency := default) True.intro) 47 | (config := { terminal := true }) 48 | 49 | /-- 50 | error: tactic 'aesop' failed, failed to prove the goal after exhaustive search. 51 | -/ 52 | #guard_msgs in 53 | example : U := by 54 | aesop (add safe apply (transparency! := default) True.intro) 55 | (config := { terminal := true }) 56 | 57 | example : U := by 58 | aesop (add safe apply (transparency! := all) True.intro) 59 | -------------------------------------------------------------------------------- /AesopTest/AssumptionTransparency.lean: -------------------------------------------------------------------------------- 1 | /- 2 | Copyright (c) 2023 Jannis Limperg. All rights reserved. 3 | Released under Apache 2.0 license as described in the file LICENSE. 4 | Authors: Jannis Limperg 5 | -/ 6 | 7 | import Aesop 8 | 9 | set_option aesop.check.all true 10 | set_option aesop.smallErrorMessages true 11 | 12 | def T := Empty 13 | 14 | /-- 15 | error: tactic 'aesop' failed, failed to prove the goal after exhaustive search. 16 | -/ 17 | #guard_msgs in 18 | example (h : T) : Empty := by 19 | aesop (erase Aesop.BuiltinRules.applyHyps) 20 | (config := { assumptionTransparency := .reducible, terminal := true }) 21 | 22 | /-- 23 | error: tactic 'aesop' failed, failed to prove the goal after exhaustive search. 24 | -/ 25 | #guard_msgs in 26 | example (h : T) : Empty := by 27 | aesop (erase Aesop.BuiltinRules.applyHyps) 28 | (config := { assumptionTransparency := .reducible, terminal := true }) 29 | 30 | example (h : T) : Empty := by 31 | aesop (erase Aesop.BuiltinRules.applyHyps) 32 | 33 | @[irreducible] def U := False 34 | 35 | /-- 36 | error: tactic 'aesop' failed, failed to prove the goal after exhaustive search. 37 | -/ 38 | #guard_msgs in 39 | example (h : U) : False := by 40 | aesop (config := { terminal := true }) 41 | 42 | example (h : U) : False := by 43 | aesop (config := { assumptionTransparency := .all }) 44 | -------------------------------------------------------------------------------- /AesopTest/AuxDecl.lean: -------------------------------------------------------------------------------- 1 | /- 2 | Copyright (c) 2025 Jannis Limperg. All rights reserved. 3 | Released under Apache 2.0 license as described in the file LICENSE. 4 | Authors: Jannis Limperg 5 | -/ 6 | 7 | import Aesop 8 | 9 | /- 10 | This test case checks whether tactics that make use of auxiliary declarations, 11 | such as `omega` and `bv_decide`, work with Aesop. Auxiliary declarations are 12 | problematic because we can't allows tactics on different branches of the search 13 | tree to add auxiliary declarations with the same name. 14 | -/ 15 | 16 | theorem foo (m n : Nat) : n + m = m + n ∧ m + n = n + m := by 17 | fail_if_success aesop (config := { terminal := true }) 18 | aesop (add safe (by omega)) 19 | 20 | local instance [Add α] [Add β] : Add (α × β) := 21 | ⟨λ (a, b) (a', b') => (a + a', b + b')⟩ 22 | 23 | theorem bar 24 | (fst snd fst_1 snd_1 fst_2 snd_2 w w_1 w_2 w_3 : BitVec 128) 25 | (left : w_1.uaddOverflow w_3 = true) 26 | (left_1 : w.uaddOverflow w_2 = true) 27 | (right : (w_1 + w_3).uaddOverflow 1#128 = true) : 28 | (fst_2, snd_2) = (fst, snd) + (fst_1, snd_1) := by 29 | aesop (add safe (by bv_decide)) 30 | 31 | theorem baz (a b : BitVec 1) : (a = 0 ∨ a = 1) ∧ (b = 0 ∨ b = 1) := by 32 | aesop (add safe 1000 (by bv_decide)) 33 | -------------------------------------------------------------------------------- /AesopTest/BigStep.lean: -------------------------------------------------------------------------------- 1 | /- 2 | Copyright (c) 2024 Asei Inoue. All rights reserved. 3 | Released under Apache 2.0 license as described in the file LICENSE. 4 | Authors: Asei Inoue, Jannis Limperg 5 | -/ 6 | 7 | import Aesop 8 | 9 | set_option aesop.check.all true 10 | 11 | abbrev Variable := String 12 | 13 | def State := Variable → Nat 14 | 15 | inductive Stmt : Type where 16 | | skip : Stmt 17 | | assign : Variable → (State → Nat) → Stmt 18 | | seq : Stmt → Stmt → Stmt 19 | | ifThenElse : (State → Prop) → Stmt → Stmt → Stmt 20 | | whileDo : (State → Prop) → Stmt → Stmt 21 | 22 | infix:60 ";; " => Stmt.seq 23 | 24 | export Stmt (skip assign seq ifThenElse whileDo) 25 | 26 | set_option quotPrecheck false in 27 | notation s:70 "[" x:70 "↦" n:70 "]" => (fun v ↦ if v = x then n else s v) 28 | 29 | inductive BigStep : Stmt → State → State → Prop where 30 | | protected skip (s : State) : BigStep skip s s 31 | | protected assign (x : Variable) (a : State → Nat) (s : State) : BigStep (assign x a) s (s[x ↦ a s]) 32 | | protected seq {S T : Stmt} {s t u : State} (hS : BigStep S s t) (hT : BigStep T t u) : 33 | BigStep (S;; T) s u 34 | | protected if_true {B : State → Prop} {s t : State} (hcond : B s) (S T : Stmt) (hbody : BigStep S s t) : 35 | BigStep (ifThenElse B S T) s t 36 | | protected if_false {B : State → Prop} {s t : State} (hcond : ¬ B s) (S T : Stmt) (hbody : BigStep T s t) : 37 | BigStep (ifThenElse B S T) s t 38 | | while_true {B S s t u} (hcond : B s) (hbody : BigStep S s t) (hrest : BigStep (whileDo B S) t u) : 39 | BigStep (whileDo B S) s u 40 | | while_false {B S s} (hcond : ¬ B s) : BigStep (whileDo B S) s s 41 | 42 | notation:55 "(" S:55 "," s:55 ")" " ==> " t:55 => BigStep S s t 43 | 44 | add_aesop_rules safe [BigStep.skip, BigStep.assign, BigStep.seq, BigStep.while_false] 45 | add_aesop_rules 50% [apply BigStep.while_true] 46 | add_aesop_rules safe [ 47 | (by apply BigStep.if_true (hcond := by assumption) (hbody := by assumption)), 48 | (by apply BigStep.if_false (hcond := by assumption) (hbody := by assumption)) 49 | ] 50 | 51 | namespace BigStep 52 | 53 | @[aesop safe destruct] 54 | theorem cases_if_of_true {B S T s t} (hcond : B s) : (ifThenElse B S T, s) ==> t → (S, s) ==> t := by 55 | intro h; cases h <;> aesop 56 | 57 | @[aesop safe destruct] 58 | theorem cases_if_of_false {B S T s t} (hcond : ¬ B s) : (ifThenElse B S T, s) ==> t → (T, s) ==> t := by 59 | intro h; cases h <;> aesop 60 | 61 | @[aesop 30%] 62 | theorem and_excluded {P Q R : Prop} (hQ : P → Q) (hR : ¬ P → R) : (P ∧ Q ∨ ¬ P ∧ R) := by 63 | by_cases h : P <;> aesop 64 | 65 | theorem if_iff {B S T s t} : (ifThenElse B S T, s) ==> t ↔ 66 | (B s ∧ (S, s) ==> t) ∨ (¬ B s ∧ (T, s) ==> t) := by 67 | aesop 68 | 69 | end BigStep 70 | -------------------------------------------------------------------------------- /AesopTest/Cases.lean: -------------------------------------------------------------------------------- 1 | /- 2 | Copyright (c) 2022 Jannis Limperg. All rights reserved. 3 | Released under Apache 2.0 license as described in the file LICENSE. 4 | Authors: Jannis Limperg 5 | -/ 6 | import Aesop 7 | 8 | set_option aesop.check.all true 9 | 10 | @[aesop 50% cases] 11 | inductive FancyAnd (α β : Prop) : Prop 12 | | dummy (p : Empty) 13 | | and (a : α) (b : β) 14 | 15 | attribute [aesop safe -51 cases] Empty 16 | 17 | example {α β} (h : FancyAnd α β) : α ∧ β := by 18 | aesop 19 | 20 | @[aesop safe cases (cases_patterns := [All _ [], All _ (_ :: _)])] 21 | inductive All (P : α → Prop) : List α → Prop 22 | | nil : All P [] 23 | | cons : P x → All P xs → All P (x :: xs) 24 | 25 | @[aesop 99% constructors] 26 | structure MyTrue : Prop 27 | 28 | -- Without the patterns on the `cases` rule for `All`, this test would loop 29 | -- since the `constructors` rule would never be applied. 30 | example {P : α → Prop} (h : All P (x :: xs)) : MyTrue := by 31 | aesop 32 | -------------------------------------------------------------------------------- /AesopTest/CasesScript.lean: -------------------------------------------------------------------------------- 1 | /- 2 | Copyright (c) 2024 Jannis Limperg. All rights reserved. 3 | Released under Apache 2.0 license as described in the file LICENSE. 4 | Authors: Jannis Limperg 5 | -/ 6 | 7 | import Aesop 8 | 9 | set_option aesop.check.all true 10 | 11 | @[aesop 50% cases] 12 | inductive FancyAnd (α β : Prop) : Prop 13 | | dummy (p : Empty) 14 | | and (a : α) (b : β) 15 | 16 | /-- 17 | info: Try this: 18 | apply And.intro 19 | · 20 | cases h with 21 | | dummy p => 22 | have fwd : False := Aesop.BuiltinRules.empty_false p 23 | simp_all only 24 | | and a b => simp_all only 25 | · 26 | cases h with 27 | | dummy p => 28 | have fwd : False := Aesop.BuiltinRules.empty_false p 29 | simp_all only 30 | | and a b => simp_all only 31 | -/ 32 | #guard_msgs in 33 | example {α β} (h : FancyAnd α β) : α ∧ β := by 34 | aesop? 35 | 36 | @[aesop safe cases (cases_patterns := [All _ [], All _ (_ :: _)])] 37 | inductive All (P : α → Prop) : List α → Prop 38 | | nil : All P [] 39 | | cons : P x → All P xs → All P (x :: xs) 40 | 41 | @[aesop 99% constructors] 42 | structure MyTrue : Prop 43 | 44 | /-- 45 | info: Try this: 46 | rcases h with ⟨⟩ | @⟨x_1, xs_1, a, a_1⟩ 47 | apply MyTrue.mk 48 | -/ 49 | #guard_msgs in 50 | example {P : α → Prop} (h : All P (x :: xs)) : MyTrue := by 51 | aesop? 52 | -------------------------------------------------------------------------------- /AesopTest/CasesTransparency.lean: -------------------------------------------------------------------------------- 1 | /- 2 | Copyright (c) 2023 Jannis Limperg. All rights reserved. 3 | Released under Apache 2.0 license as described in the file LICENSE. 4 | Authors: Jannis Limperg 5 | -/ 6 | 7 | import Aesop 8 | 9 | set_option aesop.check.all true 10 | set_option aesop.smallErrorMessages true 11 | 12 | example (h : False) : α := by 13 | aesop 14 | 15 | def T := False 16 | 17 | variable {α : Type} 18 | 19 | /-- 20 | error: tactic 'aesop' failed, failed to prove the goal after exhaustive search. 21 | -/ 22 | #guard_msgs in 23 | example (h : T) : α := by 24 | aesop (config := { terminal := true }) 25 | 26 | /-- 27 | error: tactic 'aesop' failed, failed to prove the goal after exhaustive search. 28 | -/ 29 | #guard_msgs in 30 | example (h : T) : α := by 31 | aesop (add safe cases (transparency! := reducible) False) 32 | (config := { terminal := true }) 33 | 34 | /-- 35 | error: tactic 'aesop' failed, failed to prove the goal after exhaustive search. 36 | -/ 37 | #guard_msgs in 38 | example (h : T) : α := by 39 | aesop (add safe cases (transparency := default) False) 40 | (config := { terminal := true }) 41 | 42 | example (h : T) : α := by 43 | aesop (add safe cases (transparency! := default) False) 44 | 45 | def U := T 46 | 47 | example (h : U) : α := by 48 | aesop (add safe cases (transparency! := default) False) 49 | -------------------------------------------------------------------------------- /AesopTest/Com.lean: -------------------------------------------------------------------------------- 1 | /- 2 | Copyright (c) 2022 Jannis Limperg. All rights reserved. 3 | Released under Apache 2.0 license as described in the file LICENSE. 4 | Authors: Jannis Limperg 5 | -/ 6 | import Aesop 7 | 8 | set_option aesop.check.all false 9 | -- With this option, the test becomes unbearably slow. 10 | 11 | abbrev State := String → Int 12 | 13 | inductive Com where 14 | | Skip : Com 15 | | Seq : Com → Com → Com 16 | 17 | declare_syntax_cat com 18 | 19 | syntax "SKIP" : com 20 | syntax com ";" com : com 21 | syntax "(" com ")" : com 22 | syntax term : com 23 | 24 | syntax "[Com|" com "]" : term 25 | 26 | macro_rules 27 | | `([Com| SKIP]) => `(Com.Skip) 28 | | `([Com| $x ; $y]) => `(Com.Seq [Com| $x] [Com| $y]) 29 | | `([Com| ( $x:com )]) => `([Com| $x]) 30 | | `([Com| $x:term ]) => `($x) 31 | 32 | @[aesop safe [constructors, -100 cases (index := [hyp BigStep Com.Skip _ _, hyp BigStep [Com| _;_] _ _])]] 33 | inductive BigStep : Com → State → State → Prop where 34 | | Skip : BigStep Com.Skip s s 35 | | Seq (h1 : BigStep c₁ s t) (h2 : BigStep c₂ t u) : BigStep [Com| c₁;c₂] s u 36 | 37 | theorem seq_assoc : 38 | BigStep [Com| (c1;c2);c3] s s' ↔ BigStep [Com| c1;c2;c3] s s' := by 39 | aesop 40 | -------------------------------------------------------------------------------- /AesopTest/CompositeLocalRuleTerm.lean: -------------------------------------------------------------------------------- 1 | /- 2 | Copyright (c) 2024 Jannis Limperg. All rights reserved. 3 | Released under Apache 2.0 license as described in the file LICENSE. 4 | Authors: Jannis Limperg 5 | -/ 6 | 7 | import Aesop 8 | 9 | set_option aesop.check.all true 10 | set_option aesop.smallErrorMessages true 11 | 12 | -- Composite terms are supported by the `apply` and `forward` builders... 13 | 14 | structure A where 15 | 16 | /-- 17 | error: tactic 'aesop' failed, made no progress 18 | -/ 19 | #guard_msgs in 20 | example (h : A → β → γ) (b : β) : γ := by 21 | aesop 22 | 23 | example (h : A → β → γ) (b : β) : γ := by 24 | aesop (add safe apply (h {})) 25 | 26 | example (h : A → β → γ) (b : β) : γ := by 27 | aesop (add safe forward (h {})) 28 | 29 | -- ... and also by the `simp` builder. 30 | 31 | /-- 32 | error: tactic 'aesop' failed, made no progress 33 | -/ 34 | #guard_msgs in 35 | example {P : α → Prop} (h : A → x = y) (p : P x) : P y := by 36 | aesop 37 | 38 | example {P : α → Prop} (h : A → x = y) (p : P x) : P y := by 39 | aesop (add simp (h {})) 40 | 41 | /-- 42 | error: tactic 'aesop' failed, made no progress 43 | -/ 44 | #guard_msgs in 45 | example {P : α → Prop} (h₁ : A → x = y) (h₂ : A → y = z) (p : P x) : P z := by 46 | aesop 47 | 48 | example {P : α → Prop} (h₁ : A → x = y) (h₂ : A → y = z) (p : P x) : P z := by 49 | aesop (add simp [(h₁ {}), (h₂ {})]) 50 | -------------------------------------------------------------------------------- /AesopTest/ConstructorEquations.lean: -------------------------------------------------------------------------------- 1 | /- 2 | Copyright (c) 2023 Jannis Limperg. All rights reserved. 3 | Released under Apache 2.0 license as described in the file LICENSE. 4 | Authors: Jannis Limperg 5 | -/ 6 | 7 | import Aesop 8 | 9 | set_option aesop.check.all true 10 | 11 | -- From an equation where both sides contain only constructor applications 12 | -- and variables, Aesop should derive equations about the variables. 13 | 14 | example (h : Nat.zero = Nat.succ n) : False := by 15 | aesop 16 | 17 | example (h : Nat.succ (Nat.succ n) = Nat.succ Nat.zero) : False := by 18 | aesop 19 | 20 | example (h : (Nat.succ m, Nat.succ n) = (Nat.succ a, Nat.succ b)) : 21 | m = a ∧ n = b := by 22 | aesop 23 | 24 | structure MyProd (A B : Type _) where 25 | toProd : A × B 26 | 27 | example (h : MyProd.mk (x, y) = .mk (a, b)) : x = a ∧ y = b := by 28 | aesop 29 | -------------------------------------------------------------------------------- /AesopTest/Constructors.lean: -------------------------------------------------------------------------------- 1 | /- 2 | Copyright (c) 2021 Jannis Limperg. All rights reserved. 3 | Released under Apache 2.0 license as described in the file LICENSE. 4 | Authors: Jannis Limperg 5 | -/ 6 | 7 | import Aesop 8 | 9 | set_option aesop.check.all true 10 | set_option aesop.smallErrorMessages true 11 | 12 | @[aesop safe] 13 | inductive Even : Nat → Type 14 | | zero : Even 0 15 | | plusTwo : Even n → Even (n + 2) 16 | 17 | example : Even 6 := by 18 | aesop 19 | 20 | attribute [-aesop] Even 21 | 22 | def T n := Even n 23 | 24 | /-- 25 | error: tactic 'aesop' failed, failed to prove the goal after exhaustive search. 26 | -/ 27 | #guard_msgs in 28 | example : T 6 := by 29 | aesop (config := { terminal := true }) 30 | 31 | example : T 6 := by 32 | aesop (add safe constructors (transparency! := default) Even) 33 | -------------------------------------------------------------------------------- /AesopTest/CustomIndexing.lean: -------------------------------------------------------------------------------- 1 | /- 2 | Copyright (c) 2022 Jannis Limperg. All rights reserved. 3 | Released under Apache 2.0 license as described in the file LICENSE. 4 | Authors: Jannis Limperg 5 | -/ 6 | 7 | import Aesop 8 | 9 | set_option aesop.check.all true 10 | set_option aesop.smallErrorMessages true 11 | 12 | variable {α : Type} 13 | 14 | /-- 15 | error: tactic 'aesop' failed, failed to prove the goal after exhaustive search. 16 | -/ 17 | #guard_msgs in 18 | example (h : α) : α := by 19 | aesop (rule_sets := [-builtin,-default]) 20 | (add safe h apply (index := [target False])) 21 | (config := { terminal := true }) 22 | 23 | example (h : α) : α := by 24 | aesop (rule_sets := [-builtin,-default]) (add safe h apply) 25 | 26 | -- See Com test for a more realistic scenario. 27 | -------------------------------------------------------------------------------- /AesopTest/CustomTactic.lean: -------------------------------------------------------------------------------- 1 | /- 2 | Copyright (c) 2022 Jannis Limperg. All rights reserved. 3 | Released under Apache 2.0 license as described in the file LICENSE. 4 | Authors: Jannis Limperg 5 | -/ 6 | 7 | import Aesop 8 | 9 | set_option aesop.check.all true 10 | set_option aesop.smallErrorMessages true 11 | 12 | def Foo := True 13 | 14 | /-- 15 | error: tactic 'aesop' failed, failed to prove the goal after exhaustive search. 16 | -/ 17 | #guard_msgs in 18 | example : Foo := by 19 | aesop (config := { terminal := true }) 20 | 21 | example : Foo := by 22 | simp [Foo] 23 | 24 | open Lean.Elab.Tactic in 25 | @[aesop safe] 26 | def myTactic : TacticM Unit := do 27 | evalTactic $ ← `(tactic| rw [Foo]) 28 | 29 | example : Foo := by 30 | set_option aesop.check.script false in 31 | set_option aesop.check.script.steps false in 32 | aesop 33 | -------------------------------------------------------------------------------- /AesopTest/DefaultRuleSets.lean: -------------------------------------------------------------------------------- 1 | /- 2 | Copyright (c) 2023 Jannis Limperg. All rights reserved. 3 | Released under Apache 2.0 license as described in the file LICENSE. 4 | Authors: Jannis Limperg 5 | -/ 6 | 7 | import AesopTest.DefaultRuleSetsInit 8 | 9 | set_option aesop.check.all true 10 | set_option aesop.smallErrorMessages true 11 | 12 | @[aesop norm unfold (rule_sets := [regular₁])] 13 | def T := True 14 | 15 | /-- 16 | error: tactic 'aesop' failed, failed to prove the goal after exhaustive search. 17 | -/ 18 | #guard_msgs in 19 | example : T := by 20 | aesop (config := { terminal := true }) 21 | 22 | example : T := by 23 | aesop (rule_sets := [regular₁]) 24 | 25 | @[aesop norm unfold (rule_sets := [regular₂, dflt₁])] 26 | def U := True 27 | 28 | /-- 29 | error: tactic 'aesop' failed, failed to prove the goal after exhaustive search. 30 | -/ 31 | #guard_msgs in 32 | example : U := by 33 | aesop (rule_sets := [-dflt₁]) (config := { terminal := true }) 34 | 35 | example : U := by 36 | aesop 37 | -------------------------------------------------------------------------------- /AesopTest/DefaultRuleSetsInit.lean: -------------------------------------------------------------------------------- 1 | /- 2 | Copyright (c) 2023 Jannis Limperg. All rights reserved. 3 | Released under Apache 2.0 license as described in the file LICENSE. 4 | Authors: Jannis Limperg 5 | -/ 6 | 7 | import Aesop 8 | 9 | set_option aesop.check.all true 10 | 11 | declare_aesop_rule_sets [regular₁, regular₂] 12 | 13 | declare_aesop_rule_sets [dflt₁, dflt₂] (default := true) 14 | -------------------------------------------------------------------------------- /AesopTest/DestructProducts.lean: -------------------------------------------------------------------------------- 1 | /- 2 | Copyright (c) 2022 Jannis Limperg. All rights reserved. 3 | Released under Apache 2.0 license as described in the file LICENSE. 4 | Authors: Jannis Limperg 5 | -/ 6 | 7 | -- This file tests the builtin rules that destruct hypotheses with product-like 8 | -- types. 9 | 10 | import Aesop 11 | 12 | set_option aesop.check.all true 13 | 14 | @[aesop safe constructors] 15 | inductive Ex (α : Sort u) (β : α → Prop) : Prop 16 | | intro (fst : α) (snd : β fst) 17 | 18 | @[aesop safe constructors] 19 | structure Sig (α : Sort u) (β : α → Sort v) : Sort _ where 20 | fst : α 21 | snd : β fst 22 | 23 | example (h : α ∧ β) : Sig α (λ _ => β) := by 24 | aesop 25 | 26 | example (h : α × β) : Sig α (λ _ => β) := by 27 | aesop 28 | 29 | example (h : PProd α β) : Sig α (λ _ => β) := by 30 | aesop 31 | 32 | example (h : MProd α β) : Sig α (λ _ => β) := by 33 | aesop 34 | 35 | example {p : α → Prop} (h : ∃ a, p a) : Ex α p := by 36 | aesop 37 | 38 | example {p : α → Prop} (h : { a // p a }) : Sig α p := by 39 | aesop 40 | 41 | example {p : α → Type} (h : Σ a, p a) : Sig α p := by 42 | aesop 43 | 44 | example {p : α → Type} (h : Σ' a, p a) : Sig α p := by 45 | aesop 46 | 47 | -- Test case for a bug reported by Jakob von Raumer. 48 | example (x y : α × α) (h: p ∧ (x = y)) : x = y := by 49 | aesop (config := { enableSimp := false }) 50 | -------------------------------------------------------------------------------- /AesopTest/DestructProductsTransparency.lean: -------------------------------------------------------------------------------- 1 | /- 2 | Copyright (c) 2023 Jannis Limperg. All rights reserved. 3 | Released under Apache 2.0 license as described in the file LICENSE. 4 | Authors: Jannis Limperg 5 | -/ 6 | 7 | import Aesop 8 | 9 | set_option aesop.check.all true 10 | set_option aesop.smallErrorMessages true 11 | 12 | @[aesop safe constructors] 13 | structure Sig (α : Sort u) (β : α → Sort v) : Sort _ where 14 | fst : α 15 | snd : β fst 16 | 17 | def T α β := α ∧ β 18 | 19 | /-- 20 | error: tactic 'aesop' failed, failed to prove the goal after exhaustive search. 21 | -/ 22 | #guard_msgs in 23 | example (h : T α β) : Sig α (λ _ => β) := by 24 | aesop (config := { terminal := true }) 25 | 26 | example (h : T α β) : Sig α (λ _ => β) := by 27 | aesop (config := { destructProductsTransparency := .default }) 28 | 29 | def U := T 30 | 31 | /-- 32 | error: tactic 'aesop' failed, failed to prove the goal after exhaustive search. 33 | -/ 34 | #guard_msgs in 35 | example (h : U α β) : Sig α (λ _ => β) := by 36 | aesop (config := { terminal := true }) 37 | 38 | example (h : U α β) : Sig α (λ _ => β) := by 39 | aesop (config := { destructProductsTransparency := .default }) 40 | 41 | /-- 42 | error: tactic 'aesop' failed, failed to prove the goal after exhaustive search. 43 | -/ 44 | #guard_msgs in 45 | example (h : U α β ∧ U γ δ) : Sig α (λ _ => γ) := by 46 | aesop (config := { terminal := true }) 47 | 48 | example (h : U α β ∧ U γ δ) : Sig α (λ _ => γ) := by 49 | aesop (config := { destructProductsTransparency := .default }) 50 | -------------------------------------------------------------------------------- /AesopTest/DocLists.lean: -------------------------------------------------------------------------------- 1 | /- 2 | Copyright (c) 2022 Jannis Limperg. All rights reserved. 3 | Released under Apache 2.0 license as described in the file LICENSE. 4 | Authors: Jannis Limperg 5 | -/ 6 | 7 | -- NOTE: This file contains examples for, and therefore should be kept in sync 8 | -- with, the README. 9 | 10 | import Aesop 11 | 12 | set_option aesop.check.all true 13 | 14 | inductive MyList (α : Type _) 15 | | nil 16 | | cons (hd : α) (tl : MyList α) 17 | 18 | namespace MyList 19 | 20 | protected def append : (_ _ : MyList α) → MyList α 21 | | nil, ys => ys 22 | | cons x xs, ys => cons x (MyList.append xs ys) 23 | 24 | instance : Append (MyList α) := 25 | ⟨MyList.append⟩ 26 | 27 | @[simp] 28 | theorem nil_append : nil ++ xs = xs := rfl 29 | 30 | @[simp] 31 | theorem cons_append : cons x xs ++ ys = cons x (xs ++ ys) := rfl 32 | 33 | @[aesop safe [constructors, cases]] 34 | inductive NonEmpty : MyList α → Prop 35 | | cons : NonEmpty (cons x xs) 36 | 37 | @[aesop 50%] 38 | theorem nonEmpty_append₁ {xs : MyList α} ys : 39 | NonEmpty xs → NonEmpty (xs ++ ys) := by 40 | aesop 41 | 42 | /-- 43 | info: Try this: 44 | intro a 45 | obtain @⟨x, xs_1⟩ := a 46 | simp_all only [cons_append] 47 | apply MyList.NonEmpty.cons 48 | -/ 49 | #guard_msgs in 50 | theorem nonEmpty_append₁' {xs : MyList α} ys : 51 | NonEmpty xs → NonEmpty (xs ++ ys) := by 52 | aesop? 53 | 54 | example {α : Type _} {xs : MyList α} ys zs : 55 | NonEmpty xs → NonEmpty (xs ++ ys ++ zs) := by 56 | aesop 57 | 58 | theorem nil_not_nonEmpty (xs : MyList α) : xs = nil → ¬ NonEmpty xs := by 59 | aesop (add unsafe 10% cases MyList) 60 | 61 | @[simp] 62 | theorem append_nil {xs : MyList α} : 63 | xs ++ nil = xs := by 64 | induction xs <;> aesop 65 | 66 | theorem append_assoc {xs ys zs : MyList α} : 67 | (xs ++ ys) ++ zs = xs ++ (ys ++ zs) := by 68 | induction xs <;> aesop 69 | 70 | end MyList 71 | -------------------------------------------------------------------------------- /AesopTest/DroppedMVars.lean: -------------------------------------------------------------------------------- 1 | /- 2 | Copyright (c) 2022 Jannis Limperg. All rights reserved. 3 | Released under Apache 2.0 license as described in the file LICENSE. 4 | Authors: Jesse Vogel, Jannis Limperg 5 | -/ 6 | 7 | -- Thanks to Jesse Vogel for this test case. It demonstrates the handling of 8 | -- 'dropped' mvars. A rapp drops an mvar if the mvar appears in the parent goal 9 | -- of the rapp, but not in any of its subgoals. In this case, we add an 10 | -- additional regular goal for the mvar. 11 | 12 | import Aesop 13 | 14 | set_option aesop.check.all true 15 | 16 | axiom Ring : Type 17 | 18 | axiom RingHom (R S : Ring) : Type 19 | 20 | @[aesop 99%] 21 | axiom RingId (R : Ring) : RingHom R R 22 | 23 | @[aesop 99%] 24 | axiom ZZ : Ring 25 | 26 | example : ∃ (R : Ring) (_ : RingHom R R), True := by 27 | aesop (add safe True.intro) (config := { enableSimp := false }) 28 | -------------------------------------------------------------------------------- /AesopTest/ElabConfig.lean: -------------------------------------------------------------------------------- 1 | /- 2 | Copyright (c) 2022 Jannis Limperg. All rights reserved. 3 | Released under Apache 2.0 license as described in the file LICENSE. 4 | Authors: Jannis Limperg 5 | -/ 6 | 7 | import Aesop 8 | 9 | set_option aesop.check.all true 10 | 11 | /-- 12 | error: 'noSuchOption' is not a field of structure 'Aesop.Options' 13 | -/ 14 | #guard_msgs in 15 | example : True := by 16 | aesop (config := { noSuchOption := true }) 17 | 18 | /-- 19 | error: 'noSuchOption' is not a field of structure 'Lean.Meta.Simp.ConfigCtx' 20 | -/ 21 | #guard_msgs in 22 | example : True := by 23 | aesop (simp_config := { noSuchOption := true }) 24 | -------------------------------------------------------------------------------- /AesopTest/EnableUnfold.lean: -------------------------------------------------------------------------------- 1 | /- 2 | Copyright (c) 2023 Jannis Limperg. All rights reserved. 3 | Released under Apache 2.0 license as described in the file LICENSE. 4 | Authors: Jannis Limperg 5 | -/ 6 | 7 | import Aesop 8 | 9 | set_option aesop.check.all true 10 | set_option aesop.smallErrorMessages true 11 | 12 | @[aesop norm unfold] 13 | def T := True 14 | 15 | /-- 16 | error: tactic 'aesop' failed, failed to prove the goal after exhaustive search. 17 | -/ 18 | #guard_msgs in 19 | example : T := by 20 | aesop (config := { enableUnfold := false, terminal := true }) 21 | 22 | example : T := by 23 | aesop 24 | -------------------------------------------------------------------------------- /AesopTest/EqualUpToIds.lean: -------------------------------------------------------------------------------- 1 | /- 2 | Copyright (c) 2023 Jannis Limperg. All rights reserved. 3 | Released under Apache 2.0 license as described in the file LICENSE. 4 | Authors: Jannis Limperg 5 | -/ 6 | 7 | import Aesop.Util.Basic 8 | import Aesop.Util.EqualUpToIds 9 | import Aesop.Tree.RunMetaM 10 | 11 | -- Some simple test cases for the EqualUpToIds module. The module is mostly 12 | -- tested by using it in script validation, which is run on almost all Aesop 13 | -- calls in the test suite. 14 | 15 | open Aesop Lean Lean.Elab.Tactic 16 | 17 | def assertEqualTactics (t₁ t₂ : TacticM Unit) : TacticM Unit := do 18 | let commonMCtx ← getMCtx 19 | let preState ← show MetaM _ from saveState 20 | let preGoals ← getGoals 21 | let (state₁, goals₁) ← runTacticMCapturingPostState t₁ preState preGoals 22 | let (state₂, goals₂) ← runTacticMCapturingPostState t₂ preState preGoals 23 | let eq ← 24 | tacticStatesEqualUpToIds commonMCtx state₁.meta.mctx state₂.meta.mctx 25 | goals₁.toArray goals₂.toArray 26 | if ! eq then 27 | throwError "Tactics produced different tactic states.\nTactic 1:{indentD $ ← ppTacticState state₁ goals₁}\nTactic 2:{indentD $ ← ppTacticState state₂ goals₂}\n" 28 | where 29 | ppTacticState (state : Meta.SavedState) (goals : List MVarId) : 30 | MetaM MessageData := 31 | state.runMetaM' do 32 | addMessageContext $ .joinSep (goals.map toMessageData) "\n" 33 | 34 | open Lean.Elab.Tactic in 35 | elab &"assert_equal_tactics " 36 | " { " ts₁:tacticSeq " } " " { " ts₂:tacticSeq " } " : tactic => do 37 | assertEqualTactics (evalTactic ts₁) (evalTactic ts₂) 38 | 39 | example : True := by 40 | assert_equal_tactics { trivial } { trivial } 41 | trivial 42 | 43 | example : True := by 44 | assert_equal_tactics { open Classical in trivial } { trivial } 45 | trivial 46 | 47 | /-- 48 | error: Tactics produced different tactic states. 49 | Tactic 1: 50 | case zero 51 | m : Nat 52 | ⊢ True 53 | case succ 54 | m n✝ : Nat 55 | ⊢ True 56 | Tactic 2: 57 | case zero 58 | n : Nat 59 | ⊢ True 60 | case succ 61 | n n✝ : Nat 62 | ⊢ True 63 | -/ 64 | #guard_msgs in 65 | example (n m : Nat) : True := by 66 | assert_equal_tactics { cases n } { cases m } 67 | trivial 68 | 69 | example : 0 < 3 := by 70 | apply Nat.lt_trans 71 | assert_equal_tactics { exact Nat.lt_succ_self 0 } { exact Nat.lt_succ_self 0 } 72 | (case m => exact 1); all_goals decide 73 | 74 | example : 0 < 3 := by 75 | apply Nat.lt_trans 76 | assert_equal_tactics { apply Nat.lt_trans } { apply Nat.lt_trans } 77 | (case m => exact 1); all_goals decide 78 | -------------------------------------------------------------------------------- /AesopTest/Erase.lean: -------------------------------------------------------------------------------- 1 | /- 2 | Copyright (c) 2022 Jannis Limperg. All rights reserved. 3 | Released under Apache 2.0 license as described in the file LICENSE. 4 | Authors: Jannis Limperg 5 | -/ 6 | 7 | import Aesop 8 | 9 | set_option aesop.check.all true 10 | set_option aesop.smallErrorMessages true 11 | 12 | @[aesop [10% cases, safe constructors]] 13 | inductive Even : Nat → Prop 14 | | zero : Even 0 15 | | plus_two : Even n → Even (n + 2) 16 | 17 | example : Even 2 := by 18 | aesop 19 | 20 | -- Removing the Aesop attribute erases all rules associated with the identifier 21 | -- from all rule sets. 22 | attribute [-aesop] Even 23 | 24 | /-- 25 | error: tactic 'aesop' failed, failed to prove the goal after exhaustive search. 26 | -/ 27 | #guard_msgs in 28 | example : Even 2 := by 29 | aesop (config := { terminal := true }) 30 | 31 | example : Even 2 := by 32 | aesop (add safe Even) 33 | 34 | -- We can also selectively remove rules in a certain phase or with a certain 35 | -- builder. 36 | attribute [aesop [unsafe 10% cases, safe constructors]] Even 37 | 38 | erase_aesop_rules [ unsafe Even ] 39 | 40 | example : Even 2 := by 41 | aesop 42 | 43 | erase_aesop_rules [ constructors Even ] 44 | 45 | /-- 46 | error: tactic 'aesop' failed, failed to prove the goal after exhaustive search. 47 | -/ 48 | #guard_msgs in 49 | example : Even 2 := by 50 | aesop (config := { terminal := true }) 51 | 52 | example : Even 2 := by 53 | aesop (add safe constructors Even) 54 | -------------------------------------------------------------------------------- /AesopTest/EraseSimp.lean: -------------------------------------------------------------------------------- 1 | /- 2 | Copyright (c) 2023 Jannis Limperg. All rights reserved. 3 | Released under Apache 2.0 license as described in the file LICENSE. 4 | Authors: Jannis Limperg 5 | -/ 6 | 7 | import Aesop 8 | 9 | set_option aesop.check.all true 10 | set_option aesop.smallErrorMessages true 11 | 12 | example (n : Nat) : n + m = m + n := by 13 | aesop (add simp Nat.add_comm) 14 | 15 | attribute [local simp] Nat.add_comm 16 | 17 | example (n : Nat) : n + m = m + n := by 18 | aesop 19 | 20 | /-- 21 | error: tactic 'aesop' failed, made no progress 22 | -/ 23 | #guard_msgs in 24 | example (n : Nat) : n + m = m + n := by 25 | aesop (erase Nat.add_comm) (config := { warnOnNonterminal := false }) 26 | 27 | /-- 28 | error: tactic 'aesop' failed, made no progress 29 | -/ 30 | #guard_msgs in 31 | example (n : Nat) : n + m = m + n := by 32 | aesop (erase norm simp Nat.add_comm) (config := { warnOnNonterminal := false }) 33 | 34 | /-- 35 | error: aesop: 'Nat.add_comm' is not registered (with the given features) in any rule set. 36 | -/ 37 | #guard_msgs in 38 | example (n : Nat) : n + m = m + n := by 39 | aesop (erase apply Nat.add_comm) 40 | -------------------------------------------------------------------------------- /AesopTest/EraseUnfold.lean: -------------------------------------------------------------------------------- 1 | /- 2 | Copyright (c) 2023 Jannis Limperg. All rights reserved. 3 | Released under Apache 2.0 license as described in the file LICENSE. 4 | Authors: Jireh Loreaux, Jannis Limperg 5 | -/ 6 | 7 | -- Thanks to Jireh Loreaux for reporting this MWE. 8 | 9 | import Aesop 10 | 11 | set_option aesop.check.all true 12 | set_option aesop.smallErrorMessages true 13 | 14 | @[irreducible] 15 | def foo : Nat := 37 16 | 17 | /-- 18 | error: tactic 'aesop' failed, failed to prove the goal after exhaustive search. 19 | -/ 20 | #guard_msgs in 21 | example : foo = 37 := by 22 | aesop (config := { terminal := true }) (erase Aesop.BuiltinRules.rfl) 23 | 24 | example : foo = 37 := by 25 | unfold foo 26 | rfl 27 | 28 | attribute [aesop norm unfold] foo 29 | 30 | example : foo = 37 := by aesop (erase Aesop.BuiltinRules.rfl) 31 | 32 | attribute [-aesop] foo 33 | 34 | /-- 35 | error: tactic 'aesop' failed, failed to prove the goal after exhaustive search. 36 | -/ 37 | #guard_msgs in 38 | example : foo = 37 := by 39 | aesop (config := { terminal := true }) (erase Aesop.BuiltinRules.rfl) 40 | 41 | example : foo = 37 := by 42 | unfold foo 43 | rfl 44 | -------------------------------------------------------------------------------- /AesopTest/Ext.lean: -------------------------------------------------------------------------------- 1 | /- 2 | Copyright (c) 2023 Jannis Limperg. All rights reserved. 3 | Released under Apache 2.0 license as described in the file LICENSE. 4 | Authors: Jannis Limperg 5 | -/ 6 | 7 | import Aesop 8 | 9 | example (f g : α → β) (h : ∀ a, f a = g a) : f = g := by 10 | aesop 11 | 12 | example (x y : α × β) (h₁ : x.1 = y.1) (h₂ : x.2 = y.2) : x = y := by 13 | aesop 14 | -------------------------------------------------------------------------------- /AesopTest/ExtScript.lean: -------------------------------------------------------------------------------- 1 | /- 2 | Copyright (c) 2024 Jannis Limperg. All rights reserved. 3 | Released under Apache 2.0 license as described in the file LICENSE. 4 | Authors: Jannis Limperg 5 | -/ 6 | 7 | import Aesop 8 | 9 | set_option aesop.check.all true 10 | 11 | @[ext] 12 | structure MyProd (α β : Type _) where 13 | fst : α 14 | snd : β 15 | 16 | variable {α β γ δ ι: Type} 17 | 18 | /-- 19 | info: Try this: 20 | ext : 1 21 | · sorry 22 | · sorry 23 | --- 24 | warning: declaration uses 'sorry' 25 | -/ 26 | #guard_msgs in 27 | example {p q : MyProd α β} : p = q := by 28 | aesop? (add safe Aesop.BuiltinRules.ext) 29 | (config := { warnOnNonterminal := false }) 30 | all_goals sorry 31 | 32 | /-- 33 | info: Try this: 34 | ext : 1 35 | · sorry 36 | · ext : 1 37 | · sorry 38 | · sorry 39 | --- 40 | warning: declaration uses 'sorry' 41 | -/ 42 | #guard_msgs in 43 | example {p q : MyProd α (MyProd β γ)} : p = q := by 44 | aesop? (add safe Aesop.BuiltinRules.ext) 45 | (config := { warnOnNonterminal := false }) 46 | all_goals sorry 47 | 48 | /-- 49 | info: Try this: 50 | ext : 1 51 | · ext x : 1 52 | sorry 53 | · sorry 54 | --- 55 | warning: declaration uses 'sorry' 56 | -/ 57 | #guard_msgs in 58 | example {p q : MyProd (α → β) γ} : p = q := by 59 | aesop? (add safe Aesop.BuiltinRules.ext) 60 | (config := { warnOnNonterminal := false }) 61 | all_goals sorry 62 | 63 | /-- 64 | info: Try this: 65 | ext : 1 66 | · ext x : 1 67 | sorry 68 | · ext x x_1 : 2 69 | sorry 70 | --- 71 | warning: declaration uses 'sorry' 72 | -/ 73 | #guard_msgs in 74 | example {p q : (α → β) × (γ → δ → ι)} : p = q := by 75 | aesop? (add safe Aesop.BuiltinRules.ext) 76 | (erase Aesop.BuiltinRules.destructProducts) 77 | (config := { warnOnNonterminal := false }) 78 | all_goals sorry 79 | 80 | -- This test case checks script generation for ext calls which generate subgoals 81 | -- with different numbers of new hypotheses. 82 | 83 | axiom T : Type 84 | axiom U : Type 85 | axiom u : U 86 | axiom v : U 87 | 88 | @[ext (iff := false)] 89 | axiom T_ext : ∀ x y : T, u = v → (∀ u v : U, u = v) → x = y 90 | 91 | /-- 92 | info: Try this: 93 | ext u v : 1 94 | · sorry 95 | · sorry 96 | --- 97 | warning: declaration uses 'sorry' 98 | -/ 99 | #guard_msgs in 100 | example (x y : T) : x = y := by 101 | aesop? (add safe Aesop.BuiltinRules.ext) 102 | (config := { warnOnNonterminal := false }) 103 | all_goals sorry 104 | -------------------------------------------------------------------------------- /AesopTest/Filter.lean: -------------------------------------------------------------------------------- 1 | /- 2 | Copyright (c) 2022 Jannis Limperg. All rights reserved. 3 | Released under Apache 2.0 license as described in the file LICENSE. 4 | Authors: Sebastian Ullrich, Jannis Limperg 5 | -/ 6 | 7 | import Aesop 8 | 9 | set_option aesop.check.all true 10 | 11 | namespace TBA 12 | 13 | inductive List (α : Type) where 14 | | nil : List α 15 | | cons (head : α) (tail : List α) : List α 16 | 17 | notation (priority := high) "[" "]" => List.nil 18 | infixr:67 (priority := high) " :: " => List.cons 19 | 20 | def filter (p : α → Prop) [DecidablePred p] (as : List α) : List α := 21 | match as with 22 | | [] => [] 23 | | a::as => if p a then a :: filter p as else filter p as 24 | 25 | variable {p : α → Prop} [DecidablePred p] {as bs : List α} 26 | 27 | @[simp] 28 | theorem filter_cons_true (h : p a) : filter p (a :: as) = a :: filter p as := by 29 | simp [filter, h] 30 | 31 | @[simp] 32 | theorem filter_cons_false (h : ¬ p a) : filter p (a :: as) = filter p as := by 33 | simp [filter, h] 34 | 35 | @[aesop 50% [constructors, cases]] 36 | inductive Mem (a : α) : List α → Prop where 37 | | head {as} : Mem a (a::as) 38 | | tail {as} : Mem a as → Mem a (a'::as) 39 | 40 | infix:50 " ∈ " => Mem 41 | 42 | theorem mem_filter : a ∈ filter p as ↔ a ∈ as ∧ p a := by 43 | apply Iff.intro 44 | case mp => 45 | intro h 46 | induction as with 47 | | nil => cases h 48 | | cons a' as ih => by_cases ha' : p a' <;> aesop 49 | case mpr => 50 | intro h 51 | induction as with 52 | | nil => cases h.1 53 | | cons a' as ih => 54 | cases h.1 with 55 | | head => 56 | rw [filter_cons_true h.2] 57 | constructor 58 | | tail ha => 59 | have : a ∈ filter p as := ih ⟨ha, h.2⟩ 60 | by_cases hpa' : p a' 61 | case pos => 62 | rw [filter_cons_true hpa'] 63 | exact Mem.tail this 64 | case neg => 65 | rw [filter_cons_false hpa'] 66 | exact this 67 | 68 | end TBA 69 | -------------------------------------------------------------------------------- /AesopTest/ForwardConstant.lean: -------------------------------------------------------------------------------- 1 | /- 2 | Copyright (c) 2024 Jannis Limperg. All rights reserved. 3 | Released under Apache 2.0 license as described in the file LICENSE. 4 | Authors: Jannis Limperg 5 | -/ 6 | 7 | import Aesop 8 | 9 | set_option aesop.check.all true 10 | set_option aesop.smallErrorMessages true 11 | 12 | structure Foo where 13 | foo :: 14 | 15 | /-- 16 | error: tactic 'aesop' failed, made no progress 17 | -/ 18 | #guard_msgs in 19 | example : Foo := by 20 | aesop 21 | 22 | example : Foo := by 23 | aesop (add safe forward Foo.foo) 24 | -------------------------------------------------------------------------------- /AesopTest/ForwardRedundantHypsWithMVars.lean: -------------------------------------------------------------------------------- 1 | /- 2 | Copyright (c) 2024 Jannis Limperg. All rights reserved. 3 | Released under Apache 2.0 license as described in the file LICENSE. 4 | Authors: Son Ho, Jannis Limperg 5 | -/ 6 | 7 | -- Regression test for a failure of the `forward` rule builder to realise that 8 | -- it had already added certain hypotheses. The `saturate` loop below adds the 9 | -- same hyp over and over again, only with different level mvars. Thanks to Son 10 | -- Ho for providing this test case. 11 | 12 | import Aesop 13 | 14 | set_option aesop.check.all true 15 | 16 | axiom len {α} (ls : List α) : Int 17 | 18 | @[aesop safe forward (pattern := len ls)] 19 | axiom len_pos {α} {ls : List α} : 0 ≤ len (ls : List α) 20 | 21 | /-- 22 | error: unsolved goals 23 | a : Type u_1 24 | l : List a 25 | fwd : 0 ≤ len l 26 | ⊢ 0 ≤ len l 27 | -/ 28 | #guard_msgs in 29 | example (l: List a): 0 ≤ len l := by 30 | saturate 31 | -------------------------------------------------------------------------------- /AesopTest/ForwardTransparency.lean: -------------------------------------------------------------------------------- 1 | /- 2 | Copyright (c) 2023 Jannis Limperg. All rights reserved. 3 | Released under Apache 2.0 license as described in the file LICENSE. 4 | Authors: Jannis Limperg 5 | -/ 6 | 7 | import Aesop 8 | 9 | set_option aesop.check.all true 10 | set_option aesop.smallErrorMessages true 11 | 12 | -- Forward rules always operate at reducible transparency. 13 | 14 | def T := Unit → Empty 15 | 16 | variable {α : Type} 17 | 18 | /-- 19 | error: tactic 'aesop' failed, failed to prove the goal after exhaustive search. 20 | -/ 21 | #guard_msgs in 22 | example (h : T) (u : Unit) : α := by 23 | aesop (config := { terminal := true }) 24 | 25 | def U := Unit 26 | 27 | /-- 28 | error: tactic 'aesop' failed, failed to prove the goal after exhaustive search. 29 | -/ 30 | #guard_msgs in 31 | example (h : T) (u : U) : α := by 32 | aesop (config := { terminal := true }) 33 | 34 | /-- 35 | error: tactic 'aesop' failed, failed to prove the goal after exhaustive search. 36 | -/ 37 | #guard_msgs in 38 | example (h : T) (u : U) : α := by 39 | aesop (add forward safe h) (config := { terminal := true }) 40 | 41 | abbrev V := Unit 42 | 43 | /-- 44 | error: tactic 'aesop' failed, failed to prove the goal after exhaustive search. 45 | -/ 46 | #guard_msgs in 47 | example (h : Unit → Empty) (u : V) : α := by 48 | aesop (config := { terminal := true }) 49 | 50 | example (h : Unit → Empty) (u : V) : α := by 51 | aesop (add forward safe h) 52 | -------------------------------------------------------------------------------- /AesopTest/ForwardUnknownFVar.lean: -------------------------------------------------------------------------------- 1 | /- 2 | Copyright (c) 2025 Xavier Généreux. All rights reserved. 3 | Released under Apache 2.0 license as described in the file LICENSE. 4 | Authors: Xavier Généreux, Jannis Limperg 5 | -/ 6 | 7 | import Aesop 8 | 9 | attribute [aesop safe forward 1] Nat.le_trans 10 | attribute [-aesop] Aesop.BuiltinRules.applyHyps 11 | 12 | variable (a b c d e f : Nat) 13 | 14 | /-- error: tactic 'aesop' failed, failed to prove the goal after exhaustive search. -/ 15 | #guard_msgs in 16 | set_option aesop.smallErrorMessages true in 17 | example (h₃ : c ≤ d) (h₄ : d ≤ e) (h₅ : e ≤ f) (h : a = f) : False := by 18 | aesop (config := { terminal := true }) 19 | -------------------------------------------------------------------------------- /AesopTest/GlobalRuleIdentErrorChecking.lean: -------------------------------------------------------------------------------- 1 | /- 2 | Copyright (c) 2024 Jannis Limperg. All rights reserved. 3 | Released under Apache 2.0 license as described in the file LICENSE. 4 | Authors: Jannis Limperg 5 | -/ 6 | 7 | import Aesop 8 | 9 | /-- 10 | error: duplicate rule 'Nat.add_assoc'; rule 'bar' was already given. 11 | Use [,...] to give multiple rules. 12 | -/ 13 | #guard_msgs in 14 | @[aesop norm simp Nat.add_assoc] 15 | theorem bar : True := trivial 16 | -------------------------------------------------------------------------------- /AesopTest/HeartbeatLimit.lean: -------------------------------------------------------------------------------- 1 | /- 2 | Copyright (c) 2023 Jannis Limperg. All rights reserved. 3 | Released under Apache 2.0 license as described in the file LICENSE. 4 | Authors: Jannis Limperg 5 | -/ 6 | 7 | import Aesop 8 | 9 | set_option aesop.check.all true 10 | set_option aesop.smallErrorMessages true 11 | set_option linter.unreachableTactic false 12 | 13 | @[aesop safe constructors] 14 | inductive Even : Nat → Prop 15 | | zero : Even 0 16 | | plusTwo : Even n → Even (n + 2) 17 | 18 | /-- error: tactic 'aesop' failed, failed to prove the goal after exhaustive search. -/ 19 | #guard_msgs in 20 | example : Even 10 := by 21 | aesop (config := { maxRuleHeartbeats := 1, terminal := true }) 22 | 23 | example : Even 10 := by 24 | aesop 25 | 26 | /-- 27 | error: aesop: error in norm simp: tactic 'simp' failed, nested error: 28 | (deterministic) timeout at `simp`, maximum number of heartbeats (1) has been reached 29 | Use `set_option maxHeartbeats ` to set the limit. 30 | 31 | Additional diagnostic information may be available using the `set_option diagnostics true` command. 32 | -/ 33 | #guard_msgs in 34 | example (n m k : Nat) : n + m + k = (n + m) + k := by 35 | aesop (config := { maxSimpHeartbeats := 1, terminal := true }) 36 | 37 | example (n m k : Nat) : n + m + k = (n + m) + k := by 38 | aesop 39 | -------------------------------------------------------------------------------- /AesopTest/IncompleteScript.lean: -------------------------------------------------------------------------------- 1 | /- 2 | Copyright (c) 2024 Jannis Limperg. All rights reserved. 3 | Released under Apache 2.0 license as described in the file LICENSE. 4 | Authors: Jannis Limperg 5 | -/ 6 | 7 | import Aesop 8 | 9 | set_option aesop.check.all true 10 | set_option aesop.smallErrorMessages true 11 | 12 | -- set_option aesop.check.script.steps false 13 | 14 | /-- 15 | info: Try this: 16 | intro a 17 | simp_all only [Nat.add_zero] 18 | sorry 19 | --- 20 | warning: aesop: failed to prove the goal after exhaustive search. 21 | --- 22 | error: unsolved goals 23 | n : Nat 24 | P : Nat → Prop 25 | a : P n 26 | ⊢ P (n + 1) 27 | -/ 28 | #guard_msgs in 29 | example {P : Nat → Prop} : P (n + 0) → P (n + 1) := by 30 | aesop? 31 | 32 | inductive Even : Nat → Prop where 33 | | zero : Even 0 34 | | add_two : Even n → Even (n + 2) 35 | 36 | /-- 37 | info: Try this: 38 | sorry 39 | --- 40 | error: tactic 'aesop' failed, made no progress 41 | -/ 42 | #guard_msgs in 43 | example : Even 5 := by 44 | aesop? 45 | 46 | attribute [aesop safe constructors] Even 47 | 48 | /-- 49 | info: Try this: 50 | apply Even.add_two 51 | apply Even.add_two 52 | sorry 53 | --- 54 | warning: aesop: failed to prove the goal after exhaustive search. 55 | --- 56 | error: unsolved goals 57 | case a.a 58 | ⊢ Even 1 59 | -/ 60 | #guard_msgs in 61 | example : Even 5 := by 62 | aesop? 63 | -------------------------------------------------------------------------------- /AesopTest/Intros.lean: -------------------------------------------------------------------------------- 1 | /- 2 | Copyright (c) 2022 Jannis Limperg. All rights reserved. 3 | Released under Apache 2.0 license as described in the file LICENSE. 4 | Authors: Jannis Limperg 5 | -/ 6 | 7 | import Aesop 8 | 9 | set_option aesop.check.all true 10 | set_option aesop.smallErrorMessages true 11 | 12 | def Injective₁ (f : α → β) := ∀ x y, f x = f y → x = y 13 | 14 | abbrev Injective₂ (f : α → β) := ∀ x y, f x = f y → x = y 15 | 16 | /-- 17 | error: tactic 'aesop' failed, failed to prove the goal after exhaustive search. 18 | -/ 19 | #guard_msgs in 20 | example : Injective₁ (@id Nat) := by 21 | aesop (config := { terminal := true }) 22 | 23 | /-- 24 | error: tactic 'aesop' failed, failed to prove the goal after exhaustive search. 25 | -/ 26 | #guard_msgs in 27 | example : Injective₁ (@id Nat) := by 28 | aesop (config := { introsTransparency? := some .reducible, terminal := true }) 29 | 30 | example : Injective₁ (@id Nat) := by 31 | aesop (config := { introsTransparency? := some .default }) 32 | 33 | /-- 34 | error: tactic 'aesop' failed, failed to prove the goal after exhaustive search. 35 | -/ 36 | #guard_msgs in 37 | example : Injective₂ (@id Nat) := by 38 | aesop (config := { terminal := true }) 39 | 40 | example : Injective₂ (@id Nat) := by 41 | aesop (config := { introsTransparency? := some .reducible }) 42 | 43 | example : Injective₂ (@id Nat) := by 44 | aesop (config := { introsTransparency? := some .default }) 45 | -------------------------------------------------------------------------------- /AesopTest/IntrosAllTransparency.lean: -------------------------------------------------------------------------------- 1 | /- 2 | Copyright (c) 2023 Jannis Limperg. All rights reserved. 3 | Released under Apache 2.0 license as described in the file LICENSE. 4 | Authors: Jannis Limperg 5 | -/ 6 | 7 | import Aesop 8 | 9 | set_option aesop.check.all true 10 | set_option aesop.smallErrorMessages true 11 | 12 | @[irreducible] def T := False → True 13 | 14 | /-- 15 | error: tactic 'aesop' failed, failed to prove the goal after exhaustive search. 16 | -/ 17 | #guard_msgs in 18 | example : T := by 19 | aesop (config := { terminal := true }) 20 | 21 | example : T := by 22 | aesop (config := { introsTransparency? := some .all }) 23 | -------------------------------------------------------------------------------- /AesopTest/Jesse.lean: -------------------------------------------------------------------------------- 1 | /- 2 | Copyright (c) 2022 Jannis Limperg. All rights reserved. 3 | Released under Apache 2.0 license as described in the file LICENSE. 4 | Authors: Jesse Vogel, Jannis Limperg 5 | -/ 6 | 7 | import Aesop 8 | 9 | set_option aesop.check.all true 10 | 11 | axiom Ring : Type 12 | axiom Morphism (R S : Ring) : Type 13 | 14 | @[aesop 99%] 15 | axiom ZZ : Ring 16 | 17 | @[aesop 99%] 18 | axiom f : Morphism ZZ ZZ 19 | 20 | noncomputable example : Σ (R : Ring), Morphism R R := by 21 | aesop 22 | 23 | axiom domain (R : Ring) : Prop 24 | 25 | @[aesop 99%] 26 | axiom ZZ_domain : domain ZZ 27 | 28 | noncomputable example : ∃ (R : Ring), domain R := by 29 | aesop 30 | -------------------------------------------------------------------------------- /AesopTest/LocalRuleSet.lean: -------------------------------------------------------------------------------- 1 | /- 2 | Copyright (c) 2022 Jannis Limperg. All rights reserved. 3 | Released under Apache 2.0 license as described in the file LICENSE. 4 | Authors: Jannis Limperg 5 | -/ 6 | 7 | import Aesop 8 | 9 | set_option aesop.check.all true 10 | set_option aesop.smallErrorMessages true 11 | 12 | -- We used to add local rules to the `default` rule set, but this doesn't work 13 | -- well when the default rule set is disabled. 14 | 15 | /-- 16 | error: tactic 'aesop' failed, failed to prove the goal after exhaustive search. 17 | -/ 18 | #guard_msgs in 19 | example : Unit := by 20 | aesop (rule_sets := [-default, -builtin]) (config := { terminal := true }) 21 | 22 | example : Unit := by 23 | aesop (add safe PUnit.unit) (rule_sets := [-default, -builtin]) 24 | -------------------------------------------------------------------------------- /AesopTest/LocalTactic.lean: -------------------------------------------------------------------------------- 1 | /- 2 | Copyright (c) 2024 Jannis Limperg. All rights reserved. 3 | Released under Apache 2.0 license as described in the file LICENSE. 4 | Authors: Jannis Limperg 5 | -/ 6 | 7 | import Aesop 8 | 9 | set_option aesop.check.all true 10 | set_option aesop.smallErrorMessages true 11 | 12 | /-- 13 | error: tactic 'aesop' failed, made no progress 14 | -/ 15 | #guard_msgs in 16 | example (m n : Nat) : m * n = n * m := by 17 | aesop 18 | 19 | example (m n : Nat) : m * n = n * m := by 20 | aesop (add safe (by rw [Nat.mul_comm])) 21 | 22 | example (m n : Nat) : m * n = n * m := by 23 | aesop (add safe tactic (by rw [Nat.mul_comm m n])) 24 | 25 | /-- 26 | error: tactic 'aesop' failed, made no progress 27 | -/ 28 | #guard_msgs in 29 | example (m n : Nat) : m * n = n * m := by 30 | aesop (add safe (by rw [Nat.mul_comm m m])) 31 | 32 | example (m n : Nat) : m * n = n * m := by 33 | aesop (add safe (by apply Nat.mul_comm; done)) 34 | -------------------------------------------------------------------------------- /AesopTest/Logic.lean: -------------------------------------------------------------------------------- 1 | /- 2 | Copyright (c) 2021 Jannis Limperg. All rights reserved. 3 | Released under Apache 2.0 license as described in the file LICENSE. 4 | Authors: Jannis Limperg 5 | -/ 6 | 7 | import Aesop 8 | 9 | set_option aesop.check.all true 10 | 11 | example : True := by 12 | aesop 13 | 14 | example : Unit := by 15 | aesop 16 | 17 | example : PUnit.{u} := by 18 | aesop 19 | 20 | example (h : False) : α := by 21 | aesop 22 | 23 | example (h : Empty) : α := by 24 | aesop 25 | 26 | example (h : PEmpty.{u}) : α := by 27 | aesop 28 | -------------------------------------------------------------------------------- /AesopTest/MVarsInInitialGoal.lean: -------------------------------------------------------------------------------- 1 | /- 2 | Copyright (c) 2022 Jannis Limperg. All rights reserved. 3 | Released under Apache 2.0 license as described in the file LICENSE. 4 | Authors: Jannis Limperg 5 | -/ 6 | 7 | import Aesop 8 | 9 | set_option aesop.check.all true 10 | 11 | example {n m k : Nat} (h : n < m) (h₂ : m < k) : n < k := by 12 | apply Nat.lt_trans <;> aesop 13 | -------------------------------------------------------------------------------- /AesopTest/Metas.lean: -------------------------------------------------------------------------------- 1 | /- 2 | Copyright (c) 2021 Jannis Limperg. All rights reserved. 3 | Released under Apache 2.0 license as described in the file LICENSE. 4 | Authors: Jannis Limperg 5 | -/ 6 | 7 | import Aesop 8 | 9 | set_option aesop.check.all true 10 | 11 | example (P : α → Prop) (a : α) (h : P a) : ∃ a, P a := by 12 | aesop 13 | 14 | example (P : α → β → Prop) (a : α) (b : β) (h : P a b) : ∃ a b, P a b := by 15 | aesop 16 | 17 | example (P : α → Type) (a : α) (h : P a) : Σ a, P a := by 18 | aesop 19 | 20 | example (P : α → β → Type) (a : α) (b : β) (h : P a b) : Σ a b, P a b := by 21 | aesop 22 | 23 | example (P Q : α → Prop) (hPQ : ∀ a, P a → Q a) (a : α) (h : P a) : ∃ a, Q a := by 24 | aesop 25 | 26 | example (P Q Dead R : α → Prop) 27 | (hPQ : ∀ a, P a → Q a) 28 | (hDeadR : ∀ a, Dead a → R a) 29 | (hQR : ∀ a, Q a → R a) 30 | (a : α) (h : P a) : 31 | ∃ a, R a := by 32 | aesop 33 | 34 | example (R : α → α → Prop) (R_trans : ∀ x y z, R x y → R y z → R x z) (a b c d) 35 | (hab : R a b) (hbc : R b c) (hcd : R c d) : R a d := by 36 | aesop 37 | -------------------------------------------------------------------------------- /AesopTest/NameResolution.lean: -------------------------------------------------------------------------------- 1 | /- 2 | Copyright (c) 2022 Jannis Limperg. All rights reserved. 3 | Released under Apache 2.0 license as described in the file LICENSE. 4 | Authors: Jannis Limperg 5 | -/ 6 | import Aesop 7 | 8 | set_option aesop.check.all true 9 | 10 | -- When parsing the names of declarations for local rules, Aesop should take the 11 | -- currently opened namespaces into account. 12 | 13 | example : List α := by 14 | aesop (add safe List.nil) 15 | 16 | namespace List 17 | 18 | example : List α := by 19 | aesop (add safe List.nil) 20 | 21 | example : List α := by 22 | aesop (add safe nil) 23 | 24 | end List 25 | -------------------------------------------------------------------------------- /AesopTest/NoImportClash.lean: -------------------------------------------------------------------------------- 1 | /- 2 | Copyright (c) 2023 Jannis Limperg. All rights reserved. 3 | Released under Apache 2.0 license as described in the file LICENSE. 4 | Authors: Jannis Limperg 5 | -/ 6 | 7 | /- 8 | This test case ensures that Aesop does not define anything that already exists 9 | in our dependencies. 10 | -/ 11 | 12 | import Aesop 13 | import Lean 14 | import Batteries 15 | -------------------------------------------------------------------------------- /AesopTest/NoNormSimp.lean: -------------------------------------------------------------------------------- 1 | /- 2 | Copyright (c) 2023 Jannis Limperg. All rights reserved. 3 | Released under Apache 2.0 license as described in the file LICENSE. 4 | Authors: Jannis Limperg 5 | -/ 6 | 7 | import Aesop 8 | 9 | set_option aesop.check.all true 10 | set_option aesop.smallErrorMessages true 11 | 12 | /-- 13 | error: tactic 'aesop' failed, failed to prove the goal after exhaustive search. 14 | -/ 15 | #guard_msgs in 16 | example : 1 = 2 → False := by 17 | aesop (config := { enableSimp := false, terminal := true }) 18 | 19 | example : 1 = 2 → False := by 20 | aesop 21 | -------------------------------------------------------------------------------- /AesopTest/NoProgress.lean: -------------------------------------------------------------------------------- 1 | /- 2 | Copyright (c) 2023 Jannis Limperg. All rights reserved. 3 | Released under Apache 2.0 license as described in the file LICENSE. 4 | Authors: Jannis Limperg 5 | -/ 6 | 7 | import Aesop 8 | 9 | set_option aesop.check.all true 10 | set_option aesop.smallErrorMessages true 11 | 12 | def T := True 13 | 14 | /-- 15 | error: tactic 'aesop' failed, made no progress 16 | -/ 17 | #guard_msgs in 18 | example : T := by 19 | aesop 20 | 21 | @[aesop norm simp] 22 | def F := False 23 | 24 | /-- 25 | warning: aesop: failed to prove the goal after exhaustive search. 26 | --- 27 | error: unsolved goals 28 | ⊢ False 29 | -/ 30 | #guard_msgs in 31 | example : F := by 32 | aesop 33 | 34 | def F' := False 35 | 36 | @[aesop safe apply] 37 | theorem F'_def : False → F' := id 38 | 39 | /-- 40 | warning: aesop: failed to prove the goal after exhaustive search. 41 | --- 42 | error: unsolved goals 43 | ⊢ False 44 | -/ 45 | #guard_msgs in 46 | example : F' := by 47 | aesop 48 | 49 | attribute [-aesop] F'_def 50 | attribute [aesop 100%] F'_def 51 | 52 | -- When an unsafe rule is applied, we don't count this as progress because the 53 | -- remaining goal that the user gets to see is exactly the same as the initial 54 | -- goal. 55 | 56 | /-- 57 | error: tactic 'aesop' failed, made no progress 58 | -/ 59 | #guard_msgs in 60 | example : F' := by 61 | aesop 62 | -------------------------------------------------------------------------------- /AesopTest/Nonterminal.lean: -------------------------------------------------------------------------------- 1 | /- 2 | Copyright (c) 2022 Jannis Limperg. All rights reserved. 3 | Released under Apache 2.0 license as described in the file LICENSE. 4 | Authors: Jannis Limperg 5 | -/ 6 | 7 | import Aesop 8 | 9 | set_option aesop.check.all true 10 | set_option aesop.smallErrorMessages true 11 | 12 | structure MyTrue₁ 13 | structure MyTrue₂ 14 | 15 | @[aesop safe] 16 | structure MyTrue₃ where 17 | tt : MyTrue₁ 18 | 19 | /-- 20 | warning: aesop: failed to prove the goal after exhaustive search. 21 | -/ 22 | #guard_msgs in 23 | example : MyTrue₃ := by 24 | aesop 25 | apply MyTrue₁.mk 26 | 27 | @[aesop safe] 28 | structure MyFalse where 29 | falso : False 30 | 31 | /-- 32 | warning: aesop: failed to prove the goal after exhaustive search. 33 | --- 34 | error: unsolved goals 35 | ⊢ False 36 | -/ 37 | #guard_msgs in 38 | example : MyFalse := by 39 | aesop 40 | 41 | /-- 42 | error: tactic 'aesop' failed, failed to prove the goal after exhaustive search. 43 | -/ 44 | #guard_msgs in 45 | example : MyFalse := by 46 | aesop (config := { terminal := true }) 47 | 48 | /-- 49 | error: unsolved goals 50 | ⊢ False 51 | -/ 52 | #guard_msgs in 53 | example : MyFalse := by 54 | aesop (config := { warnOnNonterminal := false }) 55 | 56 | @[aesop safe] 57 | structure MyFalse₂ where 58 | falso : False 59 | tt : MyTrue₃ 60 | 61 | /-- 62 | warning: aesop: failed to prove the goal after exhaustive search. 63 | --- 64 | error: unsolved goals 65 | case falso 66 | ⊢ False 67 | 68 | case tt 69 | ⊢ MyTrue₁ 70 | -/ 71 | #guard_msgs in 72 | example : MyFalse₂ := by 73 | aesop 74 | -------------------------------------------------------------------------------- /AesopTest/NormSimp.lean: -------------------------------------------------------------------------------- 1 | /- 2 | Copyright (c) 2022 Jannis Limperg. All rights reserved. 3 | Released under Apache 2.0 license as described in the file LICENSE. 4 | Authors: Jannis Limperg 5 | -/ 6 | 7 | import Aesop 8 | 9 | set_option aesop.check.all true 10 | set_option aesop.smallErrorMessages true 11 | 12 | -- A basic test for local simp rules. 13 | example {α : Prop} (h : α) : α := by 14 | aesop (rule_sets := [-builtin,-default]) (add h norm simp) 15 | 16 | -- This test checks that we don't 'self-simplify' hypotheses: `h` should not 17 | -- be used to simplify itself. 18 | example (h : (α ∧ β) ∨ γ) : α ∨ γ := by 19 | aesop (add h norm simp) 20 | 21 | -- Ditto, but we can omit the 'norm'. 22 | example (h : (α ∧ β) ∨ γ) : α ∨ γ := by 23 | aesop (add h simp) 24 | 25 | -- This test checks that the norm simp config is passed around properly. 26 | 27 | /-- 28 | error: tactic 'aesop' failed, failed to prove the goal after exhaustive search. 29 | -/ 30 | #guard_msgs in 31 | example {α β : Prop} (ha : α) (h : α → β) : β := by 32 | aesop (rule_sets := [-builtin,-default]) 33 | (simp_config := { maxDischargeDepth := 0 }) 34 | (config := { terminal := true, useDefaultSimpSet := false }) 35 | 36 | example {α β : Prop} (ha : α) (h : α → β) : β := by 37 | aesop (rule_sets := [-builtin,-default]) 38 | 39 | -- Ditto. 40 | 41 | /-- 42 | error: tactic 'aesop' failed, failed to prove the goal after exhaustive search. 43 | -/ 44 | #guard_msgs in 45 | example : true && false = false := by 46 | aesop (rule_sets := [-default,-builtin]) 47 | (simp_config := { decide := false }) 48 | (config := { terminal := true, useDefaultSimpSet := false }) 49 | 50 | example : true && false = false := by 51 | aesop (rule_sets := [-default,-builtin]) (simp_config := { decide := true }) 52 | 53 | -- We can use the `useSimpAll` config option to switch between `simp_all` and 54 | -- `simp at *`. 55 | 56 | /-- 57 | error: tactic 'aesop' failed, failed to prove the goal after exhaustive search. 58 | -/ 59 | #guard_msgs in 60 | example {α : Prop} (ha : α) : α := by 61 | aesop (rule_sets := [-builtin,-default]) 62 | (config := { useSimpAll := false, terminal := true }) 63 | 64 | example {α : Prop} (ha : α) : α := by 65 | aesop (rule_sets := [-builtin,-default]) 66 | 67 | -- We can give priorities to `simp` rules, corresponding to the priorities of 68 | -- `simp` lemmas. 69 | 70 | opaque T : Prop 71 | 72 | @[aesop simp 1] 73 | axiom TF : T ↔ False 74 | 75 | @[aesop simp 2] 76 | axiom TT : T ↔ True 77 | 78 | example : T := by aesop 79 | 80 | attribute [-aesop] TF 81 | attribute [aesop simp 3] TF 82 | 83 | /-- 84 | error: tactic 'aesop' failed, failed to prove the goal after exhaustive search. 85 | -/ 86 | #guard_msgs in 87 | example : T := by 88 | aesop (config := { terminal := true }) 89 | 90 | example : T := by 91 | rw [TT]; trivial 92 | -------------------------------------------------------------------------------- /AesopTest/Persistence0.lean: -------------------------------------------------------------------------------- 1 | /- 2 | Copyright (c) 2023 Jannis Limperg. All rights reserved. 3 | Released under Apache 2.0 license as described in the file LICENSE. 4 | Authors: Jannis Limperg 5 | -/ 6 | 7 | import Aesop 8 | 9 | set_option aesop.check.all true 10 | 11 | declare_aesop_rule_sets [test_persistence1] 12 | -------------------------------------------------------------------------------- /AesopTest/Persistence1.lean: -------------------------------------------------------------------------------- 1 | /- 2 | Copyright (c) 2022 Jannis Limperg. All rights reserved. 3 | Released under Apache 2.0 license as described in the file LICENSE. 4 | Authors: Jannis Limperg 5 | -/ 6 | 7 | import AesopTest.Persistence0 8 | 9 | set_option aesop.check.all true 10 | 11 | namespace Aesop 12 | 13 | @[aesop 50% constructors (rule_sets := [test_persistence1])] 14 | inductive NatOrBool where 15 | | ofNat (n : Nat) 16 | | ofBool (b : Bool) 17 | 18 | example (b : Bool) : NatOrBool := by 19 | aesop (rule_sets := [test_persistence1]) 20 | 21 | end Aesop 22 | -------------------------------------------------------------------------------- /AesopTest/Persistence2.lean: -------------------------------------------------------------------------------- 1 | /- 2 | Copyright (c) 2022 Jannis Limperg. All rights reserved. 3 | Released under Apache 2.0 license as described in the file LICENSE. 4 | Authors: Jannis Limperg 5 | -/ 6 | 7 | import Aesop 8 | 9 | set_option aesop.check.all true 10 | 11 | declare_aesop_rule_sets [test_persistence2] 12 | -------------------------------------------------------------------------------- /AesopTest/Persistence3.lean: -------------------------------------------------------------------------------- 1 | /- 2 | Copyright (c) 2022 Jannis Limperg. All rights reserved. 3 | Released under Apache 2.0 license as described in the file LICENSE. 4 | Authors: Jannis Limperg 5 | -/ 6 | 7 | import AesopTest.Persistence1 8 | import AesopTest.Persistence2 9 | 10 | set_option aesop.check.all true 11 | 12 | namespace Aesop 13 | 14 | @[aesop 50% (rule_sets := [test_persistence2])] 15 | def test (b : Bool) : NatOrBool := by 16 | aesop (rule_sets := [test_persistence1]) 17 | 18 | end Aesop 19 | -------------------------------------------------------------------------------- /AesopTest/PostponeSafeRules.lean: -------------------------------------------------------------------------------- 1 | /- 2 | Copyright (c) 2022 Jannis Limperg. All rights reserved. 3 | Released under Apache 2.0 license as described in the file LICENSE. 4 | Authors: Jannis Limperg 5 | -/ 6 | 7 | import Aesop 8 | 9 | set_option aesop.check.all true 10 | 11 | -- This is a test case for postponed safe rules. When a safe rule assigns mvars, 12 | -- it is not applied but instead postponed. Then we later try it again as an 13 | -- unsafe rule. 14 | -- 15 | -- It's hard to test this feature completely because we can't really tell from 16 | -- the outside when a rule has been applied. But we can at least look at the 17 | -- traces of the following test cases. 18 | 19 | axiom T : Nat → Prop 20 | 21 | @[aesop safe] 22 | axiom t : T 0 23 | 24 | example : ∃ (i : Nat), T i := by 25 | aesop 26 | 27 | 28 | axiom U : Nat → Prop 29 | 30 | @[aesop safe 0] 31 | axiom u₁ : U 0 32 | 33 | @[aesop safe 1] 34 | axiom u₂ : U 1 35 | 36 | example : ∃ i, U i ∧ i = 1 := by 37 | aesop 38 | -------------------------------------------------------------------------------- /AesopTest/RecursiveUnfoldRule.lean: -------------------------------------------------------------------------------- 1 | /- 2 | Copyright (c) 2023 Jannis Limperg. All rights reserved. 3 | Released under Apache 2.0 license as described in the file LICENSE. 4 | Authors: Jannis Limperg 5 | -/ 6 | 7 | import Aesop 8 | 9 | set_option aesop.check.all true 10 | 11 | -- We forbid `unfold` rules for recursive functions since they would lead to 12 | -- infinite unfolding. 13 | 14 | @[aesop norm unfold] 15 | def f₁ : Nat → Nat := 16 | λ _ => 0 17 | 18 | /-- 19 | error: Recursive definition 'f₂' cannot be used as an unfold rule (it would be unfolded infinitely often). Try adding a simp rule for it. 20 | -/ 21 | #guard_msgs in 22 | @[aesop norm unfold] 23 | def f₂ : Nat → Nat 24 | | 0 => 0 25 | | n + 1 => f₂ n 26 | 27 | /-- 28 | error: Recursive definition 'f₃' cannot be used as an unfold rule (it would be unfolded infinitely often). Try adding a simp rule for it. 29 | -/ 30 | #guard_msgs in 31 | @[aesop norm unfold] 32 | def f₃ : Nat → Nat := 33 | λ n => 34 | match n with 35 | | 0 => 0 36 | | n + 1 => f₃ n 37 | 38 | 39 | /-- 40 | error: Recursive definition 'f₄' cannot be used as an unfold rule (it would be unfolded infinitely often). Try adding a simp rule for it. 41 | -/ 42 | #guard_msgs in 43 | @[aesop norm unfold] 44 | def f₄ : Nat → Nat 45 | | 0 => 0 46 | | n + 1 => f₄ n 47 | 48 | -- A limitation of our approach to checking for recursive `unfold` rules: 49 | -- mutually recursive rules are not detected. But that's probably fine in 50 | -- practice. 51 | 52 | mutual 53 | @[aesop norm unfold] 54 | def f₅ : Nat → Nat 55 | | 0 => 0 56 | | n + 1 => f₆ n 57 | 58 | @[aesop norm unfold] 59 | def f₆ : Nat → Nat 60 | | 0 => 0 61 | | n + 1 => f₅ n 62 | end 63 | 64 | -- We also forbid `unfold` rules for declarations that don't unfold. 65 | 66 | /-- 67 | error: Declaration 'test' cannot be unfolded. 68 | -/ 69 | #guard_msgs in 70 | @[aesop norm unfold] 71 | axiom test : Nat 72 | -------------------------------------------------------------------------------- /AesopTest/RulePatternLooseBVar.lean: -------------------------------------------------------------------------------- 1 | /- 2 | Copyright (c) 2024 Son Ho. All rights reserved. 3 | Released under Apache 2.0 license as described in the file LICENSE. 4 | Authors: Son Ho, Jannis Limperg 5 | -/ 6 | 7 | -- A regression test for a bug in rule pattern matching. Thanks to Son Ho for 8 | -- the bug report. 9 | 10 | import Aesop 11 | 12 | inductive ScalarTy 13 | | U32 14 | 15 | @[simp] 16 | def U32.min : Int := 0 17 | def U32.max : Int := 4294967295 18 | 19 | def Scalar.min (ty : ScalarTy) : Int := 20 | match ty with 21 | | .U32 => U32.min 22 | 23 | def Scalar.max (ty : ScalarTy) : Int := 24 | match ty with 25 | | .U32 => U32.max 26 | 27 | structure Scalar (ty : ScalarTy) where 28 | val : Int 29 | hmin : Scalar.min ty ≤ val 30 | hmax : val ≤ Scalar.max ty 31 | 32 | @[reducible] def U32 := Scalar .U32 33 | 34 | def Scalar.ofIntCore {ty : ScalarTy} (x : Int) 35 | (h : Scalar.min ty ≤ x ∧ x ≤ Scalar.max ty) : Scalar ty := 36 | { val := x, hmin := h.left, hmax := h.right } 37 | 38 | def U32.ofIntCore := @Scalar.ofIntCore .U32 39 | 40 | @[aesop safe forward (pattern := x)] 41 | theorem Scalar.bounds {ty : ScalarTy} (x : Scalar ty) : 42 | Scalar.min ty ≤ x.val ∧ x.val ≤ Scalar.max ty := 43 | And.intro x.hmin x.hmax 44 | 45 | /-- 46 | error: unsolved goals 47 | x : Int 48 | h0 : 0 ≤ x 49 | h1 : x ≤ U32.max 50 | fwd : Scalar.min ScalarTy.U32 ≤ (U32.ofIntCore x ⋯).val ∧ (U32.ofIntCore x ⋯).val ≤ Scalar.max ScalarTy.U32 51 | ⊢ (U32.ofIntCore x ⋯).val ≤ U32.max 52 | -/ 53 | #guard_msgs in 54 | example (x : Int) (h0 : 0 ≤ x) (h1 : x ≤ U32.max) : 55 | (U32.ofIntCore x 56 | (⟨Eq.mp (congrArg (fun x_1 => x_1 ≤ x) rfl) h0, h1⟩)).val ≤ U32.max := by 57 | saturate 58 | -------------------------------------------------------------------------------- /AesopTest/RulePatternUniverseBug.lean: -------------------------------------------------------------------------------- 1 | /- 2 | Copyright (c) 2024 Jannis Limperg. All rights reserved. 3 | Released under Apache 2.0 license as described in the file LICENSE. 4 | Authors: Son Ho, Jannis Limperg 5 | -/ 6 | 7 | -- A regression test for an issue where universe mvars were not correctly 8 | -- assigned during rule pattern matching. Thanks to Son Ho for reporting this 9 | -- issue. 10 | 11 | import Aesop.Frontend.Attribute 12 | import Aesop.Frontend.Saturate 13 | 14 | set_option aesop.check.all true 15 | 16 | namespace List 17 | 18 | @[simp] 19 | def len (ls : List α) : Int := 20 | match ls with 21 | | [] => 0 22 | | _ :: tl => 1 + len tl 23 | 24 | @[aesop safe forward (pattern := len ls)] 25 | theorem len_pos : 0 ≤ len (ls : List α) := by 26 | induction ls <;> simp [*] 27 | omega 28 | 29 | def indexOpt (ls : List α) (i : Int) : Option α := 30 | match ls with 31 | | [] => none 32 | | hd :: tl => if i = 0 then some hd else indexOpt tl (i - 1) 33 | 34 | theorem indexOpt_bounds (ls : List α) (i : Int) : 35 | ls.indexOpt i = none ↔ i < 0 ∨ ls.len ≤ i := by 36 | match ls with 37 | | [] => simp [indexOpt]; omega 38 | | _ :: tl => 39 | have := indexOpt_bounds tl (i - 1) 40 | if h : i = 0 then 41 | simp [indexOpt, *] 42 | saturate 43 | omega 44 | else 45 | simp [indexOpt, len, *] 46 | constructor <;> intro a <;> cases a 47 | . left 48 | saturate 49 | omega 50 | . right; omega 51 | . left; omega 52 | . right; omega 53 | -------------------------------------------------------------------------------- /AesopTest/RuleSetNameHygiene0.lean: -------------------------------------------------------------------------------- 1 | /- 2 | Copyright (c) 2023 Jannis Limperg. All rights reserved. 3 | Released under Apache 2.0 license as described in the file LICENSE. 4 | Authors: Jannis Limperg 5 | -/ 6 | 7 | import Aesop 8 | 9 | set_option aesop.check.all true 10 | 11 | declare_aesop_rule_sets [test] 12 | -------------------------------------------------------------------------------- /AesopTest/RuleSetNameHygiene1.lean: -------------------------------------------------------------------------------- 1 | /- 2 | Copyright (c) 2023 Jannis Limperg. All rights reserved. 3 | Released under Apache 2.0 license as described in the file LICENSE. 4 | Authors: Jannis Limperg 5 | -/ 6 | 7 | import AesopTest.RuleSetNameHygiene0 8 | 9 | set_option aesop.check.all true 10 | set_option aesop.smallErrorMessages true 11 | 12 | macro "aesop_test" : tactic => `(tactic| aesop (rule_sets := [test])) 13 | 14 | @[aesop safe (rule_sets := [test])] 15 | structure TT where 16 | 17 | /-- 18 | error: tactic 'aesop' failed, failed to prove the goal after exhaustive search. 19 | -/ 20 | #guard_msgs in 21 | example : TT := by 22 | aesop (config := { terminal := true }) 23 | 24 | example : TT := by 25 | aesop_test 26 | -------------------------------------------------------------------------------- /AesopTest/RuleSets0.lean: -------------------------------------------------------------------------------- 1 | /- 2 | Copyright (c) 2022-2023 Jannis Limperg. All rights reserved. 3 | Released under Apache 2.0 license as described in the file LICENSE. 4 | Authors: Jannis Limperg 5 | -/ 6 | 7 | import Aesop 8 | 9 | set_option aesop.check.all true 10 | 11 | declare_aesop_rule_sets [test_A, test_B, test_C, test_D] 12 | -------------------------------------------------------------------------------- /AesopTest/Safe.lean: -------------------------------------------------------------------------------- 1 | /- 2 | Copyright (c) 2022 Jannis Limperg. All rights reserved. 3 | Released under Apache 2.0 license as described in the file LICENSE. 4 | Authors: Jannis Limperg 5 | -/ 6 | 7 | import Aesop 8 | 9 | set_option aesop.check.all true 10 | set_option aesop.smallErrorMessages true 11 | 12 | @[aesop 50%] 13 | inductive Even : Nat → Prop 14 | | zero : Even 0 15 | | plus_two {n} : Even n → Even (n + 2) 16 | 17 | structure Even' (n) : Prop where 18 | even : Even n 19 | 20 | theorem even'_of_false : False → Even' n 21 | | h => nomatch h 22 | 23 | theorem even'_of_even' : Even' n → Even' n := 24 | id 25 | 26 | -- Once a safe rule is applied, the corresponding goal is marked as inactive 27 | -- and never visited again. 28 | 29 | /-- 30 | error: tactic 'aesop' failed, failed to prove the goal after exhaustive search. 31 | -/ 32 | #guard_msgs in 33 | example : Even' 2 := by 34 | aesop (add safe [even'_of_false 0, Even'.mk 1]) 35 | (config := { terminal := true }) 36 | 37 | /-- 38 | error: tactic 'aesop' failed, maximum number of rule applications (10) reached. Set the 'maxRuleApplications' option to increase the limit. 39 | -/ 40 | #guard_msgs in 41 | example : Even' 2 := by 42 | aesop (add safe [even'_of_even'], unsafe [Even'.mk 100%]) 43 | (config := { maxRuleApplications := 10, terminal := true }) 44 | 45 | example : Even' 2 := by 46 | aesop (add safe [even'_of_false 1, Even'.mk 0]) 47 | -------------------------------------------------------------------------------- /AesopTest/SafeExtractionCopyIntroducedMVars.lean: -------------------------------------------------------------------------------- 1 | /- 2 | Copyright (c) 2023 Jannis Limperg. All rights reserved. 3 | Released under Apache 2.0 license as described in the file LICENSE. 4 | Authors: Jannis Limperg 5 | -/ 6 | 7 | import Aesop 8 | 9 | set_option aesop.check.all true 10 | 11 | axiom P : Prop 12 | axiom T : Prop 13 | axiom Q : Nat → Prop 14 | axiom R : Nat → Prop 15 | axiom S : Nat → Prop 16 | 17 | @[aesop safe] 18 | axiom q_r_p : ∀ x, Q x → R x → P 19 | 20 | @[aesop safe] 21 | axiom s_q : ∀ x y, S y → Q x 22 | 23 | @[aesop safe] 24 | axiom s_r : ∀ x y, S y → R x 25 | 26 | axiom s : S 0 27 | 28 | /-- 29 | warning: aesop: failed to prove the goal after exhaustive search. 30 | --- 31 | error: unsolved goals 32 | case a.a 33 | ⊢ S ?a.y✝ 34 | 35 | case a.a 36 | ⊢ S ?a.y✝ 37 | 38 | case x 39 | ⊢ Nat 40 | 41 | case a.a 42 | ⊢ S ?a.y✝ 43 | 44 | case a.a 45 | ⊢ S ?a.y✝ 46 | 47 | case x 48 | ⊢ Nat 49 | --- 50 | error: (kernel) declaration has metavariables '_example' 51 | -/ 52 | #guard_msgs in 53 | example : P := by 54 | aesop 55 | -------------------------------------------------------------------------------- /AesopTest/SafePrefixExpansionRappLimit.lean: -------------------------------------------------------------------------------- 1 | /- 2 | Copyright (c) 2024 Jannis Limperg. All rights reserved. 3 | Released under Apache 2.0 license as described in the file LICENSE. 4 | Authors: Jannis Limperg 5 | -/ 6 | 7 | import Aesop 8 | 9 | @[aesop safe] 10 | axiom loopy {α : Prop} : α ∨ α → α 11 | 12 | /-- 13 | warning: aesop: failed to prove the goal. Some goals were not explored because the maximum rule application depth (30) was reached. Set option 'maxRuleApplicationDepth' to increase the limit. 14 | --- 15 | warning: aesop: safe prefix was not fully expanded because the maximum number of rule applications (50) was reached. 16 | --- 17 | error: unsolved goals 18 | case a.a.a.a.a.a.a.a.a.a.a.a.a.a.a.a.a.a.a.a.a.a.a.a.a.a.a.a.a.a.a.a.a.a.a.a.a.a.a.a.a.a.a.a.a.a.a.a.a.a.a.a.a.a.a.a.a.a.a.a.a.a.a.a.a.a.a.a.a.a.a.a.a.a.a.a.a.a.a.a 19 | ⊢ False 20 | -/ 21 | #guard_msgs in 22 | example : False := by 23 | aesop 24 | 25 | /-- 26 | error: tactic 'aesop' failed, failed to prove the goal. Some goals were not explored because the maximum rule application depth (30) was reached. Set option 'maxRuleApplicationDepth' to increase the limit. 27 | Initial goal: 28 | ⊢ False 29 | Remaining goals after safe rules: 30 | case a.a.a.a.a.a.a.a.a.a.a.a.a.a.a.a.a.a.a.a.a.a.a.a.a.a.a.a.a.a.a.a.a.a.a.a.a.a.a.a.a.a.a.a.a.a.a.a.a.a.a.a.a.a.a.a.a.a.a.a.a.a.a.a.a.a.a.a.a.a.a.a.a.a.a.a.a.a.a.a 31 | ⊢ False 32 | The safe prefix was not fully expanded because the maximum number of rule applications (50) was reached. 33 | -/ 34 | #guard_msgs in 35 | example : False := by 36 | aesop (config := { terminal := true }) 37 | -------------------------------------------------------------------------------- /AesopTest/SafePrefixInTerminalError.lean: -------------------------------------------------------------------------------- 1 | import Aesop 2 | 3 | /-- 4 | error: tactic 'aesop' failed, failed to prove the goal after exhaustive search. 5 | Initial goal: 6 | ⊢ ∀ (a b : Nat), a + b = b + 2 * a 7 | Remaining goals after safe rules: 8 | a b : Nat 9 | ⊢ a + b = b + 2 * a 10 | -/ 11 | #guard_msgs in 12 | example : ∀ (a b : Nat), a + b = b + 2 * a := by 13 | aesop (config := { terminal := true }) 14 | -------------------------------------------------------------------------------- /AesopTest/ScriptWithOptions.lean: -------------------------------------------------------------------------------- 1 | /- 2 | Copyright (c) 2023 Jannis Limperg. All rights reserved. 3 | Released under Apache 2.0 license as described in the file LICENSE. 4 | Authors: Jannis Limperg 5 | -/ 6 | 7 | import Aesop 8 | 9 | set_option aesop.check.all true 10 | 11 | /-- 12 | info: Try this: 13 | simp_all (config := { }) only 14 | -/ 15 | #guard_msgs in 16 | example : True := by 17 | aesop? (config := {}) (simp_config := {}) 18 | -------------------------------------------------------------------------------- /AesopTest/SimpLetHypotheses.lean: -------------------------------------------------------------------------------- 1 | /- 2 | Copyright (c) 2024 Jannis Limperg. All rights reserved. 3 | Released under Apache 2.0 license as described in the file LICENSE. 4 | Authors: Jannis Limperg 5 | -/ 6 | 7 | import Aesop 8 | 9 | set_option aesop.check.all true 10 | 11 | -- Aesop `simp` substitutes `let` hypotheses even when `zetaDelta` is disabled 12 | -- (which is the default). 13 | 14 | /-- 15 | error: unsolved goals 16 | P : Nat → Type 17 | h : P 0 18 | x : Nat := 0 19 | ⊢ P 0 20 | -/ 21 | #guard_msgs in 22 | example (P : Nat → Type) (h : P 0) : let x := 0; P x := by 23 | intro x 24 | aesop (rule_sets := [-builtin,-default]) 25 | (config := { warnOnNonterminal := false }) 26 | 27 | /-- 28 | error: unsolved goals 29 | P : Nat → Type 30 | h : P 0 31 | x : Nat := 0 32 | ⊢ P 0 33 | -/ 34 | #guard_msgs in 35 | example (P : Nat → Type) (h : P 0) : let x := 0; P x := by 36 | intro x 37 | aesop (rule_sets := [-builtin,-default]) 38 | (config := { useSimpAll := false, warnOnNonterminal := false }) 39 | -------------------------------------------------------------------------------- /AesopTest/Simprocs.lean: -------------------------------------------------------------------------------- 1 | /- 2 | Copyright (c) 2024 Jannis Limperg. All rights reserved. 3 | Released under Apache 2.0 license as described in the file LICENSE. 4 | Authors: Jannis Limperg 5 | -/ 6 | 7 | import Aesop 8 | 9 | open Lean Lean.Meta 10 | 11 | set_option aesop.check.all true 12 | set_option aesop.smallErrorMessages true 13 | 14 | def unfoldConst («from» to : Name) : Simp.Simproc := λ e => 15 | if e.isConstOf «from» then 16 | return .done { expr := .const to [] } 17 | else 18 | return .continue 19 | 20 | @[irreducible] def T₁ := True 21 | 22 | /-- 23 | error: tactic 'aesop' failed, made no progress 24 | -/ 25 | #guard_msgs in 26 | example : T₁ := by 27 | aesop 28 | 29 | simproc unfoldT₁ (T₁) := unfoldConst ``T₁ ``True 30 | 31 | example : T₁ := by 32 | aesop 33 | 34 | /-- 35 | error: tactic 'aesop' failed, made no progress 36 | -/ 37 | #guard_msgs in 38 | example : T₁ := by 39 | aesop (config := { useDefaultSimpSet := false }) 40 | 41 | @[irreducible] def T₂ := True 42 | 43 | /-- 44 | error: tactic 'aesop' failed, made no progress 45 | -/ 46 | #guard_msgs in 47 | example : T₂ := by 48 | aesop 49 | 50 | simproc [aesop_builtin] unfoldT₂ (T₂) := unfoldConst ``T₂ ``True 51 | 52 | /-- 53 | error: tactic 'aesop' failed, made no progress 54 | -/ 55 | #guard_msgs in 56 | example : T₂ := by 57 | aesop (rule_sets := [-builtin]) 58 | 59 | example : T₂ := by 60 | aesop 61 | -------------------------------------------------------------------------------- /AesopTest/Split.lean: -------------------------------------------------------------------------------- 1 | /- 2 | Copyright (c) 2022 Jannis Limperg. All rights reserved. 3 | Released under Apache 2.0 license as described in the file LICENSE. 4 | Authors: Jannis Limperg 5 | -/ 6 | 7 | import Aesop 8 | 9 | set_option aesop.check.all true 10 | 11 | -- Tests the builtin split rules. 12 | 13 | @[aesop norm unfold] 14 | def id' (n : Nat) : Nat := 15 | match n with 16 | | .zero => .zero 17 | | .succ n => .succ n 18 | 19 | 20 | -- First, the rule for splitting targets. 21 | example : id' n = n := by 22 | aesop 23 | 24 | -- Second, the rule for splitting hypotheses. We introduce a custom equality to 25 | -- make sure that in the following example, Aesop doesn't just substitute the 26 | -- hypothesis during norm simp. 27 | inductive MyEq (x : α) : α → Prop 28 | | rfl : MyEq x x 29 | 30 | @[aesop unsafe] 31 | def MyEq.elim : MyEq x y → x = y 32 | | rfl => by rfl 33 | 34 | example (h : MyEq n (id' m)) : n = m := by 35 | aesop 36 | -------------------------------------------------------------------------------- /AesopTest/SplitScript.lean: -------------------------------------------------------------------------------- 1 | /- 2 | Copyright (c) 2024 Jannis Limperg. All rights reserved. 3 | Released under Apache 2.0 license as described in the file LICENSE. 4 | Authors: Jannis Limperg 5 | -/ 6 | 7 | import Aesop 8 | 9 | set_option aesop.check.all true 10 | 11 | open Classical 12 | 13 | inductive MyFalse : Prop 14 | 15 | /-- 16 | info: Try this: 17 | split 18 | next h => sorry 19 | next h => sorry 20 | --- 21 | warning: declaration uses 'sorry' 22 | -/ 23 | #guard_msgs in 24 | example {A B : Prop} : if P then A else B := by 25 | aesop? (config := { warnOnNonterminal := false }) 26 | all_goals sorry 27 | 28 | /-- 29 | info: Try this: 30 | split at h 31 | next h_1 => simp_all only [true_or] 32 | next h_1 => simp_all only [or_true] 33 | -/ 34 | #guard_msgs in 35 | example (h : if P then A else B) : A ∨ B := by 36 | aesop? 37 | 38 | /-- 39 | info: Try this: 40 | split at h 41 | next n => simp_all only [true_or] 42 | next n => simp_all only [true_or, or_true] 43 | next n_1 x x_1 => simp_all only [imp_false, or_true] 44 | -/ 45 | #guard_msgs in 46 | theorem foo (n : Nat) (h : match n with | 0 => A | 1 => B | _ => C) : 47 | A ∨ B ∨ C := by 48 | set_option aesop.check.rules false in -- TODO simp introduces mvar 49 | set_option aesop.check.tree false in 50 | aesop? 51 | -------------------------------------------------------------------------------- /AesopTest/Stats.lean: -------------------------------------------------------------------------------- 1 | /- 2 | Copyright (c) 2025 Jannis Limperg. All rights reserved. 3 | Released under Apache 2.0 license as described in the file LICENSE. 4 | Authors: Jannis Limperg 5 | -/ 6 | 7 | import Aesop 8 | 9 | /- 10 | Regression test for a bug involving the stats extension and async elaboration. 11 | The commands below should not panic and should show non-zero stats. 12 | -/ 13 | 14 | set_option aesop.collectStats true 15 | set_option trace.aesop.stats true 16 | 17 | #guard_msgs (drop trace) in 18 | theorem mem_spec {o : Option α} : a ∈ o ↔ o = some a := by 19 | aesop (add norm simp Membership.mem) 20 | 21 | #guard_msgs (drop info) in 22 | #aesop_stats 23 | -------------------------------------------------------------------------------- /AesopTest/Strategy.lean: -------------------------------------------------------------------------------- 1 | /- 2 | Copyright (c) 2022 Jannis Limperg. All rights reserved. 3 | Released under Apache 2.0 license as described in the file LICENSE. 4 | Authors: Jannis Limperg 5 | -/ 6 | 7 | import Aesop 8 | 9 | set_option aesop.check.all true 10 | set_option aesop.smallErrorMessages true 11 | 12 | @[aesop 50% constructors] 13 | inductive I₁ 14 | | ofI₁ : I₁ → I₁ 15 | | ofTrue : True → I₁ 16 | 17 | example : I₁ := by 18 | aesop 19 | 20 | example : I₁ := by 21 | aesop (config := { strategy := .bestFirst }) 22 | 23 | example : I₁ := by 24 | aesop (config := { strategy := .breadthFirst }) 25 | 26 | /-- 27 | error: tactic 'aesop' failed, maximum number of rule applications (10) reached. Set the 'maxRuleApplications' option to increase the limit. 28 | -/ 29 | #guard_msgs in 30 | example : I₁ := by 31 | aesop (config := 32 | { strategy := .depthFirst 33 | maxRuleApplicationDepth := 0 34 | maxRuleApplications := 10, 35 | terminal := true }) 36 | 37 | example : I₁ := by 38 | aesop (config := { strategy := .depthFirst, maxRuleApplicationDepth := 10 }) 39 | -------------------------------------------------------------------------------- /AesopTest/Subst.lean: -------------------------------------------------------------------------------- 1 | /- 2 | Copyright (c) 2022 Jannis Limperg. All rights reserved. 3 | Released under Apache 2.0 license as described in the file LICENSE. 4 | Authors: Jannis Limperg 5 | -/ 6 | 7 | import Aesop 8 | 9 | set_option aesop.check.all true 10 | set_option aesop.smallErrorMessages true 11 | 12 | -- These test cases test the builtin subst tactic. 13 | 14 | /-- 15 | error: tactic 'aesop' failed, failed to prove the goal after exhaustive search. 16 | -/ 17 | #guard_msgs in 18 | example (h₁ : x = 5) (h₂ : y = 5) : x = y := by 19 | aesop (erase Aesop.BuiltinRules.subst) 20 | (config := { useSimpAll := false, terminal := true }) 21 | 22 | example (h₁ : x = 5) (h₂ : y = 5) : x = y := by 23 | aesop (config := { useSimpAll := false }) 24 | 25 | variable {α : Type} 26 | 27 | /-- 28 | error: tactic 'aesop' failed, failed to prove the goal after exhaustive search. 29 | -/ 30 | #guard_msgs in 31 | example {x y z : α} (h₁ : x = y) (h₂ : y = z) : x = z := by 32 | aesop (erase Aesop.BuiltinRules.subst) 33 | (config := { useSimpAll := false, terminal := true }) 34 | 35 | example {x y z : α} (h₁ : x = y) (h₂ : y = z) : x = z := by 36 | aesop (config := { useSimpAll := false }) 37 | 38 | /-- 39 | error: tactic 'aesop' failed, failed to prove the goal after exhaustive search. 40 | -/ 41 | #guard_msgs in 42 | example {x y : α }(P : ∀ x y, x = y → Prop) (h₁ : x = y) (h₂ : P x y h₁) : 43 | x = y := by 44 | aesop 45 | (erase Aesop.BuiltinRules.subst, Aesop.BuiltinRules.assumption, 46 | Aesop.BuiltinRules.applyHyps) 47 | (config := { useSimpAll := false, terminal := true }) 48 | 49 | example {x y : α }(P : ∀ x y, x = y → Prop) (h₁ : x = y) (h₂ : P x y h₁) : 50 | x = y := by 51 | aesop 52 | (erase Aesop.BuiltinRules.assumption, 53 | Aesop.BuiltinRules.applyHyps) 54 | (config := { useSimpAll := false }) 55 | 56 | -- Subst also works for bi-implications. 57 | 58 | /-- 59 | error: tactic 'aesop' failed, failed to prove the goal after exhaustive search. 60 | -/ 61 | #guard_msgs in 62 | example (h₁ : P ↔ Q) (h₂ : Q ↔ R) (h₃ : P) : R := by 63 | aesop (erase Aesop.BuiltinRules.subst) 64 | (config := { useSimpAll := false, terminal := true }) 65 | 66 | example (h₁ : P ↔ Q) (h₂ : Q ↔ R) (h₃ : P) : R := by 67 | aesop (config := { useSimpAll := false }) 68 | 69 | -- Subst also works for morally-homogeneous heterogeneous equalities (using a 70 | -- builtin simp rule which turns these into actual homogeneous equalities). 71 | 72 | /-- 73 | error: tactic 'aesop' failed, failed to prove the goal after exhaustive search. 74 | -/ 75 | #guard_msgs in 76 | example {P : α → Prop} {x y z : α} (h₁ : HEq x y) (h₂ : HEq y z) (h₃ : P x) : 77 | P z := by 78 | aesop (erase Aesop.BuiltinRules.subst) 79 | (config := { useSimpAll := false, terminal := true }) 80 | 81 | example {P : α → Prop} {x y z : α} (h₁ : HEq x y) (h₂ : HEq y z) (h₃ : P x) : 82 | P z := by 83 | aesop (config := { useSimpAll := false }) 84 | -------------------------------------------------------------------------------- /AesopTest/TacGen.lean: -------------------------------------------------------------------------------- 1 | import Aesop 2 | 3 | open Lean 4 | open Lean.Meta 5 | 6 | theorem foo (a b c : Nat) : a + b + c = c + b + a := by 7 | rw [Nat.add_assoc, Nat.add_comm, Nat.add_comm b] 8 | 9 | @[aesop 100%] 10 | def tacGen₁ : Aesop.TacGen := λ _ => do 11 | return #[("apply foo", 1.0)] 12 | 13 | example (a b c : Nat) : a + b + c = c + b + a := by 14 | aesop 15 | 16 | @[aesop 100%] 17 | def tacGen₂ : Aesop.TacGen := λ _ => 18 | return #[ 19 | ("rw [Nat.add_comm b]", 0.5), 20 | ("rw [Nat.add_assoc]", 0.9), 21 | ("rw [Nat.add_comm]", 0.8) 22 | ] 23 | 24 | example (a b c : Nat) : a + b + c = c + b + a := by 25 | aesop (erase tacGen₁) 26 | -------------------------------------------------------------------------------- /AesopTest/TacticConfig.lean: -------------------------------------------------------------------------------- 1 | /- 2 | Copyright (c) 2021 Jannis Limperg. All rights reserved. 3 | Released under Apache 2.0 license as described in the file LICENSE. 4 | Authors: Jannis Limperg 5 | -/ 6 | 7 | import Aesop 8 | 9 | set_option aesop.check.all true 10 | set_option aesop.smallErrorMessages true 11 | 12 | inductive Even : Nat → Prop 13 | | zero : Even 0 14 | | plus_two {n} : Even n → Even (n + 2) 15 | 16 | inductive Odd : Nat → Prop 17 | | one : Odd 1 18 | | plus_two {n} : Odd n → Odd (n + 2) 19 | 20 | inductive EvenOrOdd : Nat → Prop 21 | | even {n} : Even n → EvenOrOdd n 22 | | odd {n} : Odd n → EvenOrOdd n 23 | 24 | -- We can add constants as rules. 25 | example : EvenOrOdd 3 := by 26 | aesop 27 | (add safe [Even.zero, Even.plus_two, Odd.one, Odd.plus_two], 28 | unsafe [apply 50% EvenOrOdd.even, 50% EvenOrOdd.odd]) 29 | 30 | -- Same with local hypotheses, or a mix. 31 | example : EvenOrOdd 3 := by 32 | have h : ∀ n, Odd n → EvenOrOdd n := λ _ p => EvenOrOdd.odd p 33 | aesop 34 | (add safe [Odd.one, Odd.plus_two], unsafe [EvenOrOdd.even 50%, h 50%]) 35 | (erase Aesop.BuiltinRules.applyHyps) -- This rule subsumes h. 36 | 37 | attribute [aesop 50%] Even.zero Even.plus_two 38 | 39 | -- We can also erase global rules... 40 | 41 | /-- 42 | error: tactic 'aesop' failed, failed to prove the goal after exhaustive search. 43 | -/ 44 | #guard_msgs in 45 | example : EvenOrOdd 2 := by 46 | aesop (add safe EvenOrOdd.even) (erase Even.zero) 47 | (config := { terminal := true }) 48 | 49 | example : EvenOrOdd 2 := by 50 | aesop (add safe EvenOrOdd.even) 51 | 52 | -- ... as well as local ones (but what for?). 53 | 54 | 55 | /-- 56 | error: tactic 'aesop' failed, failed to prove the goal after exhaustive search. 57 | -/ 58 | #guard_msgs in 59 | example : EvenOrOdd 2 := by 60 | have h : ∀ n, Even n → EvenOrOdd n := λ _ p => EvenOrOdd.even p 61 | aesop (add safe h) (erase Aesop.BuiltinRules.applyHyps, h) 62 | (config := { terminal := true }) 63 | 64 | example : EvenOrOdd 2 := by 65 | have h : ∀ n, Even n → EvenOrOdd n := λ _ p => EvenOrOdd.even p 66 | aesop (add safe h) (erase Aesop.BuiltinRules.applyHyps) 67 | -------------------------------------------------------------------------------- /AesopTest/Tauto.lean: -------------------------------------------------------------------------------- 1 | /- 2 | Copyright (c) 2022 Jannis Limperg. All rights reserved. 3 | Released under Apache 2.0 license as described in the file LICENSE. 4 | Authors: David Renshaw, Jannis Limperg 5 | -/ 6 | 7 | import Aesop 8 | 9 | set_option aesop.check.all true 10 | 11 | -- This is an example which is currently challenging for Lean 4 `tauto`. 12 | example {α : Type} [LE α] (a b c : α) (x₀ x₁ x₂ : Prop) 13 | (this1 : x₀ → x₁ → a ≤ c) 14 | (this2 : x₁ → x₂ → b ≤ a) 15 | (this3 : x₂ → x₀ → c ≤ b) : 16 | ((x₀ ∧ ¬b ≤ a) ∧ x₁ ∧ ¬c ≤ b ∨ 17 | (x₁ ∧ ¬c ≤ b) ∧ x₂ ∧ ¬a ≤ c ∨ (x₂ ∧ ¬a ≤ c) ∧ x₀ ∧ ¬b ≤ a ↔ 18 | (x₀ ∧ x₁ ∨ x₁ ∧ x₂ ∨ x₂ ∧ x₀) ∧ 19 | ¬(c ≤ b ∧ b ≤ a ∨ b ≤ a ∧ a ≤ c ∨ a ≤ c ∧ c ≤ b)) := 20 | by aesop 21 | -------------------------------------------------------------------------------- /AesopTest/TerminalError.lean: -------------------------------------------------------------------------------- 1 | /- 2 | Copyright (c) 2023 Jannis Limperg. All rights reserved. 3 | Released under Apache 2.0 license as described in the file LICENSE. 4 | Authors: Jannis Limperg 5 | -/ 6 | 7 | import Aesop 8 | 9 | set_option aesop.check.all true 10 | 11 | /-- 12 | error: tactic 'aesop' failed, failed to prove the goal after exhaustive search. 13 | Initial goal: 14 | ⊢ False 15 | Remaining goals after safe rules: 16 | ⊢ False 17 | -/ 18 | #guard_msgs in 19 | example : False := by 20 | aesop (config := { terminal := true }) 21 | -------------------------------------------------------------------------------- /AesopTest/TraceProof.lean: -------------------------------------------------------------------------------- 1 | /- 2 | Copyright (c) 2022 Jannis Limperg. All rights reserved. 3 | Released under Apache 2.0 license as described in the file LICENSE. 4 | Authors: Jannis Limperg 5 | -/ 6 | 7 | import Aesop 8 | 9 | set_option aesop.check.all true 10 | set_option aesop.smallErrorMessages true 11 | 12 | set_option trace.aesop.proof true 13 | 14 | /-- 15 | error: tactic 'aesop' failed, made no progress 16 | --- 17 | trace: [aesop.proof] 18 | -/ 19 | #guard_msgs in 20 | example : α := by 21 | aesop 22 | 23 | @[aesop norm simp] 24 | def F := False 25 | 26 | set_option pp.mvars false in 27 | /-- 28 | warning: aesop: failed to prove the goal after exhaustive search. 29 | --- 30 | error: unsolved goals 31 | ⊢ False 32 | --- 33 | trace: [aesop.proof] id ?_ 34 | -/ 35 | #guard_msgs in 36 | example : F := by 37 | aesop 38 | -------------------------------------------------------------------------------- /AesopTest/TryThisIndentation.lean: -------------------------------------------------------------------------------- 1 | /- 2 | Copyright (c) 2023 Jannis Limperg. All rights reserved. 3 | Released under Apache 2.0 license as described in the file LICENSE. 4 | Authors: Jannis Limperg 5 | -/ 6 | 7 | import Aesop 8 | 9 | /-- 10 | info: Try this: 11 | intro a_1 a_2 12 | simp_all only [ne_eq, List.mem_cons, or_self, not_false_eq_true] 13 | -/ 14 | #guard_msgs in 15 | example {a y : α} {l : List α} : a ≠ y → a ∉ l → a ∉ y::l := by 16 | aesop? 17 | 18 | /-- 19 | info: Try this: 20 | simp_all only [ne_eq, List.mem_cons, or_self, not_false_eq_true] 21 | -/ 22 | #guard_msgs in 23 | example {a y : α} {l : List α} : a ≠ y → a ∉ l → a ∉ y::l := by 24 | intros 25 | have : ¬ a ∈ y :: l := by 26 | aesop? 27 | exact this 28 | -------------------------------------------------------------------------------- /AesopTest/Unfold.lean: -------------------------------------------------------------------------------- 1 | /- 2 | Copyright (c) 2023 Jannis Limperg. All rights reserved. 3 | Released under Apache 2.0 license as described in the file LICENSE. 4 | Authors: Jannis Limperg 5 | -/ 6 | 7 | -- Inspired by this Zulip discussion: 8 | -- https://leanprover.zulipchat.com/#narrow/stream/287929-mathlib4/topic/Goal.20state.20not.20updating.2C.20bugs.2C.20etc.2E/near/338356062 9 | 10 | import Aesop 11 | 12 | set_option aesop.check.all true 13 | 14 | structure WalkingPair 15 | 16 | def WidePullbackShape A B := Sum A B 17 | 18 | abbrev WalkingCospan : Type := WidePullbackShape Empty Empty 19 | 20 | @[aesop norm destruct] 21 | def WalkingCospan_elim : WalkingCospan → Sum Empty Empty := id 22 | 23 | /-- 24 | info: Try this: 25 | unfold WalkingCospan at h 26 | unfold WidePullbackShape at h 27 | cases h with 28 | | inl val => 29 | have fwd : False := Aesop.BuiltinRules.empty_false val 30 | simp_all only 31 | | inr val_1 => 32 | have fwd : False := Aesop.BuiltinRules.empty_false val_1 33 | simp_all only 34 | -/ 35 | #guard_msgs in 36 | example (h : WalkingCospan) : α := by 37 | aesop? (add norm unfold [WidePullbackShape, WalkingCospan]) 38 | 39 | example (h : WalkingCospan) : α := by 40 | aesop (add norm simp [WidePullbackShape, WalkingCospan]) 41 | 42 | @[aesop norm unfold] 43 | def Foo := True 44 | 45 | @[aesop norm unfold] 46 | def Bar := False 47 | 48 | /-- 49 | info: Try this: 50 | unfold Bar Foo 51 | unfold Foo at h₁ 52 | unfold Foo Bar at h₂ 53 | simp_all only 54 | cases h₂ with 55 | | inl val => 56 | simp_all only 57 | apply PSum.inr 58 | simp_all only 59 | | inr val_1 => simp_all only 60 | -/ 61 | #guard_msgs in 62 | example (h₁ : Foo) (h₂ : PSum Foo Bar) : PSum Bar Foo := by 63 | aesop? 64 | -------------------------------------------------------------------------------- /AesopTest/UnreachableTacticLinter.lean: -------------------------------------------------------------------------------- 1 | /- 2 | Copyright (c) 2024 Jannis Limperg. All rights reserved. 3 | Released under Apache 2.0 license as described in the file LICENSE. 4 | Authors: Jannis Limperg 5 | -/ 6 | 7 | import Aesop 8 | 9 | set_option aesop.check.all true 10 | 11 | add_aesop_rules safe (by simp) 12 | -------------------------------------------------------------------------------- /AesopTest/WarnApplyIff.lean: -------------------------------------------------------------------------------- 1 | /- 2 | Copyright (c) 2024 Jannis Limperg. All rights reserved. 3 | Released under Apache 2.0 license as described in the file LICENSE. 4 | Authors: Jannis Limperg 5 | -/ 6 | 7 | import Aesop 8 | 9 | set_option aesop.check.all true 10 | 11 | /-- 12 | warning: Apply builder was used for a theorem with conclusion A ↔ B. 13 | You probably want to use the simp builder or create an alias that applies the theorem in one direction. 14 | Use `set_option aesop.warn.applyIff false` to disable this warning. 15 | -/ 16 | #guard_msgs in 17 | @[aesop safe apply] 18 | axiom foo : True ↔ True 19 | 20 | @[aesop simp] 21 | axiom bar : True ↔ True 22 | 23 | set_option aesop.warn.applyIff false in 24 | @[aesop 1% apply] 25 | axiom baz : True ↔ True 26 | -------------------------------------------------------------------------------- /lake-manifest.json: -------------------------------------------------------------------------------- 1 | {"version": "1.1.0", 2 | "packagesDir": ".lake/packages", 3 | "packages": 4 | [{"url": "https://github.com/leanprover-community/batteries", 5 | "type": "git", 6 | "subDir": null, 7 | "scope": "", 8 | "rev": "08681ddeb7536a50dea8026c6693cb9b07f01717", 9 | "name": "batteries", 10 | "manifestFile": "lake-manifest.json", 11 | "inputRev": "v4.21.0-rc3", 12 | "inherited": false, 13 | "configFile": "lakefile.toml"}], 14 | "name": "aesop", 15 | "lakeDir": ".lake"} 16 | -------------------------------------------------------------------------------- /lakefile.toml: -------------------------------------------------------------------------------- 1 | name = "aesop" 2 | defaultTargets = ["Aesop"] 3 | testDriver = "AesopTest" 4 | precompileModules = false # We would like to turn this on, but it breaks the Mathlib cache. 5 | platformIndependent = true 6 | 7 | [[require]] 8 | name = "batteries" 9 | git = "https://github.com/leanprover-community/batteries" 10 | rev = "v4.21.0-rc3" 11 | 12 | [[lean_lib]] 13 | name = "Aesop" 14 | 15 | [[lean_lib]] 16 | name = "AesopTest" 17 | globs = ["AesopTest.+"] 18 | leanOptions = {linter.unusedVariables = false} 19 | -------------------------------------------------------------------------------- /lean-toolchain: -------------------------------------------------------------------------------- 1 | leanprover/lean4:v4.21.0-rc3 2 | --------------------------------------------------------------------------------