├── lean-toolchain ├── Lean ├── Egg │ ├── Tests │ │ ├── Math │ │ │ ├── lean-toolchain │ │ │ ├── Math │ │ │ │ ├── Comparisons │ │ │ │ │ ├── Simp.lean │ │ │ │ │ └── Rotman Basic.lean │ │ │ │ ├── Figures │ │ │ │ │ ├── List.Perm.sum_eq.lean │ │ │ │ │ ├── Lie Algebra.lean │ │ │ │ │ └── Boolean Algebra.lean │ │ │ │ ├── Zulip │ │ │ │ │ ├── AC.lean │ │ │ │ │ └── Binders.lean │ │ │ │ ├── PushNeg.lean │ │ │ │ ├── Cast.lean │ │ │ │ ├── Mathlib │ │ │ │ │ └── Algebra │ │ │ │ │ │ └── Group │ │ │ │ │ │ └── Basic.lean │ │ │ │ ├── Andrés │ │ │ │ │ └── RingSupport.lean │ │ │ │ └── Rotman.lean │ │ │ ├── Math.lean │ │ │ └── lakefile.toml │ │ ├── Dataflow Rewriter │ │ │ ├── lean-toolchain │ │ │ ├── DataflowRewriter.lean │ │ │ ├── lakefile.toml │ │ │ └── DataflowRewriter │ │ │ │ └── One.lean │ │ ├── Rudi.lean │ │ ├── Weak Vars Bug.lean │ │ ├── EndOfInput 1.lean │ │ ├── EndOfInput 2.lean │ │ ├── Int.lean │ │ ├── Ambient Expr MVar.lean │ │ ├── Non-Recursive Def.lean │ │ ├── Recursive.lean │ │ ├── Polymorphic.lean │ │ ├── Explosion Failure.lean │ │ ├── Unexpected Bound Var.lean │ │ ├── Binder BVars.lean │ │ ├── TC Stuck.lean │ │ ├── Reconstruction Beta.lean │ │ ├── Zulip 2.lean │ │ ├── Condition Kind.lean │ │ ├── Ground Eqs.lean │ │ ├── Proof Irrelevance.lean │ │ ├── Erase TC Instances Bug 1.lean │ │ ├── Normalize.lean │ │ ├── Crash 2.lean │ │ ├── Stuck 1.lean │ │ ├── Ambient Level MVar.lean │ │ ├── Stuck 2.lean │ │ ├── Class Def.lean │ │ ├── Shapes DefEqs.lean │ │ ├── WIP AC.lean │ │ ├── Soundness.lean │ │ ├── Baskets.lean │ │ ├── Fact Kind.lean │ │ ├── Type Level Defeq Rws.lean │ │ ├── Stuck 3.lean │ │ ├── Constants.lean │ │ ├── Eta Conservative Intersection.lean │ │ ├── Binders.lean │ │ ├── Intro Conditional.lean │ │ ├── Intro Hygiene.lean │ │ ├── Thomas.lean │ │ ├── Zulip 1.lean │ │ ├── Boundless BVars.lean │ │ ├── Beta.lean │ │ ├── Subgoals.lean │ │ ├── Eta Conservative.lean │ │ ├── LMVar Encoding.lean │ │ ├── Readme.lean │ │ ├── BVar Index Correction.lean │ │ ├── StrLit.lean │ │ ├── Cond Valid MVar Dirs.lean │ │ ├── Hypotheses.lean │ │ ├── Reconstruction.lean │ │ ├── Crash 3.lean │ │ ├── Calc.lean │ │ ├── Nate.lean │ │ ├── Builtins.lean │ │ ├── Crash 1.lean │ │ ├── Shapes Rerun.lean │ │ ├── Crash 4.lean │ │ ├── Level Defeq.lean │ │ ├── batteries │ │ │ ├── SimpOnlyOverride.lean │ │ │ └── mk_batteries.sh │ │ ├── mathlib4 │ │ │ ├── SimpOnlyOverride.lean │ │ │ └── mk_mathlib.sh │ │ ├── Condition Subgoals.lean │ │ ├── Application Order.lean │ │ ├── Propositions.lean │ │ ├── Inst Erasure.lean │ │ ├── Block Invalid Matches.lean │ │ ├── Basic.lean │ │ ├── Cast.lean │ │ ├── Activations.lean │ │ ├── Reduce.lean │ │ ├── Blocking.lean │ │ ├── Tags.lean │ │ ├── Star.lean │ │ ├── Calcify.lean │ │ ├── Groups.lean │ │ ├── Equality Conditions.lean │ │ ├── Pattern MVar.lean │ │ ├── WIP ValidateTagged.lean │ │ ├── Shift Loop.lean │ │ ├── Rudi Groups.lean │ │ ├── Freshman Calc.lean │ │ ├── Derived Guides.lean │ │ ├── WIP.lean │ │ ├── TC Proj Binders.lean │ │ ├── EraseProofs.lean │ │ ├── DataflowRewriter.lean │ │ ├── Eta.lean │ │ ├── Prune.lean │ │ ├── Conditional.lean │ │ ├── Explosion.lean │ │ ├── Functional.lean │ │ ├── Freshman.lean │ │ ├── Freshman Tags.lean │ │ ├── NatLit.lean │ │ ├── Freshman Pow.lean │ │ ├── TestDriver.lean │ │ ├── MVars Properties.lean │ │ └── FunctionalTc.lean │ ├── Core │ │ ├── Request │ │ │ ├── EGraph.lean │ │ │ ├── Synth.lean │ │ │ ├── Term.lean │ │ │ └── Equiv.lean │ │ ├── Blocks.lean │ │ ├── Encode │ │ │ ├── Blocks.lean │ │ │ ├── Guides.lean │ │ │ ├── Shapes.lean │ │ │ ├── Rules.lean │ │ │ ├── EncodeM.lean │ │ │ └── Basic.lean │ │ ├── Guides.lean │ │ ├── Gen │ │ │ ├── Tagged.lean │ │ │ ├── Builtins.lean │ │ │ ├── Guides.lean │ │ │ ├── Explosion.lean │ │ │ ├── GoalTypeSpecialization.lean │ │ │ └── StructProjs.lean │ │ ├── Direction.lean │ │ ├── Explanation │ │ │ ├── Parse │ │ │ │ └── Basic.lean │ │ │ └── Basic.lean │ │ ├── Rewrite │ │ │ ├── Prerewrite.lean │ │ │ └── Condition.lean │ │ ├── Normalize.lean │ │ ├── Congr.lean │ │ ├── Config.lean │ │ └── MVars │ │ │ └── Subst.lean │ └── Tactic │ │ ├── Calcify.lean │ │ ├── Rewrite │ │ └── Rules │ │ │ └── Gen │ │ │ ├── Intros.lean │ │ │ └── Basic.lean │ │ ├── Blocks.lean │ │ ├── Guides.lean │ │ ├── Goal.lean │ │ └── Config │ │ └── Option.lean ├── gen_root_file.sh └── Egg.lean ├── Docs ├── icon.png └── Lean Together 2025.pdf ├── C ├── ffi.h └── rev_ffi.c ├── Rust ├── Egg │ ├── Cargo.toml │ ├── build.rs │ └── src │ │ ├── levels.rs │ │ ├── result.rs │ │ ├── beta.rs │ │ ├── expl.rs │ │ ├── util.rs │ │ └── eta.rs └── Slotted │ ├── Cargo.toml │ └── src │ ├── eta.rs │ ├── levels.rs │ ├── result.rs │ ├── beta.rs │ ├── analysis.rs │ ├── rewrite.rs │ └── basic.rs ├── .gitignore ├── lake-manifest.json ├── .github └── workflows │ ├── ci.yml │ └── release.yml └── lakefile.lean /lean-toolchain: -------------------------------------------------------------------------------- 1 | leanprover/lean4:v4.23.0 2 | -------------------------------------------------------------------------------- /Lean/Egg/Tests/Math/lean-toolchain: -------------------------------------------------------------------------------- 1 | leanprover/lean4:v4.23.0 -------------------------------------------------------------------------------- /Docs/icon.png: -------------------------------------------------------------------------------- https://raw.githubusercontent.com/marcusrossel/lean-egg/HEAD/Docs/icon.png -------------------------------------------------------------------------------- /Lean/Egg/Tests/Dataflow Rewriter/lean-toolchain: -------------------------------------------------------------------------------- 1 | leanprover/lean4:v4.20.0-rc5 2 | -------------------------------------------------------------------------------- /Docs/Lean Together 2025.pdf: -------------------------------------------------------------------------------- https://raw.githubusercontent.com/marcusrossel/lean-egg/HEAD/Docs/Lean Together 2025.pdf -------------------------------------------------------------------------------- /C/ffi.h: -------------------------------------------------------------------------------- 1 | typedef struct env { 2 | void* x1; 3 | void* x2; 4 | void* x3; 5 | void* x4; 6 | void* x5; 7 | } env; 8 | -------------------------------------------------------------------------------- /Lean/Egg/Tests/Rudi.lean: -------------------------------------------------------------------------------- 1 | import Egg 2 | 3 | example : (fun x => (fun t _y => t) (fun z => x z)) = (fun (x : α → α) (_y : α) => x) := by 4 | egg 5 | -------------------------------------------------------------------------------- /Lean/Egg/Tests/Weak Vars Bug.lean: -------------------------------------------------------------------------------- 1 | import Egg 2 | 3 | example {P : Prop} (h : ∀ x y : Nat, x = y → P) {a b : Nat} : a = b → P := by 4 | sorry -- egg [h] 5 | -------------------------------------------------------------------------------- /Lean/Egg/Tests/EndOfInput 1.lean: -------------------------------------------------------------------------------- 1 | import Egg 2 | 3 | -- This tests checks that `egg` can be the final syntax in a file without throwing the error 4 | -- "Unexpected end of input". 5 | 6 | example : 0 = 0 := by 7 | egg 8 | -------------------------------------------------------------------------------- /Lean/Egg/Tests/Dataflow Rewriter/DataflowRewriter.lean: -------------------------------------------------------------------------------- 1 | -- This module serves as the root of the `DataflowRewriter` library. 2 | -- Import modules here that should be built as part of the library. 3 | import DataflowRewriter.One 4 | -------------------------------------------------------------------------------- /Lean/Egg/Tests/EndOfInput 2.lean: -------------------------------------------------------------------------------- 1 | import Egg 2 | 3 | -- This tests checks that `egg` can be the final syntax in a file without throwing the error 4 | -- "Unexpected end of input". 5 | 6 | example : 0 = 0 := by 7 | egg (config := { }) 8 | -------------------------------------------------------------------------------- /Lean/Egg/Tests/Int.lean: -------------------------------------------------------------------------------- 1 | import Egg 2 | 3 | variable (a b c d : Int) 4 | 5 | example : ((a * b) - (2 * c)) * d - (a * b) = (d - 1) * (a * b) - (2 * c * d) := by 6 | egg [Int.sub_mul, Int.sub_sub, Int.add_comm, Int.mul_comm, Int.one_mul] 7 | -------------------------------------------------------------------------------- /Lean/Egg/Tests/Ambient Expr MVar.lean: -------------------------------------------------------------------------------- 1 | import Egg 2 | 3 | variable (h : ∀ x : Prop, (x ∧ x) = True) 4 | 5 | example : True = (True ∧ True) := by 6 | refine Eq.trans (b := ?b ∧ ?b) ?eq₁ ?eq₂ 7 | case eq₁ => egg [h] 8 | case eq₂ => exact rfl 9 | -------------------------------------------------------------------------------- /Lean/Egg/Tests/Non-Recursive Def.lean: -------------------------------------------------------------------------------- 1 | import Egg 2 | 3 | -- Tests for the equations generated for non-recursive functions. 4 | 5 | def f : Bool → Nat 6 | | false => 0 7 | | true => 1 8 | 9 | example : f false = 0 := by 10 | egg [f] 11 | -------------------------------------------------------------------------------- /Lean/Egg/Tests/Recursive.lean: -------------------------------------------------------------------------------- 1 | import Egg 2 | 3 | -- This case shows up in the proof of `List.zip_map'` in Batteries. 4 | 5 | /-- error: egg does not support using auxiliary declarations -/ 6 | #guard_msgs in 7 | theorem t : a = b := by 8 | egg [t] 9 | -------------------------------------------------------------------------------- /Lean/gen_root_file.sh: -------------------------------------------------------------------------------- 1 | #!/usr/bin/env bash 2 | 3 | lean_dir="$(realpath -s "$(dirname "$0")")" 4 | cd "$lean_dir" 5 | find Egg -type f | grep -v -e 'Tests' -e '.DS_Store' -e 'Egg/Lean.lean' | tr "/" "." | sed 's/^/import /' | sed 's/\.lean$//' > "Egg.lean" -------------------------------------------------------------------------------- /Rust/Egg/Cargo.toml: -------------------------------------------------------------------------------- 1 | [package] 2 | name = "egg-for-lean" 3 | version = "0.1.0" 4 | edition = "2021" 5 | 6 | [lib] 7 | crate-type = ["staticlib"] 8 | 9 | [dependencies] 10 | egg = "0.10" 11 | libc = "0.2" 12 | 13 | [build-dependencies] 14 | cc = "1.0" 15 | -------------------------------------------------------------------------------- /.gitignore: -------------------------------------------------------------------------------- 1 | /Lean/Egg/Tests/Dataflow\ Rewriter/.lake 2 | /Lean/Egg/Tests/batteries/batteries 3 | /Lean/Egg/Tests/mathlib4/mathlib4 4 | /Lean/Egg/Tests/Math/.lake 5 | /Rust/Egg/Cargo.lock 6 | /Rust/Egg/target 7 | /Rust/Slotted/Cargo.lock 8 | /Rust/Slotted/target 9 | /.lake -------------------------------------------------------------------------------- /Lean/Egg/Tests/Math/Math/Comparisons/Simp.lean: -------------------------------------------------------------------------------- 1 | import Lean 2 | 3 | initialize do pure () <* Lean.Meta.registerSimpAttr `bool_simp "" 4 | initialize do pure () <* Lean.Meta.registerSimpAttr `lie_simp "" 5 | initialize do pure () <* Lean.Meta.registerSimpAttr `rotman_simp "" 6 | -------------------------------------------------------------------------------- /Rust/Slotted/Cargo.toml: -------------------------------------------------------------------------------- 1 | [package] 2 | name = "slotted-for-lean" 3 | version = "0.1.0" 4 | edition = "2021" 5 | 6 | [lib] 7 | crate-type = ["staticlib"] 8 | 9 | [dependencies] 10 | slotted-egraphs = { version = "=0.0.26", features = ["explanations"] } 11 | libc = "0.2" -------------------------------------------------------------------------------- /Lean/Egg/Tests/Math/Math/Figures/List.Perm.sum_eq.lean: -------------------------------------------------------------------------------- 1 | import Mathlib 2 | import Egg 3 | 4 | set_option pp.explicit true 5 | set_option trace.egg true 6 | 7 | example {M} [AddCommMonoid M] {l₁ l₂ : List M} (h : l₁.Perm l₂) : l₁.sum = l₂.sum := by 8 | egg [List.Perm.sum_eq, *] 9 | -------------------------------------------------------------------------------- /Lean/Egg/Tests/Math/Math.lean: -------------------------------------------------------------------------------- 1 | import Math.Mathlib.Algebra.Group.Basic 2 | import Math.Mathlib.Algebra.Group.Defs 3 | import Math.Mathlib.Algebra.Lie.Basic 4 | import Math.Mathlib.Order.BooleanAlgebra 5 | import Math.Zulip.AC 6 | import Math.Cast 7 | import Math.PushNeg 8 | import Math.Rotman 9 | -------------------------------------------------------------------------------- /Lean/Egg/Tests/Polymorphic.lean: -------------------------------------------------------------------------------- 1 | import Egg 2 | 3 | -- Tests involving polymorphic types. 4 | 5 | example : ([] : List α) = [] := by 6 | egg 7 | 8 | example {l₁ l₂ : List α} : l₁ ++ l₂ = (l₂.reverse ++ l₁.reverse).reverse := by 9 | egg [List.reverse_append, List.reverse_reverse] 10 | -------------------------------------------------------------------------------- /Lean/Egg/Core/Request/EGraph.lean: -------------------------------------------------------------------------------- 1 | namespace Egg 2 | 3 | private opaque EGraph.Pointed : NonemptyType.{0} 4 | 5 | def EGraph.Obj := EGraph.Pointed.type 6 | 7 | instance : Nonempty EGraph.Obj := EGraph.Pointed.property 8 | 9 | structure EGraph where 10 | obj : EGraph.Obj 11 | slotted : Bool 12 | -------------------------------------------------------------------------------- /Lean/Egg/Tests/Explosion Failure.lean: -------------------------------------------------------------------------------- 1 | import Egg 2 | 3 | -- This used to fail because explosion generation only generated explosions with up to one 4 | -- instantiation (because of a bug). 5 | set_option egg.explosion true 6 | 7 | example [Mul G] (h : ∀ x y z w : G, x * y = (z * w) * w) : ∀ x y z w : G, x * y = z * w := by 8 | egg [h] 9 | -------------------------------------------------------------------------------- /Lean/Egg/Tests/Math/lakefile.toml: -------------------------------------------------------------------------------- 1 | name = "Math" 2 | version = "0.1.0" 3 | keywords = ["math"] 4 | defaultTargets = ["Math"] 5 | 6 | [[require]] 7 | name = "mathlib" 8 | scope = "leanprover-community" 9 | rev = "v4.23.0" 10 | 11 | [[require]] 12 | name = "egg" 13 | path = "../../../.." 14 | 15 | [[lean_lib]] 16 | name = "Math" 17 | -------------------------------------------------------------------------------- /Lean/Egg/Tests/Unexpected Bound Var.lean: -------------------------------------------------------------------------------- 1 | import Egg 2 | 3 | -- This tests that we correctly encode bound variables appearing in binder domains. This previously 4 | -- failed with `unexpected bound variable #0`. 5 | 6 | variable (h : True ↔ ∀ (a : Nat) (_ : a > 0), True) 7 | 8 | example : True ↔ ∀ (a : Nat) (_ : a > 0), True := by 9 | egg [h] 10 | -------------------------------------------------------------------------------- /Lean/Egg/Tests/Binder BVars.lean: -------------------------------------------------------------------------------- 1 | import Egg 2 | 3 | -- This test checks that we correctly encode bvars in binder types. We had a bug where the binder 4 | -- type would have an off-by-one bvar index. 5 | 6 | example : 7 | (∀ (α : Type) (l : List α), l.length = l.length) ↔ 8 | (∀ (α : Type) (l : List α), l.length = l.length) := by 9 | egg 10 | -------------------------------------------------------------------------------- /Lean/Egg/Tests/TC Stuck.lean: -------------------------------------------------------------------------------- 1 | import Egg 2 | 3 | -- This test checks that equations of definitions involving type classes are correctly elaborated. 4 | -- This previously failed with `typeclass instance problem is stuck, it is often due to metavariables`. 5 | 6 | def f [Inhabited α] : α := Inhabited.default 7 | 8 | example : 0 = 0 := by 9 | egg [f] 10 | -------------------------------------------------------------------------------- /Lean/Egg/Tests/Math/Math/Zulip/AC.lean: -------------------------------------------------------------------------------- 1 | import Mathlib.Data.Real.Basic 2 | import Egg 3 | 4 | -- https://leanprover.zulipchat.com/#narrow/channel/113488-general/topic/.22Missing.20Tactics.22.20list/near/506705609 5 | 6 | attribute [egg ac] add_assoc add_comm 7 | 8 | example {x y z a : ℝ} (h : x + y + z = 0) : z + (a + x) + y = a := by 9 | egg +ac [h, zero_add] 10 | -------------------------------------------------------------------------------- /Lean/Egg/Core/Blocks.lean: -------------------------------------------------------------------------------- 1 | import Egg.Core.Source 2 | import Egg.Core.Normalize 3 | 4 | open Lean Meta 5 | 6 | namespace Egg 7 | 8 | abbrev Block := Expr 9 | 10 | def Block.from? (expr : Expr) (cfg : Config.Normalization) : MetaM (Option Block) := do 11 | unless (← inferType expr).isProp do return none 12 | normalize expr cfg 13 | 14 | abbrev Blocks := Array Block 15 | -------------------------------------------------------------------------------- /Lean/Egg/Tests/Reconstruction Beta.lean: -------------------------------------------------------------------------------- 1 | import Egg 2 | 3 | -- This once started causing issues when we moved to small-step β-reduction (perhaps because then 4 | -- β-reduction actually started working). Turning off β-reduction fixed the problem in that case. 5 | -- Now it just seems to work. 6 | example (arr : Array α) (i : Nat) (h₁ h₂ : i < arr.size) : arr[i]'h₁ = arr[i]'h₂ := by 7 | egg 8 | -------------------------------------------------------------------------------- /Lean/Egg/Tests/Zulip 2.lean: -------------------------------------------------------------------------------- 1 | import Egg 2 | 3 | -- https://leanprover.zulipchat.com/#narrow/channel/113489-new-members/topic/completing.20the.20proof.20somewhat.20automatically/with/520934553 4 | 5 | variable {A B C D E F G : Prop} 6 | variable (a : A → B → D) (b : C → A → F) (c_1 : F → A → G) (c : D → G → F → E) 7 | 8 | example (ho : A) (h1 : B) (h2 : C) : E := by 9 | egg [*] 10 | -------------------------------------------------------------------------------- /Lean/Egg/Core/Encode/Blocks.lean: -------------------------------------------------------------------------------- 1 | import Egg.Core.Blocks 2 | import Egg.Core.Encode.Basic 3 | import Lean 4 | open Lean 5 | 6 | namespace Egg 7 | 8 | abbrev Block.Encoded := Expression 9 | 10 | abbrev Blocks.Encoded := Array Block.Encoded 11 | 12 | def Blocks.encode (blocks : Blocks) (cfg : Config.Encoding) : MetaM Blocks.Encoded := 13 | blocks.mapM fun block => Egg.encode block cfg 14 | -------------------------------------------------------------------------------- /Lean/Egg/Core/Guides.lean: -------------------------------------------------------------------------------- 1 | import Egg.Core.Source 2 | import Egg.Core.Normalize 3 | import Lean 4 | open Lean 5 | 6 | namespace Egg 7 | 8 | structure Guide where 9 | private mk :: 10 | expr : Expr 11 | src : Source 12 | 13 | def Guide.from (expr : Expr) (src : Source) : MetaM Guide := 14 | return { expr := ← normalize expr .noReduce, src } 15 | 16 | abbrev Guides := Array Guide 17 | -------------------------------------------------------------------------------- /Lean/Egg/Tests/Condition Kind.lean: -------------------------------------------------------------------------------- 1 | import Egg 2 | 3 | -- This test should fail, because both directions of `h` are invalid, because the `Nat` argument is 4 | -- not covered, because it can't be a condition, because it's neither a proof nor an instance. 5 | 6 | /-- error: egg failed to prove the goal (saturated) -/ 7 | #guard_msgs in 8 | example (h : Nat → 0 = 1) : 0 = 1 := by 9 | egg [h] 10 | -------------------------------------------------------------------------------- /Lean/Egg/Core/Encode/Guides.lean: -------------------------------------------------------------------------------- 1 | import Egg.Core.Guides 2 | import Egg.Core.Encode.Basic 3 | import Lean 4 | open Lean 5 | 6 | namespace Egg 7 | 8 | abbrev Guide.Encoded := Expression 9 | 10 | abbrev Guides.Encoded := Array Guide.Encoded 11 | 12 | def Guides.encode (guides : Guides) (cfg : Config.Encoding) : MetaM Guides.Encoded := 13 | guides.mapM fun guide => Egg.encode guide.expr cfg 14 | -------------------------------------------------------------------------------- /Lean/Egg/Core/Gen/Tagged.lean: -------------------------------------------------------------------------------- 1 | import Egg.Tactic.Rewrite.Rules.Parse 2 | import Egg.Tactic.Baskets 3 | import Lean 4 | open Lean Elab Tactic 5 | 6 | namespace Egg 7 | 8 | def genTagged (cfg : Config) : TacticM Rewrite.Rules := do 9 | let mut prems := #[] 10 | for basket in cfg.baskets do 11 | prems := prems ++ (← extension.getAllBasketPremises basket) 12 | Rewrite.Rules.elabTagged prems cfg 13 | -------------------------------------------------------------------------------- /Lean/Egg/Tests/Ground Eqs.lean: -------------------------------------------------------------------------------- 1 | import Egg 2 | 3 | /-- error: egg failed to prove the goal (saturated) -/ 4 | #guard_msgs in 5 | set_option egg.groundEqs false in 6 | example (h₁ : ∀ p, p ∧ p) (h₂ : (∀ p, p ∧ p) → q = True) : q = True := by 7 | egg [h₁, h₂] 8 | 9 | set_option egg.groundEqs true in 10 | example (h₁ : ∀ p, p ∧ p) (h₂ : (∀ p, p ∧ p) → q = True) : q = True := by 11 | egg [h₁, h₂] 12 | -------------------------------------------------------------------------------- /Rust/Slotted/src/eta.rs: -------------------------------------------------------------------------------- 1 | use slotted_egraphs::*; 2 | use crate::analysis::*; 3 | 4 | pub fn eta_reduction_rw() -> LeanRewrite { 5 | Rewrite::new_if("≡η", "(λ $0 ?t (app ?f (bvar $0)))", "?f", |subst, _| { 6 | !subst["f"].slots().contains(&Slot::numeric(0)) 7 | }) 8 | } 9 | 10 | pub fn eta_expansion_rw() -> LeanRewrite { 11 | Rewrite::new("≡η+", "?e", "(λ $0 _ (app ?e (bvar $0)))") 12 | } -------------------------------------------------------------------------------- /Lean/Egg/Core/Request/Synth.lean: -------------------------------------------------------------------------------- 1 | import Egg.Core.Request.Term 2 | import Lean 3 | open Lean 4 | 5 | namespace Egg 6 | 7 | @[export lean_is_synthable] 8 | def isSynthable (ty : String) : MetaM Bool := do 9 | let tyExpr ← parse ty (eagerSynth := true) 10 | -- TODO: Is there a way to resolve the bvars? 11 | if tyExpr.hasLooseBVars then return false 12 | let inst? ← Meta.synthInstance? tyExpr 13 | return inst?.isSome 14 | -------------------------------------------------------------------------------- /Lean/Egg/Tests/Proof Irrelevance.lean: -------------------------------------------------------------------------------- 1 | import Egg 2 | 3 | 4 | example : @Eq (Fin 3) ⟨1, Nat.one_lt_succ_succ 1⟩ ⟨1, by omega⟩ := by 5 | rfl 6 | 7 | example : @Eq (Fin 3) ⟨1, Nat.one_lt_succ_succ 1⟩ ⟨1, by omega⟩ := by 8 | egg 9 | 10 | example (l l' : List Nat) (a : Nat) : 11 | (a :: (l ++ l')).getLast (List.cons_ne_nil a _) = ((a :: l) ++ l').getLast (List.cons_ne_nil a _) := 12 | by egg [List.append] 13 | -------------------------------------------------------------------------------- /Lean/Egg/Tests/Erase TC Instances Bug 1.lean: -------------------------------------------------------------------------------- 1 | import Egg 2 | import Lean 3 | 4 | -- This used to fail with `egg: final proof contains expression mvar`, because we were replacing 5 | -- proof steps of rfl-theorems with defeq steps during proof reconstruction. As a result, type class 6 | -- instances appearing in rfl-theorems would not resolve erased type class instance mvars. 7 | 8 | example : (fun _ => 1) 0 = 1 := by 9 | egg 10 | -------------------------------------------------------------------------------- /Lean/Egg/Tests/Normalize.lean: -------------------------------------------------------------------------------- 1 | import Egg 2 | 3 | -- This tests that `Egg.normalize` doesn't introduce loose fvars. 4 | 5 | variable {I : Type u} {f : I → Type v₁} (x y : (i : I) → f i) (i : I) 6 | 7 | instance instMul [∀ i, Mul <| f i] : Mul (∀ i : I, f i) := 8 | ⟨fun f g i => f i * g i⟩ 9 | 10 | theorem mul_apply [∀ i, Mul <| f i] : (x * y) i = x i * y i := 11 | rfl 12 | 13 | example : True = True := by 14 | egg [mul_apply] 15 | -------------------------------------------------------------------------------- /Lean/Egg/Tests/Crash 2.lean: -------------------------------------------------------------------------------- 1 | import Egg 2 | 3 | -- This used to crash, before we changed proof erasure to encode the type of propositions. 4 | 5 | /-- error: egg failed to prove the goal (saturated) -/ 6 | #guard_msgs in 7 | example (a : Array α) (i : Fin a.size) {j : Nat} (v : α) (hj : j < a.size) 8 | (h : i.1 ≠ j) : (a.set i v)[j]'(by simp [*]) = a[j] := by 9 | egg [set, Array.getElem_fin_eq_getElem_toList, List.getElem_set_ne h] 10 | -------------------------------------------------------------------------------- /Lean/Egg/Tests/Stuck 1.lean: -------------------------------------------------------------------------------- 1 | import Egg 2 | 3 | -- This used to get stuck due to missing bvar e-class analyses for the subst and shift constructs. 4 | 5 | theorem getElem_fin [GetElem Cont Nat Elem Dom] (a : Cont) (i : Fin n) (h : Dom a i) : 6 | a[i] = a[i.val] := rfl 7 | 8 | /-- error: egg failed to prove the goal (reached time limit) -/ 9 | #guard_msgs in 10 | example (a : Array α) {i j : Fin a.size} : a[i] = a[j] := by 11 | egg [getElem_fin] 12 | -------------------------------------------------------------------------------- /Lean/Egg/Tests/Ambient Level MVar.lean: -------------------------------------------------------------------------------- 1 | import Egg 2 | 3 | variable (thm₁ : {α : Type} → [Add α] → (a b : α) → a + b = b + a) 4 | variable (thm₂ : {α : Type _} → [Add α] → (a b : α) → a + b = b + a) 5 | 6 | example {a b : Nat} : a + b = b + a := by 7 | egg [thm₁] 8 | 9 | -- TODO: This started failing after bumping to egg v0.10.0. 10 | -- Interestingly `have := @thm₂ Nat` works. 11 | example {a b : Nat} : a + b = b + a := by 12 | sorry -- egg [thm₂] 13 | -------------------------------------------------------------------------------- /Lean/Egg/Tests/Stuck 2.lean: -------------------------------------------------------------------------------- 1 | import Egg 2 | 3 | -- This used to get stuck due to missing bvar e-class analyses for the subst and shift constructs. 4 | 5 | variable (h : ∀ {α β} {f : α → List β} {b} {l : List α}, b ∈ l.flatMap f ↔ ∃ a, a ∈ l ∧ b ∈ f a) 6 | 7 | /-- error: egg failed to prove the goal (reached time limit) -/ 8 | #guard_msgs in 9 | example {xs : List α} {ys : List β} : xy ∈ (xs.flatMap fun a => ys.map (Prod.mk a)) ↔ True := by 10 | egg [h] 11 | -------------------------------------------------------------------------------- /Rust/Egg/build.rs: -------------------------------------------------------------------------------- 1 | use std::process::Command; 2 | 3 | extern crate cc; 4 | 5 | fn main() { 6 | let mut cc = cc::Build::new(); 7 | cc.file("../../C/rev_ffi.c"); 8 | let bytes = Command::new("leanc").args(["--print-cflags"]).output().unwrap().stdout; 9 | let s = String::from_utf8(bytes).unwrap(); 10 | for flag in s.split(" ") { 11 | let flag = flag.trim(); 12 | cc.flag(flag); 13 | } 14 | cc.compile("rev_ffi"); 15 | } 16 | -------------------------------------------------------------------------------- /Lean/Egg/Tests/Class Def.lean: -------------------------------------------------------------------------------- 1 | import Egg 2 | 3 | def Mul.pow [Mul α] [One α] (a : α) : Nat → α 4 | | 0 => 1 5 | | n + 1 => a * (pow a n) 6 | 7 | instance mulPow [Mul α] [One α] : Pow α Nat where 8 | pow := Mul.pow 9 | 10 | example [Mul α] [One α] (a : α) : Mul.pow a 0 = (1 : α) := by 11 | egg [Mul.pow] 12 | 13 | -- Note: This relies on generated type class projection reductions. 14 | example [Mul α] [One α] (a : α) : a ^ 0 = (1 : α) := by 15 | egg [Mul.pow] 16 | -------------------------------------------------------------------------------- /Lean/Egg/Tests/Dataflow Rewriter/lakefile.toml: -------------------------------------------------------------------------------- 1 | name = "DataflowRewriter" 2 | version = "0.1.0" 3 | keywords = ["math"] 4 | defaultTargets = ["DataflowRewriter"] 5 | 6 | [[require]] 7 | name = "mathlib" 8 | scope = "leanprover-community" 9 | 10 | [[require]] 11 | name = "dataflow-rewriter" 12 | git = "https://github.com/VCA-EPFL/dataflow-rewriter/" 13 | rev = "main" 14 | 15 | [[require]] 16 | name = "egg" 17 | path = "../../../.." 18 | 19 | [[lean_lib]] 20 | name = "DataflowRewriter" 21 | -------------------------------------------------------------------------------- /Rust/Egg/src/levels.rs: -------------------------------------------------------------------------------- 1 | use egg::*; 2 | use crate::analysis::*; 3 | 4 | pub fn level_rws() -> Vec { 5 | let mut rws = vec![]; 6 | rws.append(&mut rewrite!("≡maxS"; "(max (succ ?a) (succ ?b))" <=> "(succ (max ?b ?a))")); 7 | rws.push( rewrite!("≡max↔"; "(max ?a ?b)" => "(max ?b ?a)")); 8 | rws.push( rewrite!("≡imax0"; "(imax ?a 0)" => "0")); 9 | rws.push( rewrite!("≡imaxS"; "(imax ?a (succ ?b))" => "(max ?a (succ ?b))")); 10 | rws 11 | } -------------------------------------------------------------------------------- /Lean/Egg/Tests/Shapes DefEqs.lean: -------------------------------------------------------------------------------- 1 | import Egg 2 | 3 | set_option egg.shapes true 4 | 5 | section NatLit 6 | 7 | example : 0 = Nat.zero := by 8 | egg 9 | 10 | example : 1 = Nat.succ 0 := by 11 | egg 12 | 13 | example : Nat.succ 1 = Nat.succ (Nat.succ Nat.zero) := by 14 | egg 15 | 16 | example : Int.ofNat (Nat.succ 1) = Int.ofNat (Nat.succ (Nat.succ Nat.zero)) := by 17 | egg 18 | 19 | example (h : ∀ n, Nat.succ n = n + 1) : 1 = Nat.zero + 1 := by 20 | egg [h] 21 | 22 | end NatLit 23 | -------------------------------------------------------------------------------- /Lean/Egg/Tests/WIP AC.lean: -------------------------------------------------------------------------------- 1 | import Egg 2 | 3 | variable (a : Nat) 4 | 5 | -- TODO: This used to succeed prior to the `better-facts` PR. 6 | 7 | set_option egg.timeLimit 10 in 8 | set_option egg.reporting true in 9 | example : 10 | a + b + c + d + e + f + g + h + i + j + k + l + m + n + o + p + q + r + s + t + u + v + w + x + y + z = 11 | z + y + x + w + v + u + t + s + r + q + p + o + n + m + l + k + j + i + h + g + f + e + d + c + b + a := by 12 | egg [Nat.add_comm, Nat.add_assoc] 13 | -------------------------------------------------------------------------------- /lake-manifest.json: -------------------------------------------------------------------------------- 1 | {"version": "1.1.0", 2 | "packagesDir": ".lake/packages", 3 | "packages": 4 | [{"url": "https://github.com/nomeata/lean-calcify", 5 | "type": "git", 6 | "subDir": null, 7 | "scope": "nomeata", 8 | "rev": "b89b823f26eb35a1d9ed57af2663128d6b3a35c2", 9 | "name": "calcify", 10 | "manifestFile": "lake-manifest.json", 11 | "inputRev": "master", 12 | "inherited": false, 13 | "configFile": "lakefile.toml"}], 14 | "name": "egg", 15 | "lakeDir": ".lake"} 16 | -------------------------------------------------------------------------------- /Lean/Egg/Tests/Soundness.lean: -------------------------------------------------------------------------------- 1 | import Egg 2 | 3 | -- Tests containing examples of backend soundness issues. 4 | 5 | -- In this example egg finds a proof, but we're not performing proof reconstruction (which would be 6 | -- impossible) as a result of setting `exitPoint := some .beforeProof`. 7 | example : False := by 8 | have h : ∀ (f : Unit → Unit) (x : Unit), f x = .unit := fun _ _ => rfl 9 | have : 1 + 2 = 0 := by egg (config := { exitPoint := some .beforeProof }) [h] 10 | contradiction 11 | -------------------------------------------------------------------------------- /Lean/Egg/Tests/Baskets.lean: -------------------------------------------------------------------------------- 1 | import Egg 2 | 3 | attribute [egg add] Nat.add_comm Nat.add_assoc Nat.zero_add 4 | 5 | attribute [egg add_comm] Nat.add_comm 6 | attribute [egg add_assoc] Nat.add_assoc 7 | attribute [egg zero_add] Nat.zero_add 8 | 9 | example : 0 + a + b + c = c + b + a := by 10 | egg +add 11 | 12 | example (h : a = b) : 0 + a + b + c = c + a + a := by 13 | egg +add [h] 14 | 15 | example (h : a = b) : 0 + a + b + c = c + a + a := by 16 | egg +add_comm +add_assoc +zero_add [h] 17 | -------------------------------------------------------------------------------- /Lean/Egg/Tests/Fact Kind.lean: -------------------------------------------------------------------------------- 1 | import Egg 2 | 3 | /-- 4 | error: egg requires arguments to be (proofs of) propositions or (non-propositional) definitions 5 | -/ 6 | #guard_msgs in 7 | example (h : Nat) : 0 = 0 := by 8 | egg [h] 9 | 10 | set_option linter.unusedVariables false in 11 | example (h : True ∧ False) : 0 = 0 := by 12 | egg [h] 13 | 14 | /-- 15 | error: egg requires arguments to be (proofs of) propositions or (non-propositional) definitions 16 | -/ 17 | #guard_msgs in 18 | example (h : Inhabited Nat) : 0 = 0 := by 19 | egg [h] 20 | -------------------------------------------------------------------------------- /Lean/Egg/Tests/Type Level Defeq Rws.lean: -------------------------------------------------------------------------------- 1 | import Egg 2 | 3 | -- This test case ensures that we allow type-level defeq rewrite in erased proof terms. 4 | 5 | variable 6 | {n : Nat} {p : Nat → Prop} [inst : DecidablePred p] 7 | {find : (∃ n, p n) → Nat} 8 | {find_lt_iff : ∀ (h : ∃ n, p n) n, find h < n ↔ ∃ m < n, p m} 9 | 10 | set_option linter.unusedSectionVars false in 11 | include find_lt_iff in 12 | theorem find_le_iff (h : ∃ n : Nat, p n) (n : Nat) : find h ≤ n ↔ ∃ m ≤ n, p m := by 13 | egg [exists_prop, Nat.lt_succ_iff, find_lt_iff] 14 | -------------------------------------------------------------------------------- /Lean/Egg/Tests/Stuck 3.lean: -------------------------------------------------------------------------------- 1 | import Egg 2 | 3 | -- This used to get stuck before eqsat because of a bug causing non-termination in the loop 4 | -- generating derived rewrites. 5 | 6 | variable 7 | {n : Nat} {p : Nat → Prop} [inst : DecidablePred p] 8 | {find : [DecidablePred p] → (∃ n, p n) → Nat} 9 | {find_lt_iff : ∀ (h : ∃ y : Nat, p y) (n : Nat), find h < n ↔ ∃ m < n, p m} 10 | 11 | include find_lt_iff in 12 | theorem find_le_iff (h : ∃ z : Nat, p z) (n : Nat) : find h ≤ n ↔ ∃ m ≤ n, p m := by 13 | egg [exists_prop, Nat.lt_succ_iff, find_lt_iff] 14 | -------------------------------------------------------------------------------- /Lean/Egg/Tactic/Calcify.lean: -------------------------------------------------------------------------------- 1 | import Egg.Lean 2 | import Calcify 3 | import Lean 4 | 5 | open Lean Meta Elab Tactic 6 | 7 | namespace Egg 8 | 9 | def calcify (tk : Syntax) (proof : Expr) (newFVars : Array Name) : TacticM Unit := do 10 | let proof ← simplify proof 11 | check proof 12 | let calcBlock ← delabProof proof 13 | let tactic ← if newFVars.isEmpty then 14 | pure calcBlock 15 | else 16 | let intros ← `(tactic|intro $[$(newFVars.map mkIdent)]*) 17 | mkTacticSeqPrepend intros calcBlock 18 | TryThis.addSuggestion tk tactic (origSpan? := ← getRef) 19 | -------------------------------------------------------------------------------- /Rust/Slotted/src/levels.rs: -------------------------------------------------------------------------------- 1 | use slotted_egraphs::*; 2 | use crate::analysis::*; 3 | 4 | pub fn level_rws() -> Vec { 5 | let mut rws = vec![]; 6 | rws.push(Rewrite::new("≡maxS", "(max (succ ?a) (succ ?b))", "(succ (max ?b ?a))")); 7 | rws.push(Rewrite::new("≡maxS-rev", "(succ (max ?b ?a))", "(max (succ ?a) (succ ?b))")); 8 | rws.push(Rewrite::new("≡max↔", "(max ?a ?b)", "(max ?b ?a)")); 9 | rws.push(Rewrite::new("≡imax0", "(imax ?a 0)", "0")); 10 | rws.push(Rewrite::new("≡imaxS", "(imax ?a (succ ?b))", "(max ?a (succ ?b))")); 11 | rws 12 | } -------------------------------------------------------------------------------- /Rust/Slotted/src/result.rs: -------------------------------------------------------------------------------- 1 | #[derive(Debug)] 2 | pub enum Error { 3 | Init(String), 4 | Goal(String), 5 | Guide(String), 6 | Rewrite(String), 7 | } 8 | 9 | impl ToString for Error { 10 | 11 | fn to_string(&self) -> String { 12 | match self { 13 | Error::Init(s) => format!("⚡️ {s}"), 14 | Error::Goal(s) => format!("⚡️ {s}"), 15 | Error::Guide(s) => format!("⚡️ {s}"), 16 | Error::Rewrite(s) => format!("⚡️ {s}"), 17 | } 18 | } 19 | } 20 | 21 | pub type Res = Result; -------------------------------------------------------------------------------- /Lean/Egg/Tests/Constants.lean: -------------------------------------------------------------------------------- 1 | import Egg 2 | 3 | axiom ax : 1 = 2 4 | 5 | example : 1 = 2 := by 6 | egg [ax] 7 | 8 | /-- 9 | error: egg requires arguments to be (proofs of) propositions or (non-propositional) definitions 10 | -/ 11 | #guard_msgs in 12 | example : 0 = 0 := by 13 | egg [Nat] 14 | 15 | opaque op : Nat 16 | 17 | /-- 18 | error: egg requires arguments to be (proofs of) propositions or (non-propositional) definitions 19 | -/ 20 | #guard_msgs in 21 | example : 0 = 0 := by 22 | egg [op] 23 | 24 | opaque thm : 0 = 0 := rfl 25 | 26 | example : 0 = 0 := by 27 | egg [thm] 28 | -------------------------------------------------------------------------------- /Lean/Egg/Tests/Eta Conservative Intersection.lean: -------------------------------------------------------------------------------- 1 | import Egg 2 | 3 | -- This fails because we with intersection semantics η-reduction is applied before applying 4 | -- `first_correct` which yields open terms in the explanation. 5 | 6 | set_option egg.unionSemantics false 7 | 8 | def first (a : α) (_ : β) : α := a 9 | 10 | theorem first_correct (a : α) (b : β) : first a b = a := rfl 11 | 12 | example (f : α → α) : (fun x => (first f x) x) = f := by 13 | sorry -- egg [first_correct] 14 | 15 | example : (fun (z : α → α → α) x => (fun _y => z) x x) = (fun x => x) := by 16 | sorry -- egg 17 | -------------------------------------------------------------------------------- /Lean/Egg/Core/Direction.lean: -------------------------------------------------------------------------------- 1 | import Lean 2 | open Lean 3 | 4 | namespace Egg 5 | 6 | inductive Direction where 7 | | forward 8 | | backward 9 | deriving Inhabited, BEq, Hashable, Ord, Repr 10 | 11 | namespace Direction 12 | 13 | def description : Direction → String 14 | | forward => "→" 15 | | backward => "←" 16 | 17 | def opposite : Direction → Direction 18 | | forward => backward 19 | | backward => forward 20 | 21 | def merge : Direction → Direction → Direction 22 | | forward, forward | backward, backward => forward 23 | | forward, backward | backward, forward => backward 24 | -------------------------------------------------------------------------------- /Lean/Egg/Tests/Binders.lean: -------------------------------------------------------------------------------- 1 | import Egg 2 | 3 | example : (fun _ : Nat => 0) = (fun _ => 0 + 0) := by 4 | egg [Nat.add_zero] 5 | 6 | example : (fun x => x) = (fun x => x + 0) := by 7 | egg [Nat.add_zero] 8 | 9 | example : (fun x => x) = (fun x => 0 + 0 + x) := by 10 | egg [Nat.zero_add] 11 | 12 | example : (fun x => x) = (fun x => 0 + x) := by 13 | egg [Nat.zero_add] 14 | 15 | example (f : (Nat → Nat) → Bool) : f (fun x => x) = f (fun x => x + 0) := by 16 | egg [Nat.add_zero] 17 | 18 | example (h : ∀ x y : Nat, x = y ↔ y = x) : (∀ x y : Nat, x = y) ↔ (∀ a b : Nat, b = a + 0) := by 19 | egg [h, Nat.add_zero] 20 | -------------------------------------------------------------------------------- /Lean/Egg/Tests/Intro Conditional.lean: -------------------------------------------------------------------------------- 1 | import Egg 2 | 3 | -- These tests check that automatic introduction of antecedents is compatible with conditional 4 | -- rewriting, in the sense that conditional rewrites can use the introduced hypotheses, and proof 5 | -- reconstruction works as expected. 6 | 7 | example {P₁ P₂ : Prop} (h : P₁ → P₂) : P₁ → P₂ := by 8 | egg [h] 9 | 10 | example {P : Prop} (h : 0 = 1 → P) : 0 = 1 → P := by 11 | egg [h] 12 | 13 | example {x y : Nat} {P : Prop} (h : x = y → P) : x = y → P := by 14 | egg [h] 15 | 16 | example (h : ∀ x y : Nat, x = y → y + y = x) {a b : Nat} : a = b → b + b = a := by 17 | egg [h] 18 | -------------------------------------------------------------------------------- /Lean/Egg/Tests/Intro Hygiene.lean: -------------------------------------------------------------------------------- 1 | import Egg 2 | 3 | -- These tests check that we don't override existing fvars' names with automatically introduced 4 | -- names. 5 | 6 | set_option linter.unusedVariables false 7 | 8 | /-- 9 | info: Try this: ⏎ 10 | intro x_1 11 | calc 12 | x_1 13 | _ = x_1 := Eq.refl x_1 14 | -/ 15 | #guard_msgs in 16 | example : Nat → (x : Nat) → x = x := by 17 | intro x 18 | egg? 19 | 20 | /-- 21 | info: Try this: ⏎ 22 | intro x x_1_1 23 | calc 24 | x 25 | _ = x := Eq.refl x 26 | -/ 27 | #guard_msgs in 28 | example : Nat → (x : Nat) → (x_1 : Nat) → x = x := by 29 | intro x_1 30 | egg? 31 | -------------------------------------------------------------------------------- /Lean/Egg/Tests/Thomas.lean: -------------------------------------------------------------------------------- 1 | import Egg 2 | 3 | -- This used to not terminate with union semantics, before we introduced a cutoff to the loose bvar 4 | -- e-class analysis. 5 | 6 | -- TODO: Why does this fail? 7 | set_option egg.unionSemantics true in 8 | example : 9 | ((fun x => (fun t (_y : α) => t) (fun z => x z)) (fun (z : α → α → α) x => ((fun _y => z) x) x)) 10 | = fun _y => (fun z => z) := by 11 | sorry -- egg 12 | 13 | set_option egg.unionSemantics false in 14 | example : 15 | ((fun x => (fun t (_y : α) => t) (fun z => x z)) (fun (z : α → α → α) x => ((fun _y => z) x) x)) 16 | = fun _y => (fun z => z) := by 17 | egg 18 | -------------------------------------------------------------------------------- /Lean/Egg/Tests/Zulip 1.lean: -------------------------------------------------------------------------------- 1 | import Egg 2 | 3 | -- https://leanprover.zulipchat.com/#narrow/stream/113488-general/topic/unification.20problem.20in.20rw/near/438497625 4 | 5 | variable {α} {f : α → α → α} (f_comm : ∀ a b, f a b = f b a) (f_assoc : ∀ a b c, f (f a b) c = f a (f b c)) 6 | 7 | include f_assoc f_comm in 8 | theorem foldl_descend : (head :: tail).foldl f init = f init (tail.foldl f head) := by 9 | induction tail generalizing init <;> egg [List.foldl, *] 10 | 11 | include f_assoc f_comm in 12 | theorem foldl_eq_foldr (l : List α) : l.foldl f init = l.foldr f init := by 13 | induction l <;> egg [List.foldl, List.foldr, foldl_descend, *] 14 | -------------------------------------------------------------------------------- /Rust/Egg/src/result.rs: -------------------------------------------------------------------------------- 1 | #[derive(Debug)] 2 | pub enum Error { 3 | Init(String), 4 | Goal(String), 5 | Guide(String), 6 | Condition(String), 7 | Rewrite(String), 8 | } 9 | 10 | impl ToString for Error { 11 | 12 | fn to_string(&self) -> String { 13 | match self { 14 | Error::Init(s) => format!("⚡️ {s}"), 15 | Error::Goal(s) => format!("⚡️ {s}"), 16 | Error::Guide(s) => format!("⚡️ {s}"), 17 | Error::Condition(s) => format!("⚡️ {s}"), 18 | Error::Rewrite(s) => format!("⚡️ {s}"), 19 | } 20 | } 21 | } 22 | 23 | pub type Res = Result; -------------------------------------------------------------------------------- /Lean/Egg/Tactic/Rewrite/Rules/Gen/Intros.lean: -------------------------------------------------------------------------------- 1 | import Egg.Tactic.Rewrite.Rules.Gen.GenM 2 | 3 | open Lean Meta Elab Tactic 4 | 5 | namespace Egg 6 | 7 | def genIntros (fvars : Array FVarId) (cfg : Config) : TacticM Rewrite.Rules := do 8 | let mut rules := ∅ 9 | for fvar in fvars, idx in [:fvars.size] do 10 | let proof := .fvar fvar 11 | let type ← fvar.getType 12 | let src := .intro idx 13 | if let some rules' ← rules.add? src proof type cfg then 14 | rules := rules' 15 | if cfg.groundEqs then 16 | if let some rules' ← rules.add? (.ground src) proof type cfg .ground then 17 | rules := rules' 18 | return rules 19 | -------------------------------------------------------------------------------- /Lean/Egg/Tests/Boundless BVars.lean: -------------------------------------------------------------------------------- 1 | import Egg 2 | 3 | -- We have to turn of β-reduction in rewrites, as otherwise `thm` is reduced. 4 | set_option egg.betaReduceRws false 5 | 6 | theorem thm : ∀ y x : Nat, (fun _ => (fun _ => x) x) y = x := 7 | fun _ _ => rfl 8 | 9 | -- INVESTIGATE: It's expected that this proof fails (as the theorem is false). However, the e-graph 10 | -- in this example grows infinitely while generating larger and larger bvar indices. 11 | -- How/why does that happen? It does not happen if we use intersection semantics. 12 | example : (fun x => (fun a => (fun a => a) a) 0) = (fun x => x) := by 13 | sorry -- egg [thm] 14 | -------------------------------------------------------------------------------- /Lean/Egg/Tests/Beta.lean: -------------------------------------------------------------------------------- 1 | import Egg 2 | 3 | set_option egg.beta true 4 | 5 | example : (fun x => x) 0 = 0 := by 6 | egg 7 | 8 | example : (fun _ => 1) 0 = 1 := by 9 | egg 10 | 11 | example : (fun x => (fun y => y) x) 0 = 0 := by 12 | egg 13 | 14 | example : (fun x => (fun _ => x) x) 0 = 0 := by 15 | egg 16 | 17 | example : (fun x => (fun _ => x) 0) 1 = 1 := by 18 | egg 19 | 20 | example : id ((fun x => x + 1) 2) = id (2 + 1) := by 21 | egg 22 | 23 | example (h : y + 1 = z) : (fun _ => y + 1) 0 = z := by 24 | egg [h] 25 | 26 | -- Note: This used to break when using explicit e-class substitution. 27 | example (h : y + 1 = z) : (fun x => x + 1) y = z := by 28 | egg [h] 29 | -------------------------------------------------------------------------------- /Lean/Egg/Tests/Subgoals.lean: -------------------------------------------------------------------------------- 1 | import Egg 2 | 3 | set_option egg.subgoals true 4 | 5 | #guard_msgs in 6 | example (a : Nat) (h : True → a = b) : a = b := by 7 | egg [h] 8 | 9 | /-- 10 | error: unsolved goals 11 | b a : Nat 12 | p q : Prop 13 | h : p → q → a = b 14 | ⊢ p 15 | 16 | b a : Nat 17 | p q : Prop 18 | h : p → q → a = b 19 | ⊢ q 20 | -/ 21 | #guard_msgs in 22 | example (a : Nat) (p q : Prop) (h : p → q → a = b) : a = b := by 23 | egg [h] 24 | 25 | /-- 26 | error: unsolved goals 27 | h : ∀ (m : Nat), m = 0 → 1 = 2 28 | ⊢ Nat 29 | 30 | h : ∀ (m : Nat), m = 0 → 1 = 2 31 | ⊢ ?m.26 = 0 32 | -/ 33 | #guard_msgs in 34 | example (h : ∀ m, m = 0 → 1 = 2) : 1 = 2 := by 35 | egg [h] 36 | -------------------------------------------------------------------------------- /Lean/Egg/Core/Encode/Shapes.lean: -------------------------------------------------------------------------------- 1 | import Lean 2 | open Lean 3 | 4 | namespace Egg 5 | 6 | abbrev Shape := String 7 | 8 | partial def shape : Expr → Shape 9 | | .fvar _ | .mvar _ | .sort _ | .const .. | .app .. | .lit _ | .bvar _ | .proj .. => "*" 10 | | .lam _ ty b _ | .forallE _ ty b _ => s!"(→ {shape ty} {shape b})" 11 | | .mdata _ e => shape e 12 | | .letE .. => panic! "'Egg.shape' received let-expression" 13 | 14 | open Meta Elab Term Command 15 | 16 | scoped elab "#shape" t:term : command => liftTermElabM do 17 | logInfo <| shape <| ← inferType <| ← elabTerm t none 18 | 19 | scoped elab "show_shape" t:term : tactic => do 20 | logInfo <| shape <| ← inferType <| ← elabTerm t none 21 | -------------------------------------------------------------------------------- /Lean/Egg/Tests/Eta Conservative.lean: -------------------------------------------------------------------------------- 1 | import Egg 2 | 3 | -- This fails because we first apply `first_correct` which joins the e-classes of `f` and 4 | -- `first f x`. Thus, by our (necessarily) conservative approach to η-reduction, we can't reduce. 5 | 6 | def first (a : α) (_ : β) : α := a 7 | 8 | theorem first_correct (a : α) (b : β) : first a b = a := rfl 9 | 10 | /-- error: egg failed to prove the goal (saturated) -/ 11 | #guard_msgs in 12 | example (f : α → α) : (fun x => (first f x) x) = f := by 13 | egg [first_correct] 14 | 15 | /-- error: egg failed to prove the goal (saturated) -/ 16 | #guard_msgs in 17 | example : (fun (z : α → α → α) x => (fun _y => z) x x) = (fun x => x) := by 18 | egg 19 | -------------------------------------------------------------------------------- /Lean/Egg/Tests/LMVar Encoding.lean: -------------------------------------------------------------------------------- 1 | import Egg 2 | 3 | -- This test checks that level mvars are not encoded as pattern variables when they shouldn't be. 4 | -- 5 | -- This used to crash because the encoding of the goal contained pattern variables. The underlying 6 | -- reason is that the encoding of level mvars (as a pattern variable or constant) is decided by 7 | -- `isLevelMVarAssignable`. And for certain cases we did not bump the mctx depth correctly. 8 | 9 | class AddComm (α : Type _) extends Add α where 10 | comm : ∀ (a b : α), a + b = b + a 11 | 12 | /-- error: egg failed to prove the goal (saturated) -/ 13 | #guard_msgs in 14 | instance (α : Type _) [Add α] : AddComm α where 15 | comm a b := by egg 16 | -------------------------------------------------------------------------------- /Lean/Egg/Tactic/Blocks.lean: -------------------------------------------------------------------------------- 1 | import Egg.Core.Blocks 2 | import Lean 3 | open Lean Elab Tactic 4 | 5 | namespace Egg.Blocks 6 | 7 | declare_syntax_cat egg_blocks 8 | syntax " blocking " (term,*) : egg_blocks 9 | 10 | def parseBlocks (cfg : Config.Normalization) : TSyntax `egg_blocks → TacticM Blocks 11 | | `(egg_blocks|blocking $bs,*) => do 12 | let mut blocks : Blocks := #[] 13 | for b in bs.getElems do 14 | let e ← Tactic.elabTerm b none 15 | -- See `parseGuides` for an explanation of this. 16 | if e.hasSorry then continue 17 | let some block ← Block.from? e cfg | throwErrorAt b "blocking terms must be propositions" 18 | blocks := blocks.push block 19 | return blocks 20 | | _ => unreachable! 21 | -------------------------------------------------------------------------------- /Lean/Egg/Tests/Readme.lean: -------------------------------------------------------------------------------- 1 | import Egg 2 | 3 | example : 0 = 0 := by 4 | egg 5 | 6 | example (a b c : Nat) (h₁ : a = b) (h₂ : b = c) : a = c := by 7 | egg [h₁, h₂] 8 | 9 | open List in 10 | example (as bs : List α) : reverse (as ++ bs) = (reverse bs) ++ (reverse as) := by 11 | induction as generalizing bs with 12 | | nil => egg [reverse_nil, append_nil, List.append] 13 | | cons => egg [*, append_assoc, reverse_cons, List.append] 14 | 15 | variable (a b c d : Int) 16 | 17 | example : ((a * b) - (2 * c)) * d - (a * b) = (d - 1) * (a * b) - (2 * c * d) := by 18 | egg [Int.sub_mul, Int.sub_sub, Int.add_comm, Int.mul_comm, Int.one_mul] 19 | 20 | example {p q r : Prop} (h₁ : p) (h₂ : p ↔ q) (h₃ : q → (p ↔ r)) : p ↔ r := by 21 | egg [h₁, h₂, h₃] 22 | -------------------------------------------------------------------------------- /.github/workflows/ci.yml: -------------------------------------------------------------------------------- 1 | name: Build and Test 2 | on: 3 | push: 4 | branches: 5 | - "main" 6 | pull_request: 7 | merge_group: 8 | 9 | jobs: 10 | build-and-test: 11 | name: Build and Test 12 | runs-on: ubuntu-latest 13 | steps: 14 | - name: Checkout Repository 15 | uses: actions/checkout@v3 16 | - name: Install elan 17 | run: | 18 | set -o pipefail 19 | curl https://raw.githubusercontent.com/leanprover/elan/master/elan-init.sh -sSf | sh -s -- --default-toolchain none -y 20 | "$HOME/.elan/bin/lean" --version 21 | echo "$HOME/.elan/bin" >> $GITHUB_PATH 22 | - name: Build Library 23 | run: lake build 24 | - name: Run Tests 25 | run: lake test -- ci 26 | -------------------------------------------------------------------------------- /Lean/Egg/Tests/BVar Index Correction.lean: -------------------------------------------------------------------------------- 1 | import Egg 2 | 3 | example (h : ∀ x y : Nat, x = y ↔ y = x) : (∀ x y : Nat, x = y) ↔ (∀ a b : Nat, b = a + 0) := by 4 | egg [h, Nat.add_zero] 5 | 6 | -- We have to disable β-reduction as part of normalization, as otherwise `thm₁,₂` are useless, and 7 | -- disable β-reduction in egg, as this interferes with the test cases. 8 | set_option egg.betaReduceRws false 9 | set_option egg.beta false 10 | 11 | -- This theorem is only applicable in the forward direction. 12 | theorem thm₁ : ∀ x y : Nat, (x, y).fst = (fun _ => x) (nat_lit 1) := 13 | fun _ _ => rfl 14 | 15 | /-- error: egg failed to prove the goal (reached time limit) -/ 16 | #guard_msgs in 17 | example : (fun x => (x, 5).fst) = (fun _ : Nat => (fun x => x) 1) := by 18 | egg [thm₁] 19 | -------------------------------------------------------------------------------- /Lean/Egg/Tests/StrLit.lean: -------------------------------------------------------------------------------- 1 | import Egg 2 | 3 | -- TODO: It seems that for strings not containing "(", ")", or " ", egg removes the surrounding 4 | -- double-quotes in explanations. I think the problem comes somewhere from here: 5 | -- https://docs.rs/symbolic_expressions/5.0.3/src/symbolic_expressions/formatter.rs.html#10-25 6 | 7 | example : "a" = "a" := by 8 | sorry -- egg 9 | 10 | example : "a\nb" = "a\nb" := by 11 | sorry -- egg 12 | 13 | example : "" = "" := by 14 | egg 15 | 16 | example : " " = " " := by 17 | egg 18 | 19 | example : "a b" = "a b" := by 20 | egg 21 | 22 | example : "(" = "(" := by 23 | egg 24 | 25 | example : ")" = ")" := by 26 | egg 27 | 28 | example (h : "Le " ++ " an" = "Le an") : "Le " ++ " an" = "Le an" := by 29 | egg [h] 30 | -------------------------------------------------------------------------------- /Lean/Egg/Tests/Math/Math/PushNeg.lean: -------------------------------------------------------------------------------- 1 | import Egg 2 | import Mathlib 3 | 4 | attribute [egg neg] Ne not_not _root_.not_imp not_iff not_or not_and not_and_or not_forall not_exists 5 | 6 | -- NOTE: The following example relies on backtracking for choosing between `=` and `↔` when 7 | -- constructing proofs of reified equivalences. 8 | example : ¬¬(p ↔ p) := by 9 | egg +neg 10 | 11 | example : ¬¬(p = p) := by 12 | egg +neg 13 | 14 | example : ¬(p ≠ p) := by 15 | egg +neg 16 | 17 | example : (¬p ∧ ¬q) → ¬(p ∨ q) := by 18 | egg +neg 19 | 20 | -- NOTE: This relies on the builtin theorem for modus ponens, and hence assignment of weak vars. 21 | example : ¬(p ∧ q) → (p → ¬q) := by 22 | egg +neg 23 | 24 | example (r : α → Prop) : (∀ x, ¬r x) → ¬(∃ x, r x) := by 25 | egg +neg 26 | -------------------------------------------------------------------------------- /Lean/Egg/Tests/Cond Valid MVar Dirs.lean: -------------------------------------------------------------------------------- 1 | import Egg 2 | 3 | -- This test ensures that condition mvars are correctly taken into account when determining the 4 | -- valid direction of rewrites. 5 | 6 | /-- 7 | trace: [egg.rewrites.explicit] Basic (1) 8 | [egg.rewrites.explicit] #0(⇒): _h 9 | [egg.rewrites.explicit] f ?b = f ?a 10 | [egg.rewrites.explicit] Conditions 11 | [egg.rewrites.explicit] ?a = ?b 12 | [egg.rewrites.explicit] LHS MVars 13 | [?b: [unconditionallyVisible]] 14 | [egg.rewrites.explicit] RHS MVars 15 | [?a: [unconditionallyVisible]] 16 | -/ 17 | #guard_msgs in 18 | set_option trace.egg.rewrites.explicit true in 19 | set_option egg.groundEqs false in 20 | example (f : Nat → Nat) (_h : ∀ a b : Nat, a = b → f b = f a) : true = true := by 21 | egg [_h] 22 | -------------------------------------------------------------------------------- /Lean/Egg/Tests/Hypotheses.lean: -------------------------------------------------------------------------------- 1 | import Egg 2 | 3 | -- Tests for different variants of declaring and using hypotheses. 4 | 5 | example (a b : Nat) : a + b = b + a := by 6 | have h := Nat.add_comm 7 | egg [h] 8 | 9 | example (a : Nat) : (∀ x, x + 1 = 1 + x) → a + 1 = 1 + a := 10 | fun h => by egg [h] 11 | 12 | example (a : Nat) : (∀ x, x + 1 = 1 + x) → a + 1 = 1 + a := 13 | (by egg [·]) 14 | 15 | example (a : Nat) : a + 1 = 1 + a := by 16 | egg [(fun _ => Nat.add_comm .. : ∀ x, x + 1 = 1 + x)] 17 | 18 | example (a : Nat) : (∀ x, x + 1 = 1 + x) → a + 1 = 1 + a := 19 | fun _ => by egg [‹∀ _ : Nat, _›] 20 | 21 | variable (h : ∀ x, x + 1 = 1 + x) in 22 | example (a : Nat) : a + 1 = 1 + a := by 23 | egg [h] 24 | 25 | variable {h : ∀ x, x + 1 = 1 + x} in 26 | example (a : Nat) : a + 1 = 1 + a := by 27 | egg [h] 28 | -------------------------------------------------------------------------------- /Lean/Egg/Tests/Reconstruction.lean: -------------------------------------------------------------------------------- 1 | import Egg 2 | 3 | /- TODO: Why is proof reconstruction failing here? The explanation seems totally fine: 4 | 5 | (app (app (λ (const Nat) (app f (bvar 0))) y) y) 6 | (app (app (Rewrite=> ≡η f) y) y) 7 | (app (Rewrite=> #0 g) y) 8 | (Rewrite<= #1-rev (app (app i y) (lit 0))) 9 | (app (app (Rewrite<= ≡η (λ (const Nat) (app i (bvar 0)))) y) (lit 0)) 10 | 11 | Even stranger, if you flip the LHS and RHS of the goal, it suddenly works. The produced explanation 12 | in that case is the same just backwards. So a hacky fix would be to always obtain explanations from 13 | egg in both directions and if one fails try the other. 14 | -/ 15 | example (f i : Nat → Nat → Nat) (h₁ : f y = g) (h₂ : g y = i y (nat_lit 0)) : 16 | (f ·) y y = (i ·) y (nat_lit 0) := by 17 | sorry -- egg [h₁, h₂] 18 | -------------------------------------------------------------------------------- /Lean/Egg/Tests/Crash 3.lean: -------------------------------------------------------------------------------- 1 | import Egg 2 | 3 | -- This used to crash because we were not considering proof erasure when determining a rewrite's 4 | -- mvars, which resulted in allowed directions for rewrites which should not have been valid. As a 5 | -- result the calls to `union_instantiations` in the `LeanApplier` would crash, because the 6 | -- substitution resulting from e-matching was incomplete. 7 | 8 | /-- 9 | error: egg received invalid explanation: step contains non-defeq type-level rewrite in proof or type class instance 10 | -/ 11 | #guard_msgs in 12 | theorem Array.extract_extract' {s1 e2 e1 s2 : Nat} {a : Array α} (h : s1 + e2 ≤ e1) : 13 | (a.extract s1 e1).extract s2 e2 = a.extract (s1 + s2) (s1 + e2) := by 14 | apply ext 15 | · simp only [size_extract] 16 | omega 17 | · intro i h1 h2 18 | egg [getElem_extract, Nat.add_assoc] 19 | -------------------------------------------------------------------------------- /Lean/Egg/Core/Explanation/Parse/Basic.lean: -------------------------------------------------------------------------------- 1 | import Egg.Core.Explanation.Parse.Shared 2 | import Egg.Core.Explanation.Parse.Egg 3 | import Egg.Core.Explanation.Parse.Slotted 4 | import Lean 5 | 6 | open Lean 7 | 8 | namespace Egg.Explanation 9 | 10 | -- Note: This could be generalized to any monad with an environment and exceptions. 11 | def Raw.parse (raw : Explanation.Raw) : MetaM Explanation := do 12 | let (str, stx_cat) := 13 | if raw.slotted 14 | then (raw.str.replace "$" "", `slotted_expl) -- HACK 15 | else (raw.str, `egg_expl) 16 | match Parser.runParserCategory (← getEnv) stx_cat str with 17 | | .error err => throwError s!"{ParseError.msgPrefix}\n{err}\n\n{raw.str}" 18 | | .ok stx => 19 | if raw.slotted 20 | then return { ← parseSlottedExpl ⟨stx⟩ with kind := raw.kind } 21 | else return { ← parseEggExpl ⟨stx⟩ with kind := raw.kind } 22 | -------------------------------------------------------------------------------- /Lean/Egg/Tests/Math/Math/Zulip/Binders.lean: -------------------------------------------------------------------------------- 1 | import Mathlib.Data.Set.Basic 2 | import Egg 3 | 4 | variable (s t : Set ℕ) 5 | open Set 6 | 7 | -- https://leanprover.zulipchat.com/#narrow/channel/113488-general/topic/Is.20simp.20only.20a.20sequence.20of.20intelligent.20rewrites.3F/near/530885662 8 | 9 | example (h : s ⊆ t) : s ∩ u ⊆ t ∩ u := by 10 | rw [subset_def, inter_def, inter_def] 11 | rw [subset_def] at h 12 | -- goal: ∀ x ∈ {a | a ∈ s ∧ a ∈ u}, x ∈ {a | a ∈ t ∧ a ∈ u} 13 | simp only [mem_setOf] 14 | rintro x ⟨xs, xu⟩ 15 | exact ⟨h _ xs, xu⟩ 16 | 17 | example (h : s ⊆ t) : s ∩ u ⊆ t ∩ u := by 18 | rw [subset_def, inter_def, inter_def] 19 | rw [subset_def] at h 20 | simp only [mem_setOf] 21 | have : ∀ (x : ℕ), x ∈ s ∧ x ∈ u → x ∈ t ∧ x ∈ u := by 22 | set_option trace.egg true in 23 | egg [And.left, And.right, subset_def, inter_def, mem_setOf] 24 | sorry 25 | -------------------------------------------------------------------------------- /Lean/Egg/Tactic/Guides.lean: -------------------------------------------------------------------------------- 1 | import Egg.Core.Guides 2 | import Lean 3 | open Lean Elab Tactic 4 | 5 | namespace Egg.Guides 6 | 7 | declare_syntax_cat egg_guides 8 | syntax &" using " (term,*) : egg_guides 9 | 10 | def parseGuides : TSyntax `egg_guides → TacticM Guides 11 | | `(egg_guides|using $gs,*) => do 12 | let mut guides : Guides := #[] 13 | for g in gs.getElems, idx in [:gs.getElems.size] do 14 | let g ← Tactic.elabTerm g none 15 | -- For some reason Lean introduces the term `sorry = true` if the list of guides is empty. 16 | -- This expression contains an mvar which results in a guide term with a pattern variable, 17 | -- which crashed egg. So we filter for this here. 18 | if g.hasSorry then continue 19 | let guide ← Guide.from g (.guide idx (derived := false)) 20 | guides := guides.push guide 21 | return guides 22 | | _ => unreachable! 23 | -------------------------------------------------------------------------------- /Lean/Egg/Tests/Math/Math/Cast.lean: -------------------------------------------------------------------------------- 1 | import Mathlib 2 | import Egg 3 | 4 | open Mathlib.Tactic.Zify in 5 | attribute [egg zify] natCast_eq natCast_le natCast_lt natCast_ne natCast_dvd 6 | 7 | open Mathlib.Tactic.Zify Nat Int in 8 | attribute [egg zify] cast_sub_of_add_le cast_sub_of_lt natCast_inj natAbs_cast Int.cast_id 9 | natCast_pred_of_pos natCast_div natCast_mod natCast_dvd_natCast ofNat_isUnit 10 | WithZero.coe_div abs_natCast Int.cast_sub 11 | 12 | open Mathlib.Tactic.Zify Nat Int in 13 | #check Nat.cast_sub 14 | 15 | -- TODO: I think there's some problem with namespaces and egg baskets. 16 | 17 | example (a b : Nat) : (a = b) ↔ (a : Int) = (b : Int) := by 18 | egg [Mathlib.Tactic.Zify.natCast_eq] 19 | 20 | open Mathlib.Tactic.Zify in 21 | attribute [egg z] natCast_eq 22 | 23 | example (a b : Nat) : (a = b) ↔ (a : Int) = (b : Int) := by 24 | egg +z 25 | -------------------------------------------------------------------------------- /Lean/Egg/Core/Request/Term.lean: -------------------------------------------------------------------------------- 1 | 2 | import Egg.Core.Encode.Basic 3 | import Egg.Core.Explanation.Basic 4 | import Egg.Core.Request.Basic 5 | import Egg.Core.Request.EGraph 6 | import Lean 7 | open Lean 8 | 9 | namespace Egg 10 | 11 | def parse (s : String) (eagerSynth := false) : MetaM Expr := do 12 | match Parser.runParserCategory (← getEnv) `egg_expr s with 13 | | .ok stx => 14 | let a := Explanation.parseExpr ⟨ stx ⟩ 15 | match a with 16 | | .error _ => throwError "not ok" 17 | | .ok (c, _) => c.toExpr (synthesize := eagerSynth) 18 | | .error _ => throwError "meh" 19 | 20 | structure Request.Term where 21 | enode : Nat 22 | 23 | @[extern "get_term"] 24 | private opaque getTermRaw (graph : EGraph.Obj) (slotted : Bool) (enode : Nat) : String 25 | 26 | def EGraph.get (graph : EGraph) (req : Request.Term) : MetaM Expr := do 27 | parse <| getTermRaw graph.obj graph.slotted req.enode 28 | -------------------------------------------------------------------------------- /Lean/Egg/Tests/Calc.lean: -------------------------------------------------------------------------------- 1 | import Egg 2 | 3 | example : 0 = 0 := by 4 | egg calc 0 = 0 5 | 6 | example : 0 = 0 := by 7 | egg calc _ = 0 8 | 9 | example : 0 = 0 := by 10 | egg calc 0 11 | _ = 0 12 | 13 | example : 0 = 0 := by 14 | egg calc _ = _ 15 | 16 | example (h : a = b) : a = b := by 17 | egg calc [h] _ = _ 18 | 19 | example (h₁ : a = b) (h₂ : b = c) : a = c := by 20 | egg calc [h₁, h₂] 21 | a = b 22 | _ = c 23 | 24 | example : ∀ a b c : Nat, (a = b) → (b = c) → a = c := by 25 | intro a b c h₁ h₂ 26 | egg calc [h₁, h₂] 27 | a = b 28 | _ = c 29 | 30 | example (h₁ : a = b) (h₂ : b = c) : a = c := by 31 | egg calc 32 | a = b with [h₁] 33 | _ = c with [h₂] 34 | 35 | example (h₁ : 0 = 0 → a = b) : a = b := by 36 | egg calc [h₁, (rfl : 0 = 0)] 37 | _ = _ 38 | 39 | example (h₁ : 0 = 0 → a = b) : a = b := by 40 | egg calc [h₁] 41 | _ = _ with [(rfl : 0 = 0)] 42 | -------------------------------------------------------------------------------- /Lean/Egg/Tests/Nate.lean: -------------------------------------------------------------------------------- 1 | import Egg 2 | 3 | attribute [egg nat] Nat.add_mul Nat.one_mul Nat.add_assoc Nat.zero_mul Nat.add_zero Nat.add_comm 4 | 5 | inductive Vect (α : Type _) : Nat → Type _ where 6 | | nil : Vect α 0 7 | | cons : α → Vect α n → Vect α (n + 1) 8 | 9 | namespace Vect 10 | 11 | def drop : (n : Nat) → Vect α (k + n) → Vect α k 12 | | 0, v => v 13 | | n + 1, cons _ r => drop n r 14 | 15 | def take : (n : Nat) → Vect α (k + n) → Vect α n 16 | | 0, _ => nil 17 | | n + 1, cons h r => cons h (take n r) 18 | 19 | def cast (xs : Vect α n) (h : n = m := by egg +nat) : Vect α m := 20 | h ▸ xs 21 | 22 | def slide (sz sp : Nat) (xs : Vect α <| sz + n * (sp + 1)) : Vect (Vect α sz) (n + 1) := 23 | match n with 24 | | 0 => cons xs.cast nil 25 | | m + 1 => 26 | let window := xs.cast.take sz (k := (m + 1) * (sp + 1)) 27 | let rem := slide sz sp <| xs.cast.drop <| sp + 1 28 | cons window rem 29 | -------------------------------------------------------------------------------- /Lean/Egg/Core/Request/Equiv.lean: -------------------------------------------------------------------------------- 1 | import Egg.Core.Encode.Basic 2 | import Egg.Core.Explanation.Basic 3 | import Egg.Core.Request.Basic 4 | import Egg.Core.Request.EGraph 5 | import Lean 6 | open Lean 7 | 8 | namespace Egg.Request 9 | 10 | structure Equiv where 11 | init : Expression 12 | goal : Expression 13 | 14 | def Equiv.encoding (init goal : Expr) (cfg : Config.Encoding) : MetaM Equiv := 15 | return { init := ← encode init cfg, goal := ← encode goal cfg } 16 | 17 | end Request 18 | 19 | @[extern "explain_equiv"] 20 | private opaque explainEquivRaw (graph : EGraph.Obj) (slotted : Bool) (e₁ e₂ : Expression) : Request.Result.Raw 21 | 22 | def EGraph.run (graph : EGraph) (req : Request.Equiv) : Option Explanation.Raw := Id.run do 23 | let { kind := kind?, expl, .. } := explainEquivRaw graph.obj graph.slotted req.init req.goal 24 | let some kind := kind?.toKind? | return none 25 | return some { kind, str := expl, slotted := graph.slotted } 26 | -------------------------------------------------------------------------------- /Lean/Egg/Core/Explanation/Basic.lean: -------------------------------------------------------------------------------- 1 | import Egg.Core.Source 2 | import Egg.Core.Direction 3 | import Egg.Core.Rewrite.Rule 4 | 5 | open Lean 6 | 7 | namespace Egg.Explanation 8 | 9 | inductive Kind where 10 | | sameEClass 11 | | eqTrue 12 | 13 | structure Raw where 14 | kind : Kind 15 | str : String 16 | slotted : Bool 17 | 18 | namespace Rewrite 19 | 20 | structure Descriptor where 21 | id : Rewrite.Rule.Id 22 | dir : Direction 23 | weakVars : Array (Nat × Nat) 24 | deriving Inhabited, Repr 25 | 26 | structure Info extends Descriptor where 27 | pos? : Option SubExpr.Pos 28 | deriving Inhabited, Repr 29 | 30 | end Rewrite 31 | 32 | structure Step extends Rewrite.Info where 33 | dst : Expr 34 | deriving Inhabited 35 | 36 | structure Steps where 37 | start : Expr 38 | steps : Array Explanation.Step 39 | 40 | end Explanation 41 | 42 | structure Explanation extends Explanation.Steps where 43 | kind : Explanation.Kind 44 | -------------------------------------------------------------------------------- /Lean/Egg/Tests/Builtins.lean: -------------------------------------------------------------------------------- 1 | import Egg 2 | 3 | /-- error: egg failed to prove the goal (saturated) -/ 4 | #guard_msgs in 5 | set_option egg.builtins false in 6 | example : Nat.succ a = a + 1 := by 7 | egg 8 | 9 | set_option egg.builtins true in 10 | example : Nat.succ a = a + 1 := by 11 | egg 12 | 13 | /-- error: egg failed to prove the goal (saturated) -/ 14 | #guard_msgs in 15 | set_option egg.builtins false in 16 | example (a : Nat) (h₁ : a < b → 1 = 2) (h₂ : b > a) : 1 = 2 := by 17 | egg [h₁, h₂] 18 | 19 | set_option egg.builtins true in 20 | example (a : Nat) (h₁ : a < b → 1 = 2) (h₂ : b > a) : 1 = 2 := by 21 | egg [h₁, h₂] 22 | 23 | /-- error: egg failed to prove the goal (saturated) -/ 24 | #guard_msgs in 25 | set_option egg.builtins false in 26 | example (a : Nat) (h₁ : a ≤ b → 1 = 2) (h₂ : b ≥ a) : 1 = 2 := by 27 | egg [h₁, h₂] 28 | 29 | set_option egg.builtins true in 30 | example (a : Nat) (h₁ : a ≤ b → 1 = 2) (h₂ : b ≥ a) : 1 = 2 := by 31 | egg [h₁, h₂] 32 | -------------------------------------------------------------------------------- /Lean/Egg/Tests/Crash 1.lean: -------------------------------------------------------------------------------- 1 | import Egg 2 | 3 | -- This used to crash. I don't recall why, but something related to unassigned variables during 4 | -- substitution after e-matching. This seems to have been fixed by a change in how conditions are 5 | -- handled. 6 | 7 | axiom beq_ext {α : Type _} (inst1 : BEq α) (inst2 : BEq α) 8 | (h : ∀ x y, @BEq.beq _ inst1 x y = @BEq.beq _ inst2 x y) : inst1 = inst2 9 | 10 | open Classical in 11 | axiom beq_eq_decide {α : Type _} [BEq α] [LawfulBEq α] {a b : α} : (a == b) = decide (a = b) 12 | 13 | -- TODO: This fails because the `BEq` instances on the LHS and RHS are different, but erasure makes 14 | -- them uniform, so when the goal is "proven" by reflexivity we fail. The error message here 15 | -- should be better. 16 | theorem lawful_beq_subsingleton 17 | {α : Type _} (inst1 : BEq α) (inst2 : BEq α) [@LawfulBEq α inst1] [@LawfulBEq α inst2] : 18 | inst1 = inst2 := by 19 | apply beq_ext 20 | sorry -- egg [beq_eq_decide] 21 | -------------------------------------------------------------------------------- /Lean/Egg/Tests/Shapes Rerun.lean: -------------------------------------------------------------------------------- 1 | import Egg 2 | 3 | -- Checks that we rerun with shapes when proof reconstruction fails if `egg.retryWithShapes = true`. 4 | -- Without rerunning with shapes, the first fails during proof reconstruction with the explanation 5 | -- shown when setting `exitPoint := some .beforeProof`. 6 | 7 | /-- 8 | error: egg failed to build proof step 0: unification failure for LHS of rewrite #0: 9 | 10 | id Nat.add 11 | vs 12 | ?m.30 ?m.28 13 | in 14 | id Nat.add 15 | and 16 | () 17 | 18 | • Types: ⏎ 19 | ?m.28: Unit 20 | ?m.30: Unit → Unit 21 | 22 | • Read Only Or Synthetic Opaque MVars: [] 23 | -/ 24 | #guard_msgs in 25 | set_option egg.retryWithShapes false in 26 | example (h : ∀ (f : Unit → Unit) (u : Unit), f u = .unit) : id Nat.add = id Nat.mul := by 27 | egg [h] 28 | 29 | /-- error: egg failed to prove the goal (saturated) -/ 30 | #guard_msgs in 31 | set_option egg.retryWithShapes true in 32 | example (h : ∀ (f : Unit → Unit) (u : Unit), f u = .unit) : id Nat.add = id Nat.mul := by 33 | egg [h] 34 | -------------------------------------------------------------------------------- /Lean/Egg/Tests/Crash 4.lean: -------------------------------------------------------------------------------- 1 | import Egg 2 | 3 | -- This used to fail because we were not considering proof mvars when refreshing a rewrite's mvars. 4 | -- As a result, rewrites produced by type class specialization could contain unrefreshed mvars for 5 | -- proof terms (and their types). These did not affect the computation of valid directions, but did 6 | -- affect the encoding of erased proofs which then contained mvars which should have been assigned 7 | -- by type class specialization. As a result, `union_instantiations` would crash with the 8 | -- substitution lacking an assignment for the pattern variable corresponding to the non-refreshed 9 | -- proof mvar. 10 | 11 | variable 12 | {n : Nat} {p : Nat → Prop} [inst : DecidablePred p] 13 | {find : [DecidablePred p] → (∃ n, p n) → Nat} 14 | {find_lt_iff : ∀ (h : ∃ y : Nat, p y) (n : Nat), find h < n ↔ ∃ m < n, p m} 15 | 16 | include find_lt_iff in 17 | theorem find_le_iff (h : ∃ z : Nat, p z) (n : Nat) : find h ≤ n ↔ ∃ m ≤ n, p m := by 18 | egg [exists_prop, Nat.lt_succ_iff, find_lt_iff, h] 19 | -------------------------------------------------------------------------------- /Lean/Egg/Tests/Level Defeq.lean: -------------------------------------------------------------------------------- 1 | import Egg 2 | 3 | -- The only difference between these two and the next two examples is the order of universe levels 4 | -- in `[]`. That is, the second examples require commutativity of `Level.max`. 5 | 6 | set_option egg.levels false in 7 | example (f : α → γ) (g : β → δ) : List.map (Prod.map f g) [] = [] := by 8 | egg [List.map] 9 | 10 | set_option egg.levels true in 11 | example (f : α → γ) (g : β → δ) : List.map (Prod.map f g) [] = [] := by 12 | egg [List.map] 13 | 14 | variable {α : Type _} {β : Type _} {γ : Type _} {δ : Type _} 15 | 16 | /-- error: egg failed to prove the goal (saturated) -/ 17 | #guard_msgs in 18 | set_option egg.levels false in 19 | example (f : α → γ) (g : β → δ) : List.map (Prod.map f g) [] = [] := by 20 | egg [List.map] 21 | 22 | set_option egg.levels true in 23 | example (f : α → γ) (g : β → δ) : List.map (Prod.map f g) [] = [] := by 24 | egg [List.map] 25 | 26 | -- This example requires `Level.succ (Level.max _ _) = Level.max (Level.succ _) (Level.succ _)`. 27 | example (h : ∀ γ : Type (max u v), γ = id γ) (α : Type u) (β : Type v) : (α × β) = id (α × β) := by 28 | egg [h] 29 | -------------------------------------------------------------------------------- /Rust/Slotted/src/beta.rs: -------------------------------------------------------------------------------- 1 | use slotted_egraphs::*; 2 | use crate::analysis::*; 3 | 4 | pub fn beta_reduction_rws(small_step: bool) -> Vec { 5 | if small_step { 6 | let mut rws = vec![]; 7 | rws.push(Rewrite::new("≡β", "(app (λ $0 ?t ?b) ?a)", "(↦ $0 ?a ?b)")); 8 | rws.append(&mut subst_rws()); 9 | rws 10 | } else { 11 | vec![Rewrite::new("≡β", "(app (λ $0 ?t ?b) ?a)", "?b[(bvar $0) := ?a]")] 12 | } 13 | } 14 | 15 | fn subst_rws() -> Vec { 16 | let mut rws = vec![]; 17 | rws.push(Rewrite::new("↦bvar", "(↦ $x ?z (bvar $x))", "?z")); 18 | rws.push(Rewrite::new("↦app", "(↦ $x ?z (app ?a ?b))", "(app (↦ $x ?z ?a) (↦ $x ?z ?b))")); 19 | rws.push(Rewrite::new("↦λ", "(↦ $x ?z (λ $y ?t ?b))", "(λ $y (↦ $x ?z ?t) (↦ $x ?z ?b))")); 20 | rws.push(Rewrite::new("↦∀", "(↦ $x ?z (∀ $y ?t ?b))", "(∀ $y (↦ $x ?z ?t) (↦ $x ?z ?b))")); 21 | rws.push(Rewrite::new_if("↦|", "(↦ $x ?z ?e)", "?e", |subst, _| { 22 | !subst["e"].slots().contains(&Slot::named("x")) 23 | })); 24 | // TODO: We don't propagate substitutions over erased terms at the moment. 25 | rws 26 | } -------------------------------------------------------------------------------- /Lean/Egg/Tests/batteries/SimpOnlyOverride.lean: -------------------------------------------------------------------------------- 1 | import Lean 2 | import Egg 3 | open Lean Meta Elab Parser Tactic 4 | 5 | elab_rules : tactic 6 | | `(simp| simp only $[[$lemmas:simpLemma,*]]?) => do 7 | let simpStx ← if let some lems := lemmas then `(tactic| simp only [$lems,*]) else `(tactic| simp only) 8 | let mut args ← simpOnlyBuiltins.toArray.mapM fun b => `(egg_arg|$(mkIdent b):ident) 9 | if let some lems := lemmas then 10 | for lem in lems.getElems do 11 | -- syntax simpLemma := (simpPre <|> simpPost)? patternIgnore("← " <|> "<- ")? term 12 | let lemTerm : Term := ⟨lem.raw[2]⟩ 13 | args := args.push <| ← `(egg_arg|$lemTerm:term) 14 | focus do 15 | let s ← saveState 16 | let goal ← getMainGoal 17 | evalSimp simpStx 18 | unless (← getGoals).isEmpty do return 19 | let some _ ← Egg.Congr.from? (← Egg.normalize (← goal.getType) .noReduce) | return 20 | s.restore 21 | try 22 | evalTactic (← `(tactic|egg [$args,*])) 23 | logWarning "egg succeeded" 24 | catch err => 25 | s.restore 26 | logWarning m!"egg failed: {err.toMessageData}" 27 | evalSimp simpStx 28 | -------------------------------------------------------------------------------- /Lean/Egg/Tests/mathlib4/SimpOnlyOverride.lean: -------------------------------------------------------------------------------- 1 | import Lean 2 | import Egg 3 | open Lean Meta Elab Parser Tactic 4 | 5 | elab_rules : tactic 6 | | `(simp| simp only $[[$lemmas:simpLemma,*]]?) => do 7 | let simpStx ← if let some lems := lemmas then `(tactic| simp only [$lems,*]) else `(tactic| simp only) 8 | let mut args ← simpOnlyBuiltins.toArray.mapM fun b => `(egg_arg|$(mkIdent b):ident) 9 | if let some lems := lemmas then 10 | for lem in lems.getElems do 11 | -- syntax simpLemma := (simpPre <|> simpPost)? patternIgnore("← " <|> "<- ")? term 12 | let lemTerm : Term := ⟨lem.raw[2]⟩ 13 | args := args.push <| ← `(egg_arg|$lemTerm:term) 14 | focus do 15 | let s ← saveState 16 | let goal ← getMainGoal 17 | evalSimp simpStx 18 | unless (← getGoals).isEmpty do return 19 | let some _ ← Egg.Congr.from? (← Egg.normalize (← goal.getType) .noReduce) | return 20 | s.restore 21 | try 22 | evalTactic (← `(tactic|egg [$args,*])) 23 | logWarning "egg succeeded" 24 | catch err => 25 | s.restore 26 | logWarning m!"egg failed: {err.toMessageData}" 27 | evalSimp simpStx 28 | -------------------------------------------------------------------------------- /Lean/Egg/Tests/Condition Subgoals.lean: -------------------------------------------------------------------------------- 1 | import Egg 2 | 3 | set_option egg.subgoals true 4 | 5 | example : 0 = 0 := by 6 | egg 7 | 8 | example (h : a = b) : a = b := by 9 | egg [h] 10 | 11 | -- Note: The condition is automatically proven, because during proof reconstruction we derive 12 | -- unassigned propositional mvars by querying the e-graph. 13 | example (h : True → a = b) : a = b := by 14 | egg [h] 15 | 16 | /-- 17 | error: unsolved goals 18 | α✝ : Sort u_1 19 | a b : α✝ 20 | h : False → a = b 21 | ⊢ False 22 | -/ 23 | #guard_msgs in 24 | example (h : False → a = b) : a = b := by 25 | egg [h] 26 | 27 | example (h₁ : x = y → a = b) (h₂ : x = y) : a = b := by 28 | egg [h₁, h₂] 29 | 30 | example (h₁ : x = y → a = b) (h₂ : x = y) : a = b := by 31 | egg [h₁] 32 | exact h₂ 33 | 34 | -- Note: When condition subgoals are disabled, the given rewrite is applicable in both directions, 35 | -- because the condition resolves `n` for the backward direction. But when condition subgoals 36 | -- are enabled, the backward direction becomes non-applicable. 37 | example (a b : Nat) (h : b ∣ a) : (a / b) * b = a := by 38 | egg [Nat.div_mul_cancel] 39 | assumption 40 | -------------------------------------------------------------------------------- /Lean/Egg/Tests/Application Order.lean: -------------------------------------------------------------------------------- 1 | import Egg 2 | 3 | example (p q : Nat → Prop) : ((∀ x, p x) ∧ (∀ x, q x)) ↔ ∀ x, p x ∧ q x := by 4 | egg [forall_and] 5 | 6 | example (p q : Nat → Nat → Prop) : ((∀ x, p 1 x) ∧ (∀ x, q 1 x)) ↔ ∀ x, p 1 x ∧ q 1 x := by 7 | egg [forall_and] 8 | 9 | -- TODO: egg can't apply `forall_and`, because `p x 1` and `q x 1` are not syntactially of the form 10 | -- `f x`. Of course we can write them in that form as `(fun a => p a 1) x`. How does `rw` 11 | -- figure this out? Is it through `kabstract`? 12 | -- FIX?: A heuristic to fix this might be: When a bvar appears in a sequence of applications, always 13 | -- abstract it out (as much as possible), so it's as far to the right as possible. This can 14 | -- then be "undone" be egg if needed, by applying β-reduction. 15 | example (p q : Nat → Nat → Prop) : ((∀ x, p x 1) ∧ (∀ x, q x 1)) ↔ ∀ x, p x 1 ∧ q x 1 := by 16 | sorry -- egg [forall_and] 17 | 18 | example (p q : Nat → Nat → Prop) : ((∀ x, p x 1) ∧ (∀ x, q x 1)) ↔ ∀ x, p x 1 ∧ q x 1 := by 19 | egg [@forall_and _ (fun a => p a 1) (fun a => q a 1)] 20 | 21 | example (p q : Nat → Nat → Prop) : ((∀ x, p x 1) ∧ (∀ x, q x 1)) ↔ ∀ x, p x 1 ∧ q x 1 := by 22 | rw [forall_and] 23 | -------------------------------------------------------------------------------- /Lean/Egg/Tests/Propositions.lean: -------------------------------------------------------------------------------- 1 | import Egg 2 | 3 | example : True := by 4 | egg 5 | 6 | example : True ↔ True := by 7 | egg 8 | 9 | example (p q : Prop) (h : p ↔ q) : p ↔ q := by 10 | egg [h] 11 | 12 | example (x : Nat) : (x.add (.succ .zero) = x) ↔ ((Nat.succ .zero).add x = x) := by 13 | have h (x y : Nat) : x.add y = y.add x := Nat.add_comm .. 14 | egg [h] 15 | 16 | example (h : False) : False := by 17 | egg [h] 18 | 19 | example {p q r : Prop} (h₁ : p ∧ q) (h₂ : p ∧ q → r) : r := by 20 | egg [h₁, h₂] 21 | 22 | -- These test were previously used for the legacy `from` construct. 23 | section From 24 | set_option linter.unusedVariables false 25 | 26 | example (h : True) : True := by 27 | egg [h] 28 | 29 | example (h : 0 = 0) : 0 = 0 := by 30 | egg [h] 31 | 32 | example (a b : Nat) (h : a = b) : a = b := by 33 | egg [Nat.add_comm, h] 34 | 35 | example (a b c : Nat) (h₁ : a = b) (h₂ : b = c) : a = c := by 36 | egg [h₁, h₂] 37 | 38 | example (h : p ∧ q ∧ r) : r ∧ r ∧ q ∧ p ∧ q ∧ r ∧ p := by 39 | egg [and_comm, and_self, and_assoc, h] 40 | 41 | example (P : Nat → Prop) (hp : P Nat.zero.succ) (h : ∀ n, P n ↔ P n.succ) : 42 | P Nat.zero.succ.succ.succ.succ := by 43 | egg [h, hp] 44 | 45 | end From 46 | -------------------------------------------------------------------------------- /Rust/Slotted/src/analysis.rs: -------------------------------------------------------------------------------- 1 | use slotted_egraphs::*; 2 | use crate::lean_expr::*; 3 | 4 | #[derive(Default, Clone, PartialEq, Eq)] 5 | pub struct LeanAnalysis { 6 | pub nat_val: Option, 7 | pub is_primitive: bool // A class is primitive if it represents a `Nat`, `Str` or universe level e-node. 8 | } 9 | 10 | impl Analysis for LeanAnalysis { 11 | 12 | fn merge(l: Self, r: Self) -> Self { 13 | Self { 14 | nat_val: l.nat_val.max(r.nat_val), 15 | is_primitive: l.is_primitive && r.is_primitive 16 | } 17 | } 18 | 19 | fn make(_: &EGraph, enode: &LeanExpr) -> Self { 20 | match enode { 21 | LeanExpr::Nat(n) => Self { nat_val: Some(*n), is_primitive: true }, 22 | 23 | LeanExpr::Str(_) | LeanExpr::UVar(_) | LeanExpr::Param(_) | LeanExpr::Succ(_) | 24 | LeanExpr::Max(_, _) | LeanExpr::IMax(_, _) | LeanExpr::Unknown => 25 | Self { is_primitive: true, ..Default::default() }, 26 | 27 | _ => Default::default() 28 | } 29 | } 30 | } 31 | 32 | pub type LeanEGraph = EGraph; 33 | pub type LeanRewrite = Rewrite; -------------------------------------------------------------------------------- /Lean/Egg/Core/Encode/Rules.lean: -------------------------------------------------------------------------------- 1 | import Egg.Core.Encode.Basic 2 | import Egg.Core.Rewrite.Rule 3 | import Lean 4 | open Lean 5 | 6 | namespace Egg.Rewrite 7 | 8 | -- IMPORTANT: The C interface to egg depends on the order of these fields. 9 | structure Rule.Encoded where 10 | name : String 11 | lhs : Expression 12 | rhs : Expression 13 | conds : Array Expression 14 | 15 | nonrec def Condition.encode (c : Rewrite.Condition) (cfg : Config.Encoding) : MetaM Expression := do 16 | let type ← encode c.type cfg 17 | match c.kind with 18 | | .proof => return s!"(proof {type})" 19 | | .tcInst => return s!"(inst {type})" 20 | 21 | nonrec def Rule.encode (rule : Rule) (cfg : Config.Encoding) : MetaM Rule.Encoded := 22 | return { 23 | name := rule.id.description 24 | lhs := ← encode rule.rw.lhs cfg 25 | rhs := ← encode rule.rw.rhs cfg 26 | conds := ← rule.rw.conds.active.mapM (·.encode cfg) 27 | } 28 | 29 | abbrev Rules.Encoded := Array Rule.Encoded 30 | 31 | def Rules.encode (rules : Rules) (cfg : Config.Encoding) (subgoals : Bool) : MetaM Rules.Encoded := 32 | rules.entries.filterMapM fun rule => do 33 | if ← rule.rw.isValid subgoals then 34 | rule.encode cfg 35 | else 36 | return none 37 | -------------------------------------------------------------------------------- /Lean/Egg/Tests/Inst Erasure.lean: -------------------------------------------------------------------------------- 1 | import Egg 2 | 3 | /-- 4 | trace: [egg.encoded] Encoded 5 | [egg.encoded] Goal 6 | [egg.encoded] LHS: (app (app (app (app (app (app (const "HAdd.hAdd" 0 0 0) (const "Nat")) (const "Nat")) (const "Nat")) (inst (app (app (app (const "HAdd" 0 0 0) (const "Nat")) (const "Nat")) (const "Nat")))) (fvar 71)) (fvar 71)) 7 | [egg.encoded] RHS: (app (app (app (app (app (app (const "HAdd.hAdd" 0 0 0) (const "Nat")) (const "Nat")) (const "Nat")) (inst (app (app (app (const "HAdd" 0 0 0) (const "Nat")) (const "Nat")) (const "Nat")))) (fvar 71)) (fvar 71)) 8 | [egg.encoded] Guides 9 | [egg.encoded] (= (app (app (app (app (app (app (const "HAdd.hAdd" 0 0 0) (const "Nat")) (const "Nat")) (const "Nat")) (inst (app (app (app (const "HAdd" 0 0 0) (const "Nat")) (const "Nat")) (const "Nat")))) (fvar 71)) (fvar 71)) (app (app (app (app (app (app (const "HAdd.hAdd" 0 0 0) (const "Nat")) (const "Nat")) (const "Nat")) (inst (app (app (app (const "HAdd" 0 0 0) (const "Nat")) (const "Nat")) (const "Nat")))) (fvar 71)) (fvar 71))) 10 | -/ 11 | #guard_msgs in 12 | set_option trace.egg.encoded true in 13 | set_option egg.builtins false in 14 | set_option egg.tcProjs false in 15 | example (a : Nat) : a + a = a + a := by 16 | egg 17 | -------------------------------------------------------------------------------- /Lean/Egg/Tests/Block Invalid Matches.lean: -------------------------------------------------------------------------------- 1 | import Egg 2 | 3 | -- We have to disable β-reduction as part of normalization, as otherwise `thm₁` and `thm₂` are 4 | -- useless for our testing purposes. Also, we don't want β- or η-reduction affecting these test in 5 | -- any way, so we disable them. 6 | set_option egg.betaReduceRws false 7 | set_option egg.etaReduceRws false 8 | set_option egg.beta false 9 | set_option egg.eta false 10 | 11 | -- These tests cover Condition (1) of (in-)valid matches. 12 | section «Condition 1» 13 | 14 | -- This theorem is only applicable in the backward direction. 15 | theorem thm₁ : ∀ x y : Nat, x = (fun _ => x) y := 16 | fun _ _ => rfl 17 | 18 | /-- error: egg failed to prove the goal (saturated) -/ 19 | #guard_msgs in 20 | example : (fun x => x) = (fun _ : Nat => (fun x => x) 1) := by 21 | egg [thm₁] 22 | 23 | end «Condition 1» 24 | 25 | -- These tests cover Condition (2) of (in-)valid matches. 26 | section «Condition 2» 27 | 28 | theorem thm₂ : ∀ y x : Nat, (fun _ => (fun _ => x) x) y = x := 29 | fun _ _ => rfl 30 | 31 | /-- error: egg failed to prove the goal (saturated) -/ 32 | #guard_msgs in 33 | example : (fun x => (fun a => (fun a => a) a) 0) = (fun x => x) := by 34 | egg [thm₂] 35 | 36 | end «Condition 2» 37 | -------------------------------------------------------------------------------- /Lean/Egg/Tests/Basic.lean: -------------------------------------------------------------------------------- 1 | import Egg 2 | 3 | /-- error: egg failed to prove the goal (saturated) -/ 4 | #guard_msgs in 5 | example : true = false := by 6 | egg 7 | 8 | /-- 9 | error: egg requires arguments to be (proofs of) propositions or (non-propositional) definitions 10 | -/ 11 | #guard_msgs in 12 | example : 0 = 0 := by 13 | egg [true] 14 | 15 | def foo := true 16 | 17 | -- Note: This works because the definition `foo` induces the equation (rewrite) `foo = true`. 18 | example : 0 = 0 := by 19 | egg [foo] 20 | 21 | example : 0 = 0 := by 22 | egg 23 | 24 | example (a : Nat) : a = a := by 25 | egg 26 | 27 | example (a b : Nat) (h : a = b) : a = b := by 28 | egg [h] 29 | 30 | example (a b c : Nat) (h₁ : a = b) (h₂ : b = c) : a = c := by 31 | egg [h₁, h₂] 32 | 33 | example (a b : Nat) : a + b = b + a := by 34 | egg [Nat.add_comm] 35 | 36 | example (a b c : Nat) : (a + b) + c = (c + b) + a := by 37 | egg [Nat.add_comm, Nat.add_assoc] 38 | 39 | example (a : Nat) (h : ∀ x : Nat, x + 1 = 1 + x) : a + 1 = 1 + a := by 40 | egg [h] 41 | 42 | def f : Nat → Nat 43 | | .zero => .zero 44 | | .succ n => f n 45 | 46 | def g : Nat → Nat 47 | | .zero => .zero 48 | | .succ n => g n 49 | 50 | example : f (g Nat.zero.succ.succ) = .zero := by 51 | egg [f, g] 52 | -------------------------------------------------------------------------------- /C/rev_ffi.c: -------------------------------------------------------------------------------- 1 | #include 2 | #include 3 | #include "ffi.h" 4 | 5 | // Returns `EStateM.Result Exception PUnit Bool`. 6 | extern lean_obj_res lean_is_synthable( 7 | lean_object* ty_str, 8 | lean_object* x1, 9 | lean_object* x2, 10 | lean_object* x3, 11 | lean_object* x4, 12 | lean_object* x5 13 | ); 14 | 15 | _Bool is_synthable(const void* v, const char* str) { 16 | env* e = (env*)v; 17 | lean_obj_res ty_str = lean_mk_string(str); 18 | 19 | lean_inc(e->x1); lean_inc(e->x2); lean_inc(e->x3); lean_inc(e->x4); lean_inc(e->x5); 20 | lean_obj_res s = lean_is_synthable(ty_str, e->x1, e->x2, e->x3, e->x4, e->x5); 21 | 22 | switch (s->m_tag) { 23 | /* EStateM.Result.ok */ case 0: { 24 | // Note: The `Bool` is boxed for some reason. 25 | lean_obj_res b = lean_ctor_get(s, 0); 26 | size_t result = lean_unbox(b); 27 | return result != 0; 28 | } 29 | /* EStateM.Result.err */ case 1: { 30 | return false; 31 | // TODO: We should propagate the error back, as it would also be important 32 | // to know whether we have a `MetaM` failure when we return from `runRaw`. 33 | } 34 | default: { 35 | exit(1); 36 | } 37 | } 38 | } -------------------------------------------------------------------------------- /Lean/Egg/Tests/Cast.lean: -------------------------------------------------------------------------------- 1 | import Egg 2 | 3 | namespace Simple 4 | 5 | class C (α) extends NatCast α, Add α 6 | 7 | variable (cast_add : ∀ {R} [C R] (m n : Nat), ((m + n : Nat) : R) = m + n) 8 | 9 | example (a b : Nat) [C R] : Nat.cast (a + b) = (a : R) + b := by 10 | egg [cast_add] 11 | 12 | instance : C Int where 13 | 14 | example (a b : Nat) : Nat.cast (a + b) = (a : Int) + b := by 15 | egg [cast_add] 16 | 17 | end Simple 18 | 19 | namespace NoDiamond 20 | 21 | class N (R) extends NatCast R, Add R where 22 | 23 | variable (cast_add : ∀ {R : Type _} [N R] (m n : Nat), ((m + n : Nat) : R) = m + n) 24 | 25 | example (a b : Nat) [N R] : Nat.cast (a + b) = (a : R) + b := by 26 | egg [cast_add] 27 | 28 | instance : N Int where 29 | 30 | example (a b : Nat) : Nat.cast (a + b) = (a : Int) + b := by 31 | egg [cast_add] 32 | 33 | end NoDiamond 34 | 35 | namespace Diamond 36 | 37 | class A (α) extends Add α where 38 | class B (α) extends Add α where 39 | class N (R) extends NatCast R, A R, B R where 40 | 41 | variable (cast_add : ∀ {R} [N R] (m n : Nat), ((m + n : Nat) : R) = m + n) 42 | 43 | example (a b : Nat) [N R] : Nat.cast (a + b) = (a : R) + b := by 44 | egg [cast_add] 45 | 46 | instance : N Int where 47 | 48 | example (a b : Nat) : Nat.cast (a + b) = (a : Int) + b := by 49 | egg [cast_add] 50 | 51 | end Diamond 52 | -------------------------------------------------------------------------------- /Lean/Egg/Tests/Math/Math/Mathlib/Algebra/Group/Basic.lean: -------------------------------------------------------------------------------- 1 | import Math.Mathlib.Algebra.Group.Defs 2 | 3 | -- NOTE: Starting around line 110. 4 | 5 | section CommSemigroup 6 | 7 | variable [CommSemigroup G] 8 | 9 | example (a b c : G) : a * (b * c) = b * (a * c) := by 10 | egg +comm_semigroup 11 | 12 | /- Previous -/ attribute [egg comm_semigroup] mul_left_comm 13 | 14 | example (a b c : G) : a * b * c = a * c * b := by 15 | egg +comm_semigroup 16 | 17 | /- Previous -/ attribute [egg comm_semigroup] mul_right_comm 18 | 19 | example (a b c d : G) : a * b * (c * d) = a * c * (b * d) := by 20 | egg +comm_semigroup 21 | 22 | /- Previous -/ attribute [egg comm_semigroup] mul_mul_mul_comm 23 | 24 | example (a b c : G) : a * b * c = b * c * a := by 25 | egg +comm_semigroup 26 | 27 | /- Previous -/ attribute [egg comm_semigroup] mul_rotate 28 | 29 | example (a b c : G) : a * (b * c) = b * (c * a) := by 30 | egg +comm_semigroup 31 | 32 | /- Previous -/ attribute [egg comm_semigroup] mul_rotate' 33 | 34 | end CommSemigroup 35 | 36 | -- NOTE: Skipped monoid power theorems around lines 130-200. 37 | 38 | section CommMonoid 39 | 40 | example (hy : x * y = 1) (hz : x * z = 1) : y = z := by 41 | sorry -- TODO: This used to work: egg +comm_monoid [*] 42 | 43 | /- Previous -/ attribute [egg comm_monoid] inv_unique 44 | 45 | end CommMonoid 46 | 47 | -- NOTE: Skipped lines 210+. 48 | -------------------------------------------------------------------------------- /Lean/Egg/Tests/Activations.lean: -------------------------------------------------------------------------------- 1 | import Egg 2 | 3 | set_option trace.egg.activations true 4 | set_option egg.builtins false 5 | set_option linter.unusedVariables false 6 | 7 | /-- 8 | trace: [egg.activations] Activations 9 | [egg.activations] nat-lit: false 10 | level: false 11 | lambda: false 12 | forall: false 13 | -/ 14 | #guard_msgs in 15 | example : True := by 16 | egg 17 | 18 | /-- 19 | trace: [egg.activations] Activations 20 | [egg.activations] nat-lit: true 21 | level: false 22 | lambda: false 23 | forall: false 24 | -/ 25 | #guard_msgs in 26 | example : 0 = 0 := by 27 | egg 28 | 29 | /-- 30 | trace: [egg.activations] Activations 31 | [egg.activations] nat-lit: false 32 | level: false 33 | lambda: true 34 | forall: false 35 | -/ 36 | #guard_msgs in 37 | example : (fun x : Bool => x) true = id true := by 38 | egg [id] 39 | 40 | /-- 41 | trace: [egg.activations] Activations 42 | [egg.activations] nat-lit: false 43 | level: false 44 | lambda: false 45 | forall: true 46 | -/ 47 | #guard_msgs in 48 | example (h : ∀ x : Nat, x = x) : True := by 49 | egg [h] 50 | 51 | /-- 52 | trace: [egg.activations] Activations 53 | [egg.activations] nat-lit: false 54 | level: true 55 | lambda: false 56 | forall: false 57 | -/ 58 | #guard_msgs in 59 | example (x : α × β) : id x = x := by 60 | egg [id] 61 | -------------------------------------------------------------------------------- /Lean/Egg/Tests/Reduce.lean: -------------------------------------------------------------------------------- 1 | import Egg 2 | 3 | example : true = true := by 4 | egg 5 | 6 | -- Checks that we try to introduce assumptions until the goal is an equality. 7 | example : True → true = true := by 8 | egg 9 | 10 | -- Checks that we try to introduce assumptions until the goal is an equality. 11 | example : True → False → true = true := by 12 | egg 13 | 14 | abbrev P := ∀ x y : Nat, x + y = x + y 15 | 16 | -- Checks that we try to introduce assumptions until the goal is an equality. 17 | example : P := by 18 | egg 19 | 20 | -- Checks that we do not unfold semi-reducible (or higher) definitions. 21 | 22 | set_option allowUnsafeReducibility true in attribute [semireducible] P 23 | 24 | /-- error: egg failed to prove the goal (saturated) -/ 25 | #guard_msgs in 26 | example : P := by 27 | egg 28 | 29 | abbrev Q := ∀ x : Nat, Nat.succ x = nat_lit 0 30 | 31 | -- Checks that we can "see through" definitions in rewrites where the body is an equality. 32 | example (h : Q) : 1 = 0 := by 33 | egg [h] 34 | 35 | abbrev R := true = false 36 | 37 | -- Checks that we can "see through" definitions in rewrites where the body needs unfolding. 38 | example (h : R) : true = false := by 39 | egg [h] 40 | 41 | abbrev S := ∀ _ : 0 = 0, R 42 | 43 | -- Checks that we can "see through" definitions in rewrites where the body is not an equality. 44 | example (h : S) : R := by 45 | egg [h] 46 | -------------------------------------------------------------------------------- /Lean/Egg/Tests/Blocking.lean: -------------------------------------------------------------------------------- 1 | import Egg 2 | 3 | /-- warning: conditions are only blocked when `egg.subgoals` is `true` -/ 4 | #guard_msgs in 5 | example : True := by 6 | egg blocking True 7 | 8 | /-- error: blocking terms must be propositions -/ 9 | #guard_msgs in 10 | example : True := by 11 | egg blocking 0 12 | 13 | set_option egg.subgoals true 14 | 15 | example : True := by 16 | egg blocking True 17 | 18 | variable {P : Prop} 19 | 20 | /-- 21 | error: unsolved goals 22 | P : Prop 23 | h : P → False 24 | ⊢ P 25 | -/ 26 | #guard_msgs in 27 | example (h : P → False) : False := by 28 | egg [h] 29 | 30 | /-- error: egg failed to prove the goal (saturated) -/ 31 | #guard_msgs in 32 | example (h : P → False) : False := by 33 | egg [h] blocking P 34 | 35 | /-- 36 | error: unsolved goals 37 | P Q : Prop 38 | hq : Q 39 | h₁ : P → False 40 | h₂ : Q → False 41 | ⊢ P 42 | -/ 43 | #guard_msgs in 44 | example {Q : Prop} (hq : Q) (h₁ : P → False) (h₂ : Q → False) : False := by 45 | egg [h₁, h₂] 46 | 47 | example {Q : Prop} (hq : Q) (_h₁ : P → False) (h₂ : Q → False) : False := by 48 | egg [_h₁, h₂] blocking P 49 | exact hq 50 | 51 | example (h : ∀ n, n > 0 → P) : P := by 52 | egg [h] 53 | · exact 1 54 | · simp 55 | 56 | example (h : ∀ n, n > 1 → n > 0) : 2 > 0 := by 57 | egg [h] 58 | simp 59 | 60 | example {f : Nat → Nat} (h : ∀ n, n > 2 → f n = x) : f 3 = x := by 61 | egg [h] 62 | simp 63 | -------------------------------------------------------------------------------- /Lean/Egg/Tests/Tags.lean: -------------------------------------------------------------------------------- 1 | import Egg 2 | 3 | class Group (α) extends Mul α, One α, Inv α where 4 | mul_assoc (a b c : α) : (a * b) * c = a * (b * c) 5 | one_mul (a : α) : 1 * a = a 6 | mul_one (a : α) : a * 1 = a 7 | inv_mul_self (a : α) : a⁻¹ * a = 1 8 | mul_inv_self (a : α) : a * a⁻¹ = 1 9 | 10 | variable [Group α] (a b x y : α) 11 | 12 | attribute [egg group] Group.mul_assoc 13 | attribute [egg group] Group.one_mul 14 | attribute [egg group] Group.mul_one 15 | attribute [egg group] Group.inv_mul_self 16 | attribute [egg group] Group.mul_inv_self 17 | 18 | /-- error: egg failed to prove the goal (saturated) -/ 19 | #guard_msgs in 20 | example : a⁻¹ * (a * b) = b := by egg 21 | 22 | /-- info: Group.inv_mul_self, Group.mul_assoc, Group.mul_inv_self, Group.mul_one, Group.one_mul -/ 23 | #guard_msgs in #egg_basket group 24 | 25 | @[egg group] 26 | theorem inv_mul_cancel_left : a⁻¹ * (a * b) = b := by egg +group 27 | 28 | /-- 29 | info: inv_mul_cancel_left, Group.inv_mul_self, Group.mul_assoc, Group.mul_inv_self, Group.mul_one, Group.one_mul 30 | -/ 31 | #guard_msgs in #egg_basket group 32 | 33 | @[egg group] 34 | theorem mul_inv_cancel_left : a * (a⁻¹ * b) = b := by egg +group 35 | 36 | /-- 37 | info: inv_mul_cancel_left, mul_inv_cancel_left, Group.inv_mul_self, Group.mul_assoc, Group.mul_inv_self, Group.mul_one, Group.one_mul 38 | -/ 39 | #guard_msgs in #egg_basket group 40 | -------------------------------------------------------------------------------- /Lean/Egg/Tests/Star.lean: -------------------------------------------------------------------------------- 1 | import Egg 2 | 3 | -- Tests for the `*` argument. 4 | 5 | /-- error: egg failed to prove the goal (saturated) -/ 6 | #guard_msgs in 7 | example : true = false := by 8 | egg [*] 9 | 10 | /-- error: egg failed to prove the goal (saturated) -/ 11 | #guard_msgs in 12 | example (_ : true = true) : [true] ++ [] = [true] := by 13 | egg [*] 14 | 15 | /-- error: duplicate '*' -/ 16 | #guard_msgs in 17 | example : 0 = 0 := by 18 | egg [*, *] 19 | 20 | /-- error: duplicate '*' -/ 21 | #guard_msgs in 22 | example (h : 0 = 0) : 0 = 0 := by 23 | egg [*, h, *] 24 | 25 | example : 0 = 0 := by 26 | egg [*] 27 | 28 | example (a : Nat) : a = a := by 29 | egg [*] 30 | 31 | example (a b : Nat) (_ : a = b) : a = b := by 32 | egg [*] 33 | 34 | example (a b c : Nat) (_ : a = b) (_ : b = c) : a = c := by 35 | egg [*] 36 | 37 | example (a b : Nat) : a + b = b + a := by 38 | have := Nat.add_comm 39 | egg [*] 40 | 41 | example (a b c : Nat) : (a + b) + c = (c + b) + a := by 42 | have := Nat.add_comm 43 | have := Nat.add_assoc 44 | egg [*] 45 | 46 | example (a b c : Nat) : (a + b) + c = (c + b) + a := by 47 | have := Nat.add_assoc 48 | egg [*, Nat.add_comm] 49 | 50 | example (a b c : Nat) : (a + b) + c = (c + b) + a := by 51 | have := Nat.add_assoc 52 | egg [Nat.add_comm, *] 53 | 54 | example (a b c : Nat) : (a + b) + c = (c + b) + a := by 55 | egg [Nat.add_comm, *, Nat.add_assoc] 56 | -------------------------------------------------------------------------------- /Rust/Egg/src/beta.rs: -------------------------------------------------------------------------------- 1 | use egg::*; 2 | use crate::analysis::*; 3 | use crate::lean_expr::*; 4 | 5 | pub fn beta_reduction_rw() -> LeanRewrite { 6 | rewrite!("≡β"; "(app (λ ?t ?b) ?a)" => { Beta { body : "?b".parse().unwrap(), arg : "?a".parse().unwrap() }}) 7 | } 8 | 9 | struct Beta { 10 | body: Var, 11 | arg: Var 12 | } 13 | 14 | impl Applier for Beta { 15 | 16 | fn apply_one(&self, graph: &mut LeanEGraph, _: Id, subst: &Subst, ast: Option<&PatternAst>, rule: Symbol) -> Vec { 17 | let arg_bvars = &graph[subst[self.arg]].data.loose_bvars; 18 | let body_bvars = &graph[subst[self.body]].data.loose_bvars; 19 | 20 | let shifted_arg: PatternAst = if arg_bvars.is_empty() { 21 | format!("{}", self.arg).parse().unwrap() 22 | } else { 23 | format!("(↑ + 1 0 {})", self.arg).parse().unwrap() 24 | }; 25 | 26 | let sub: PatternAst = format!("(↦ 0 {} {})", shifted_arg, self.body).parse().unwrap(); 27 | 28 | let beta: PatternAst = if !arg_bvars.is_empty() || body_bvars.iter().any(|b| *b != 0) { 29 | format!("(↑ - 1 0 {})", sub).parse().unwrap() 30 | } else { 31 | sub 32 | }; 33 | 34 | let (id, did_union) = graph.union_instantiations(ast.unwrap(), &beta, subst, rule); 35 | if did_union { vec![id] } else { vec![] } 36 | } 37 | } -------------------------------------------------------------------------------- /Lean/Egg/Core/Gen/Builtins.lean: -------------------------------------------------------------------------------- 1 | import Egg.Core.Rewrite.Rule 2 | import Lean 3 | 4 | open Lean Meta 5 | 6 | namespace Egg.Rewrite.Rules 7 | 8 | theorem imp_mp {p q : Prop} (imp : p → q) (h : p) : q := 9 | imp h 10 | 11 | -- TODO: `Nat.succ_eq_add_one` causes nat-lit rewrites to always be active. 12 | -- 13 | -- We exclude `imp_mp` when subgoals are enabled, as this theorem proves everything in that case. 14 | private def builtinTheorems (subgoals : Bool) : Array Name := Id.run do 15 | let mut thms := #[``Nat.succ_eq_add_one, ``ge_iff_le, ``gt_iff_lt] 16 | unless subgoals do thms := thms.push ``imp_mp 17 | return thms 18 | 19 | def builtins (cfg : Config.Normalization) (subgoals : Bool) : MetaM Rules := do 20 | let mut rules := ∅ 21 | let env ← getEnv 22 | let thms := builtinTheorems subgoals 23 | for thm in thms, idx in [:thms.size] do 24 | let info := env.find? thm |>.get! 25 | let lvlMVars ← List.replicateM info.numLevelParams mkFreshLevelMVar 26 | let val := info.instantiateValueLevelParams! lvlMVars 27 | let type := info.instantiateTypeLevelParams lvlMVars 28 | let some rules' ← rules.add? (.builtin idx) val type cfg .both (mkIdent thm) 29 | | throwError "egg failed to create rewrites for builtin '{thm}'" 30 | rules := rules' 31 | return rules 32 | 33 | end Rewrite.Rules 34 | 35 | def genBuiltins (cfg : Config) : MetaM Rewrite.Rules := do 36 | Rewrite.Rules.builtins cfg cfg.subgoals 37 | -------------------------------------------------------------------------------- /Lean/Egg/Tests/Calcify.lean: -------------------------------------------------------------------------------- 1 | import Egg 2 | 3 | /-- 4 | info: Try this: calc 5 | true 6 | _ = true := Eq.refl true 7 | -/ 8 | #guard_msgs in 9 | example : true = true := by 10 | egg? 11 | 12 | /-- 13 | info: Try this: calc 14 | a 15 | _ = b := h₁ 16 | _ = c := h₂ 17 | -/ 18 | #guard_msgs in 19 | example (a b c : Nat) (h₁ : a = b) (h₂ : b = c) : a = c := by 20 | egg? [h₁, h₂] 21 | 22 | variable (f : Nat → Nat → Nat) 23 | 24 | /-- 25 | info: Try this: calc 26 | f 1 2 27 | _ = f 2 1 := (h 2 1).symm 28 | -/ 29 | #guard_msgs in 30 | example (h : ∀ x y : Nat, f x y = f y x) : f 1 2 = f 2 1 := by 31 | egg? [h] 32 | 33 | /-- 34 | info: Try this: calc 35 | f a a 36 | _ = f b a := h a b 37 | -/ 38 | #guard_msgs in 39 | example (a b : Nat) (h : ∀ x y : Nat, f x x = f y x) : f a a = f b a := by 40 | egg? [h] 41 | 42 | -- This used to fail because we didn't correctly reduce and introduce variables before calling 43 | -- calcify. 44 | /-- 45 | info: Try this: ⏎ 46 | intro a b 47 | calc 48 | f a a 49 | _ = f b a := h a b 50 | -/ 51 | #guard_msgs in 52 | example (h : ∀ x y : Nat, f x x = f y x) : ∀ a b : Nat, f a a = f b a := by 53 | egg? [h] 54 | 55 | 56 | abbrev E := ∀ a b : Nat, f a a = f b a 57 | 58 | /-- 59 | info: Try this: ⏎ 60 | intro a b 61 | calc 62 | f a a 63 | _ = f b a := h a b 64 | -/ 65 | #guard_msgs in 66 | example (h : ∀ x y : Nat, f x x = f y x) : E f := by 67 | egg? [h] 68 | -------------------------------------------------------------------------------- /Lean/Egg/Tests/Groups.lean: -------------------------------------------------------------------------------- 1 | import Egg 2 | 3 | class Group (α) extends One α, Inv α, Mul α where 4 | mul_assoc (a b c : α) : (a * b) * c = a * (b * c) 5 | one_mul (a : α) : 1 * a = a 6 | mul_one (a : α) : a * 1 = a 7 | inv_mul_self (a : α) : a⁻¹ * a = 1 8 | mul_inv_self (a : α) : a * a⁻¹ = 1 9 | 10 | variable [Group G] {a b : G} 11 | 12 | open Group Egg.Guides Egg.Config.Modifier in 13 | macro "group" mod:egg_cfg_mod guides:(egg_guides)? : tactic => `(tactic| 14 | egg $mod:egg_cfg_mod [mul_assoc, one_mul, mul_one, inv_mul_self, mul_inv_self] $[$guides]? 15 | ) 16 | 17 | theorem inv_mul_cancel_left : a⁻¹ * (a * b) = b := by 18 | group 19 | 20 | theorem mul_inv_cancel_left : a * (a⁻¹ * b) = b := by 21 | group 22 | 23 | theorem inv_one : (1 : G)⁻¹ = 1 := by 24 | group using 1 * (1 : G)⁻¹ 25 | 26 | theorem inv_mul : (a * b)⁻¹ = b⁻¹ * a⁻¹ := by 27 | calc _ = b⁻¹ * a⁻¹ * (a * b) * (a * b)⁻¹ := by group 28 | _ = _ := by group 29 | 30 | theorem inv_inv : a⁻¹⁻¹ = a := by 31 | calc _ = a⁻¹⁻¹ * (a⁻¹ * a) := by group 32 | _ = _ := by group 33 | 34 | theorem inv_mul' : (a * b)⁻¹ = b⁻¹ * a⁻¹ := by 35 | group using b⁻¹ * a⁻¹ * (a * b) * (a * b)⁻¹ 36 | 37 | -- NOTE: The explanation is too long. 38 | -- theorem inv_mul'' : (a * b)⁻¹ = b⁻¹ * a⁻¹ := by 39 | -- group using a⁻¹ * (a * b) * (a * b)⁻¹ 40 | 41 | theorem inv_inv' : a⁻¹⁻¹ = a := by 42 | group using a⁻¹⁻¹ * a⁻¹ * a 43 | -------------------------------------------------------------------------------- /Lean/Egg/Tests/Equality Conditions.lean: -------------------------------------------------------------------------------- 1 | import Egg 2 | 3 | example (h : 0 = 0 → 1 = 2) : 1 = 2 := by 4 | egg [h] 5 | 6 | example (h : (0 = (fun x => x) 0) → 1 = 2) : 1 = 2 := by 7 | egg [h] 8 | 9 | variable {x : Nat} {f : Nat → Nat} 10 | 11 | example (h₁ : x = y) (h₂ : x = y → 1 = 2) : 1 = 2 := by 12 | egg [h₁, h₂] 13 | 14 | example (h₁ : x = y) (h₂ : x = y → 1 = 2) : 1 = 2 := by 15 | egg [h₁, h₂] 16 | 17 | example (h₁ : x = y) (h₂ : y = z) (h₃ : x = z → 1 = 2) : 1 = 2 := by 18 | egg [h₁, h₂, h₃] 19 | 20 | -- Note: The fact that this works is kind of accidental. By default, we shouldn't expect this to 21 | -- work, because the goal doesn't introduce `x` into the e-graph, so we shouldn't be able to 22 | -- establish `f x = x` by rewriting with `h₁`. But when the applier for `h₂` is run, in order 23 | -- to check the precondition, we need to find the e-classes of `f x` and `x`. To do this, we 24 | -- use `add_instantiation` which also adds `f x` and `x` to the e-graph (in fact, we also add 25 | -- the term `f x = x` to the e-graph). Thus, after `h₂` fails, we can actually apply `h₁` 26 | -- which then enables `h₂` to succeed. 27 | example (h₁ : ∀ a : Nat, f a = a) (h₂ : f x = x → 1 = 2) : 1 = 2 := by 28 | egg [h₁, h₂] 29 | 30 | example (h₁ : ∀ a : Nat, f a = a) (h₂ : p ∧ q) (h₃ : (f x = x) → (p ∧ q) → 1 = 2) : 1 = 2 := by 31 | egg [h₁, h₂, h₃] 32 | 33 | example (h₁ : ∀ a : Nat, f a = a) (h₂ : p ∧ q) (h₃ : p ∧ q ↔ r) (h₄ : (f x = x) → r → 1 = 2) : 34 | 1 = 2 := by 35 | egg [h₁, h₂, h₃, h₄] 36 | -------------------------------------------------------------------------------- /Lean/Egg/Core/Rewrite/Prerewrite.lean: -------------------------------------------------------------------------------- 1 | import Egg.Core.Direction 2 | import Egg.Core.Congr 3 | open Lean Meta 4 | 5 | namespace Egg 6 | 7 | -- Note: We don't create `Prerewrite`s directly, but use `Prerewrite.from?` instead. 8 | structure Prerewrite extends Congr where private mk :: 9 | qvars : MVarIdSet 10 | proof : Expr 11 | deriving Inhabited 12 | 13 | def Prerewrite.from? (proof type : Expr) (cfg : Config.Normalization) (normalize : Bool) : 14 | MetaM (Option Prerewrite) := do 15 | let type ← if normalize then Egg.normalize type cfg else pure type 16 | let (args, _, prop) ← withReducible do forallMetaTelescopeReducing type 17 | let proof := mkAppN proof args 18 | let qvars : MVarIdSet := args.foldl (·.insert ·.mvarId!) ∅ 19 | if let some cgr ← Congr.from? prop then 20 | return some { cgr with proof, qvars } 21 | -- Note: We need this to reduce abbrevs which don't unfold to `∀ ...`, but rather just `_ ~ _`. 22 | else if let some cgr ← Congr.from? (← withReducible do reduce (skipTypes := false) prop) then 23 | return some { cgr with proof, qvars } 24 | else if (← inferType prop).isProp then 25 | let cgr : Congr := { rel := .eq, lhs := prop, rhs := .const ``True [] } 26 | let proof ← mkEqTrue proof 27 | return some { cgr with proof, qvars } 28 | else 29 | return none 30 | 31 | def Prerewrite.forDir (pre : Prerewrite) : Direction → MetaM Prerewrite 32 | | .forward => return pre 33 | | .backward => return { pre with 34 | lhs := pre.rhs, 35 | rhs := pre.lhs, 36 | proof := ← pre.rel.mkSymm pre.proof 37 | } 38 | -------------------------------------------------------------------------------- /Lean/Egg/Tests/Pattern MVar.lean: -------------------------------------------------------------------------------- 1 | import Egg 2 | 3 | variable (h : ∀ (p : Nat → Nat) (x : Nat), p x = p (x + 0)) 4 | 5 | set_option egg.tcProjs false in 6 | example (f : Nat → Nat → Nat) : (f 1) x = (f 1) (x + 0) := by 7 | egg [h] 8 | 9 | example (f : Nat → Nat → Nat) : (f (nat_lit 1)) x = (f 1) (x + 0) := by 10 | egg [h] 11 | 12 | example (f : Nat → Nat → Nat) : (f 1) x = (f (nat_lit 1)) (x + 0) := by 13 | egg [h] 14 | 15 | example (f : Nat → Nat → Nat) : (f 1) x = (f 1) (x + (nat_lit 0)) := by 16 | egg [h] 17 | 18 | -- Note: This used to fail with a broken explanation, because `?m.2213 : Nat → Nat` was matched 19 | -- against `HAdd.hAdd : Nat → Nat → Nat`. That is, the types didn't match. 20 | -- This suddenly started working as part of the `better-facts` PR. 21 | set_option egg.retryWithShapes false in 22 | example (f : Nat → Nat → Nat) : (f 1) x = (f 1) (x + 0) := by 23 | egg [h] 24 | 25 | set_option egg.shapes true in 26 | example (f : Nat → Nat → Nat) : (f 1) x = (f 1) (x + 0) := by 27 | egg [h] 28 | 29 | /-- 30 | error: Invalid rewrite argument: The pattern to be substituted is a metavariable (`?p ?x`) in this equality 31 | ?p ?x = ?p (?x + 0) 32 | -/ 33 | #guard_msgs in 34 | example (f : Nat → Nat → Nat) : (f 1) x = (f 1) (x + 0) := by 35 | rw [h] 36 | 37 | /-- 38 | error: Tactic `simp` failed with a nested error: 39 | maximum recursion depth has been reached 40 | use `set_option maxRecDepth ` to increase limit 41 | use `set_option diagnostics true` to get diagnostic information 42 | -/ 43 | #guard_msgs in 44 | example (f : Nat → Nat → Nat) : (f 1) x = (f 1) (x + 0) := by 45 | simp [h] 46 | -------------------------------------------------------------------------------- /Lean/Egg/Tests/WIP ValidateTagged.lean: -------------------------------------------------------------------------------- 1 | import Egg 2 | 3 | class Group (α) extends Mul α, One α, Inv α where 4 | mul_assoc (a b c : α) : (a * b) * c = a * (b * c) 5 | one_mul (a : α) : 1 * a = a 6 | mul_one (a : α) : a * 1 = a 7 | inv_mul_self (a : α) : a⁻¹ * a = 1 8 | mul_inv_self (a : α) : a * a⁻¹ = 1 9 | 10 | variable [Group α] (a b x y : α) 11 | 12 | attribute [egg group] Group.mul_assoc 13 | attribute [egg group] Group.one_mul 14 | attribute [egg group] Group.mul_one 15 | attribute [egg group] Group.inv_mul_self 16 | attribute [egg group] Group.mul_inv_self 17 | 18 | def hPow : α → Nat → α 19 | | _, 0 => 1 20 | | a, (n+1) => a * hPow a n 21 | 22 | instance [Group α] : HPow α Nat α := ⟨hPow⟩ 23 | 24 | def OrderN (n : Nat) (a : α) : Prop := a^n = 1 25 | 26 | -- not defining the cardinality here for space reasons 27 | def card (α) [Group α] : Nat := sorry 28 | 29 | -- This one should not go through! not an equality 30 | @[egg order] 31 | theorem ex_min_order : ∃ n : Nat, OrderN n a ∧ (∀ n', n' < n → ¬ OrderN n a) := sorry 32 | 33 | -- This should also be recognized as an equality 34 | @[egg order] 35 | theorem card_order : OrderN (card α) a := by 36 | sorry 37 | 38 | def Abelian (α) [Group α] : Prop := ∀ a b : α, a * b = b * a 39 | 40 | def commutator := a*b*a⁻¹*b⁻¹ 41 | 42 | -- Ideally, egg can see through this prop that there's an equality? 43 | @[egg abel] 44 | theorem all_commutators_trivial_abelian : (∀ a b : α, commutator a b = 1) → Abelian α := by sorry 45 | 46 | /-- We should not break egg after adding these lemmas -/ 47 | example : a * b = b * a := by 48 | egg 49 | -------------------------------------------------------------------------------- /Lean/Egg/Core/Normalize.lean: -------------------------------------------------------------------------------- 1 | import Egg.Core.Config 2 | import Lean 3 | open Lean Meta Core 4 | 5 | namespace Egg 6 | 7 | -- Instantiates mvars, performs ζ-, converts `Expr.proj`s to `Expr.app`s, and removes `Expr.mdata`s. 8 | -- Depending on the `cfg`, also performs β- and η-reduction and reduction of internalized natural 9 | -- number operations. 10 | partial def normalize (e : Expr) (cfg : Config.Normalization) : MetaM Expr := do 11 | go (← instantiateMVars e) 12 | where 13 | go : Expr → MetaM Expr 14 | | .mdata _ e => go e 15 | | .app fn arg => do 16 | let mut app := .app (← go fn) (← go arg) 17 | if cfg.betaReduceRws then app ← betaReduce app 18 | if cfg.natReduceRws then app ← natReduce app 19 | return app 20 | | .lam n ty b i => do 21 | withLocalDecl n i (← go ty) fun fvar => do 22 | let body ← mkLambdaFVars #[fvar] (← go <| b.instantiate1 fvar) 23 | return if cfg.etaReduceRws then body.eta else body 24 | | .forallE n ty b i => do 25 | withLocalDecl n i (← go ty) fun fvar => do 26 | mkForallFVars #[fvar] (← go <| b.instantiate1 fvar) 27 | | e@(.letE ..) => do go (← zetaReduce e) 28 | | .proj ty idx b => do go (← expandProj ty idx b) 29 | | e => return e 30 | 31 | expandProj (ty : Name) (idx : Nat) (b : Expr) : MetaM Expr := do 32 | let some field := (getStructureFields (← getEnv) ty)[idx]? 33 | | throwError "'Egg.normalize' failed to reduce proj" 34 | mkProjection b field 35 | 36 | natReduce (e : Expr) : MetaM Expr := do 37 | if let some n ← evalNat e 38 | then return mkNatLit n 39 | else return e 40 | -------------------------------------------------------------------------------- /Lean/Egg/Tests/Shift Loop.lean: -------------------------------------------------------------------------------- 1 | import Egg 2 | 3 | -- This used to get stuck because of an integer underflow ocurring in the e-class analysis for loose 4 | -- bound variables, as a result of down shifts. I think this happened because of a shift node with a 5 | -- self-loop. I think there are multiple problems with looping shift-nodes. First, they can generate 6 | -- infinitely many new nodes. E.g. if we have a +2 shift node with a self-loop and another e-node in 7 | -- the same e-class which contains bvar 5, then we get copies of that e-node with bvar 7, 9, 11, ... 8 | -- The second problem is analogous and occurs when constructing the loose bvar e-class analysis. My 9 | -- assumption is that e-class analysis values for nodes with self-loops are constructed by 10 | -- performing all of the potentially recursive calls to nested e-class analysis values until a fixed 11 | -- point is reached. For shift nodes with self-loops there is no fixed point. 12 | -- We address the latter problem by adding an arbitrary limit of 100 for cutting off loose bvar 13 | -- e-class analyses. 14 | -- 15 | -- I used to following workaround for the lack of reliable printing to stdout from Rust: 16 | -- 17 | -- use std::fs::OpenOptions; 18 | -- use std::io::prelude::*; 19 | -- let mut file = OpenOptions::new().write(true).append(true).open("").unwrap(); 20 | -- let _ = writeln!(file, "apply {}", rule); 21 | 22 | variable {r : α → α → Prop} (s : α → Prop) 23 | 24 | /-- error: egg failed to prove the goal (reached time limit) -/ 25 | #guard_msgs in 26 | open Classical in 27 | example : ¬(∃ a, ∀ b, s b → r b a) ↔ (∀ a, ∃ b, s b → ¬r b a) := by 28 | egg [not_exists, exists_prop, not_and] 29 | -------------------------------------------------------------------------------- /Lean/Egg/Core/Gen/Guides.lean: -------------------------------------------------------------------------------- 1 | import Egg.Core.Rewrite.Basic 2 | import Egg.Core.Guides 3 | import Lean 4 | 5 | open Lean Meta 6 | 7 | namespace Egg 8 | 9 | -- Note: Mvars which are not at the current mctx depth are considered closed. 10 | private partial def closedSubexprs (expr : Expr) : MetaM ExprSet := do 11 | if expr.hasLooseBVars then 12 | closedChildSubexprs expr 13 | else if expr.hasMVar || expr.hasLevelMVar then 14 | let { result := mvars, .. } := expr.collectMVars {} 15 | unless ← mvars.allM fun m => return !(← m.isAssignable) do return (← closedChildSubexprs expr) 16 | let { result := lmvars, .. } := collectLevelMVars {} expr 17 | unless ← lmvars.allM fun m => return !(← isLevelMVarAssignable m) do return (← closedChildSubexprs expr) 18 | return {expr} 19 | else 20 | return {expr} 21 | where 22 | closedChildSubexprs (expr : Expr) : MetaM ExprSet := 23 | expr.foldlM (init := ∅) fun acc child => 24 | return acc.union (← closedSubexprs child) 25 | 26 | private def deriveGuides (rw : Rewrite) : MetaM ExprSet := do 27 | let mut result := ∅ 28 | result := result.union (← closedSubexprs rw.lhs) 29 | result := result.union (← closedSubexprs rw.rhs) 30 | for cond in rw.conds.active do result := result.union (← closedSubexprs cond.type) 31 | return result 32 | 33 | def genDerivedGuides (goal : Congr) (rws : Array Rewrite) : MetaM Guides := do 34 | let guides : ExprSet ← rws.foldlM (init := {← goal.expr}) (return ·.union <| ← deriveGuides ·) 35 | let mut result := #[] 36 | for guide in guides, idx in [:guides.size] do 37 | result := result.push <| ← Guide.from guide (.guide idx (derived := true)) 38 | return result 39 | -------------------------------------------------------------------------------- /Lean/Egg/Tests/Rudi Groups.lean: -------------------------------------------------------------------------------- 1 | import Egg 2 | 3 | class Group (α) extends One α, Inv α, Mul α where 4 | mul_assoc (a b c : α) : (a * b) * c = a * (b * c) 5 | one_mul (a : α) : 1 * a = a 6 | mul_one (a : α) : a * 1 = a 7 | inv_mul_self (a : α) : a⁻¹ * a = 1 8 | mul_inv_self (a : α) : a * a⁻¹ = 1 9 | 10 | variable [Group G] {a b : G} 11 | 12 | open Group Egg.Guides Egg.Config.Modifier in 13 | macro "group" mod:egg_cfg_mod guides:(egg_guides)? : tactic => `(tactic| 14 | egg $mod:egg_cfg_mod [mul_assoc, one_mul, mul_one, inv_mul_self, mul_inv_self] $[$guides]? 15 | ) 16 | 17 | set_option hygiene false in 18 | macro "rudify " lhs:term:51 " = " rhs:term : term => `( 19 | (fun x => (fun t _ => t) (fun z => x z)) (fun _ => $lhs) = 20 | (fun (x : G → G) (_ : α) => x) (fun _ => $rhs) 21 | ) 22 | 23 | set_option egg.timeLimit 10 24 | 25 | theorem inv_mul_cancel_left : a⁻¹ * (a * b) = b := by group 26 | 27 | theorem inv_mul_cancel_left' : rudify a⁻¹ * (a * b) = b := by group 28 | 29 | theorem mul_inv_cancel_left : a * (a⁻¹ * b) = b := by group 30 | 31 | theorem mul_inv_cancel_left' : rudify a * (a⁻¹ * b) = b := by group 32 | 33 | theorem inv_one : (1 : G)⁻¹ = 1 := by 34 | group using 1 * (1 : G)⁻¹ 35 | 36 | theorem inv_one' : rudify (1 : G)⁻¹ = 1 := by 37 | group using 1 * (1 : G)⁻¹ 38 | 39 | theorem inv_mul : (a * b)⁻¹ = b⁻¹ * a⁻¹ := by 40 | group using b⁻¹ * a⁻¹ * (a * b) * (a * b)⁻¹ 41 | 42 | -- TODO: Broken explanation 43 | theorem inv_mul' : rudify (a * b)⁻¹ = b⁻¹ * a⁻¹ := by 44 | sorry -- group using b⁻¹ * a⁻¹ * (a * b) * (a * b)⁻¹ 45 | 46 | theorem inv_inv : a⁻¹⁻¹ = a := by 47 | group using a⁻¹⁻¹ * a⁻¹ * a 48 | 49 | theorem inv_inv' : rudify a⁻¹⁻¹ = a := by 50 | group using a⁻¹⁻¹ * a⁻¹ * a 51 | -------------------------------------------------------------------------------- /Lean/Egg/Core/Congr.lean: -------------------------------------------------------------------------------- 1 | import Egg.Core.Normalize 2 | import Lean 3 | open Lean Meta 4 | 5 | namespace Egg.Congr 6 | 7 | inductive Rel where 8 | | eq 9 | | iff 10 | deriving Inhabited 11 | 12 | structure _root_.Egg.Congr where 13 | rel : Congr.Rel 14 | lhs : Expr 15 | rhs : Expr 16 | deriving Inhabited 17 | 18 | namespace Rel 19 | 20 | def relate (e₁ e₂ : Expr) : Rel → MetaM Expr 21 | | eq => mkEq e₁ e₂ 22 | | iff => return mkIff e₁ e₂ 23 | 24 | def mkRefl (expr : Expr) : Rel → MetaM Expr 25 | | eq => mkEqRefl expr 26 | | iff => mkAppM ``Iff.refl #[expr] 27 | 28 | def mkSymm (proof : Expr) : Rel → MetaM Expr 29 | | eq => mkEqSymm proof 30 | | iff => mkAppM ``Iff.symm #[proof] 31 | 32 | def mkTrans (proof₁ proof₂ : Expr) : Rel → MetaM Expr 33 | | eq => mkEqTrans proof₁ proof₂ 34 | | iff => mkAppM ``Iff.trans #[proof₁, proof₂] 35 | 36 | def mkMP (proof : Expr) : Rel → MetaM Expr 37 | | eq => mkAppM ``Eq.mp #[proof] 38 | | iff => mkAppM ``Iff.mp #[proof] 39 | 40 | end Rel 41 | 42 | def expr (cgr : Congr) : MetaM Expr := do 43 | match cgr.rel with 44 | | .eq => mkEq cgr.lhs cgr.rhs 45 | | .iff => return mkIff cgr.lhs cgr.rhs 46 | 47 | -- Since `=` and `↔` are not heterogeneous, we assume `lhs` and `rhs` to have the same type. 48 | def type (cgr : Congr) : MetaM Expr := 49 | inferType cgr.lhs 50 | 51 | def from? (type : Expr) : MetaM (Option Congr) := do 52 | let type ← normalize type .noReduce 53 | if let some (_, lhs, rhs) := type.eq? then 54 | return some { rel := .eq, lhs, rhs } 55 | else if let some (lhs, rhs) := type.iff? then 56 | return some { rel := .iff, lhs, rhs } 57 | else 58 | return none 59 | 60 | def from! (type : Expr) : MetaM Congr := do 61 | return (← Congr.from? type).get! 62 | -------------------------------------------------------------------------------- /Lean/Egg/Tests/Dataflow Rewriter/DataflowRewriter/One.lean: -------------------------------------------------------------------------------- 1 | import Egg 2 | import Mathlib 3 | 4 | -- works 5 | example {α} {l : List α} (i : Nat) (h₁ h₂ : i < l.length) : 6 | l.get (Fin.mk i h₁) = l.get (Fin.mk i h₂) := by rfl 7 | 8 | -- works 9 | example {α} {l : List α} (i : Nat) (h₁ h₂ : i < l.length) : 10 | Fin.mk i h₁ = Fin.mk i h₂ := by rfl 11 | 12 | -- works 13 | example {α} {l : List α} (i : Nat) (h₁ h₂ : i < l.length) : 14 | l.get (Fin.mk i h₁) = l.get (Fin.mk i h₂) := by egg 15 | 16 | -- works 17 | example {α} {l : List α} (i : Nat) (h₁ h₂ : i < l.length) : 18 | Fin.mk i h₁ = Fin.mk i h₂ := by egg 19 | 20 | -- works 21 | example T x1 x2 val isLt : 22 | @Eq T 23 | (@List.get T x1 24 | (@Fin.mk (@List.length T x1) val 25 | (@Fin.val_lt_of_le (@List.length T x1) (@List.length T (@Prod.mk (List T) (List T) x1 x2).1) 26 | (@Fin.mk (@List.length T (@Prod.mk (List T) (List T) x1 x2).1) val isLt) 27 | (@le_refl Nat 28 | (@PartialOrder.toPreorder Nat (@StrictOrderedSemiring.toPartialOrder Nat Nat.instStrictOrderedSemiring)) 29 | (@List.length T x1))))) 30 | (@List.get T x1 (@Fin.mk (@List.length T (@Prod.mk (List T) (List T) x1 x2).1) val isLt)) := by rfl 31 | 32 | -- fails 33 | example T x1 x2 val isLt : 34 | @Eq T 35 | (@List.get T x1 36 | (@Fin.mk (@List.length T x1) val 37 | (@Fin.val_lt_of_le (@List.length T x1) (@List.length T (@Prod.mk (List T) (List T) x1 x2).1) 38 | (@Fin.mk (@List.length T (@Prod.mk (List T) (List T) x1 x2).1) val isLt) 39 | (@le_refl Nat 40 | (@PartialOrder.toPreorder Nat (@StrictOrderedSemiring.toPartialOrder Nat Nat.instStrictOrderedSemiring)) 41 | (@List.length T x1))))) 42 | (@List.get T x1 (@Fin.mk (@List.length T (@Prod.mk (List T) (List T) x1 x2).1) val isLt)) := by egg 43 | -------------------------------------------------------------------------------- /Lean/Egg.lean: -------------------------------------------------------------------------------- 1 | import Egg.Tactic.Baskets 2 | import Egg.Tactic.Calc 3 | import Egg.Tactic.Rewrite.Rules.Gen.Intros 4 | import Egg.Tactic.Rewrite.Rules.Gen.GenM 5 | import Egg.Tactic.Rewrite.Rules.Gen.Basic 6 | import Egg.Tactic.Rewrite.Rules.Parse 7 | import Egg.Tactic.Config.Option 8 | import Egg.Tactic.Config.Modifier 9 | import Egg.Tactic.Calcify 10 | import Egg.Tactic.Basic 11 | import Egg.Tactic.Trace 12 | import Egg.Tactic.Goal 13 | import Egg.Tactic.Guides 14 | import Egg.Core.Rewrite.Basic 15 | import Egg.Core.Rewrite.Rule 16 | import Egg.Core.Rewrite.Prerewrite 17 | import Egg.Core.Rewrite.Condition 18 | import Egg.Core.Explanation.Parse.Slotted 19 | import Egg.Core.Explanation.Parse.Basic 20 | import Egg.Core.Explanation.Parse.Egg 21 | import Egg.Core.Explanation.Parse.Shared 22 | import Egg.Core.Explanation.Basic 23 | import Egg.Core.Explanation.Congr 24 | import Egg.Core.Explanation.Proof 25 | import Egg.Core.Explanation.Expr 26 | import Egg.Core.Encode.Shapes 27 | import Egg.Core.Encode.Rules 28 | import Egg.Core.Encode.Basic 29 | import Egg.Core.Encode.Blocks 30 | import Egg.Core.Encode.EncodeM 31 | import Egg.Core.Encode.Guides 32 | import Egg.Core.Congr 33 | import Egg.Core.MVars.Collect 34 | import Egg.Core.MVars.Basic 35 | import Egg.Core.MVars.Subst 36 | import Egg.Core.Gen.Tagged 37 | import Egg.Core.Gen.StructProjs 38 | import Egg.Core.Gen.GoalTypeSpecialization 39 | import Egg.Core.Gen.Explosion 40 | import Egg.Core.Gen.Builtins 41 | import Egg.Core.Gen.TcProjs 42 | import Egg.Core.Gen.Guides 43 | import Egg.Core.Direction 44 | import Egg.Core.Normalize 45 | import Egg.Core.Source 46 | import Egg.Core.Request.EGraph 47 | import Egg.Core.Request.Basic 48 | import Egg.Core.Request.Term 49 | import Egg.Core.Request.Equiv 50 | import Egg.Core.Request.Synth 51 | import Egg.Core.Config 52 | import Egg.Core.Blocks 53 | import Egg.Core.Guides 54 | -------------------------------------------------------------------------------- /Lean/Egg/Tests/Freshman Calc.lean: -------------------------------------------------------------------------------- 1 | import Egg 2 | 3 | class CommRing (α) extends Zero α, One α, Add α, Sub α, Mul α, Div α, Pow α Nat, Inv α, Neg α where 4 | comm_add (a b : α) : a + b = b + a 5 | comm_mul (a b : α) : a * b = b * a 6 | add_assoc (a b c : α) : a + (b + c) = (a + b) + c 7 | mul_assoc (a b c : α) : a * (b * c) = (a * b) * c 8 | sub_canon (a b : α) : a - b = a + -b 9 | neg_add (a : α) : a + -a = 0 10 | div_canon (a b : α) : a / b = a * b⁻¹ 11 | zero_add (a : α) : a + 0 = a 12 | zero_mul (a : α) : a * 0 = 0 13 | one_mul (a : α) : a * 1 = a 14 | distrib (a b c : α) : a * (b + c) = (a * b) + (a * c) 15 | pow_zero (a : α) : a ^ 0 = 1 16 | pow_succ (a : α) : a ^ (n + 1) = a * (a ^ n) 17 | 18 | class CharTwoRing (α) extends CommRing α where 19 | char_two (a : α) : a + a = 0 20 | 21 | open CommRing CharTwoRing 22 | variable [CharTwoRing α] (x y : α) 23 | 24 | theorem freshmans_dream₂ : (x + y) ^ 2 = x ^ 2 + y ^ 2 := by 25 | egg calc [comm_add, comm_mul, add_assoc, mul_assoc, sub_canon, neg_add, div_canon, zero_add, 26 | zero_mul, one_mul, distrib, pow_zero, pow_succ, char_two, Nat.succ_eq_add_one] 27 | _ = (x + y) * (x + y) 28 | _ = x * (x + y) + y * (x + y) 29 | _ = _ 30 | 31 | -- NOTE: The explanation is too long. 32 | -- theorem freshmans_dream₃ : (x + y) ^ 3 = x ^ 3 + x * y ^ 2 + x ^ 2 * y + y ^ 3 := by 33 | -- egg calc [comm_add, comm_mul, add_assoc, mul_assoc, sub_canon, neg_add, div_canon, zero_add, zero_mul, one_mul, distrib, pow_zero, pow_succ, char_two, Nat.succ_eq_add_one] 34 | -- _ = (x + y) * (x + y) * (x + y) 35 | -- _ = (x + y) * (x * (x + y) + y * (x + y)) 36 | -- _ = (x + y) * (x ^ 2 + y ^ 2) 37 | -- _ = x * (x ^ 2 + y ^ 2) + y * (x ^ 2 + y ^ 2) 38 | -- _ = (x * x ^ 2) + x * y ^ 2 + y * x ^ 2 + y * y ^ 2 39 | -- _ = _ 40 | -------------------------------------------------------------------------------- /Lean/Egg/Tests/Math/Math/Figures/Lie Algebra.lean: -------------------------------------------------------------------------------- 1 | import Mathlib.Algebra.Lie.Basic 2 | import Egg 3 | 4 | set_option egg.timeLimit 10 5 | 6 | attribute [egg lie_external] neg_eq_iff_add_eq_zero zero_add add_zero smul_neg sub_eq_zero neg_neg 7 | sub_neg_eq_add neg_add_cancel sub_eq_add_neg add_sub_cancel_right 8 | smul_sub 9 | 10 | -- TODO: extends AddCommGroup 11 | attribute [egg lie_ring] LieRing.add_lie LieRing.lie_add LieRing.lie_self LieRing.leibniz_lie 12 | 13 | -- TODO: extends CommRing and Module 14 | egg_basket lie_alg extends lie_ring with LieAlgebra.lie_smul 15 | 16 | egg_basket lie_ring_mod extends lie_ring, lie_alg with 17 | LieRingModule.add_lie, LieRingModule.lie_add, LieRingModule.leibniz_lie 18 | 19 | egg_basket lie_mod extends lie_ring, lie_alg, lie_ring_mod with 20 | LieModule.smul_lie, LieModule.lie_smul 21 | 22 | attribute [egg lie_tower] leibniz_lie lie_swap_lie 23 | 24 | egg_basket lie extends lie_ring_mod, lie_tower, lie_external with 25 | add_lie, lie_add, smul_lie, lie_smul, lie_zero, zero_lie, lie_self 26 | 27 | section BasicProperties 28 | 29 | variable {R : Type u} {L : Type v} {M : Type w} {N : Type w₁} 30 | variable [CommRing R] [LieRing L] [LieAlgebra R L] 31 | variable [AddCommGroup M] [Module R M] [LieRingModule L M] [LieModule R L M] 32 | variable [AddCommGroup N] [Module R N] [LieRingModule L N] [LieModule R L N] 33 | variable (t : R) (x y z : L) (m n : M) 34 | 35 | theorem L1' : ⁅x, y⁆ = -⁅y, x⁆ := by 36 | have h := by egg +lie calc 37 | 0 = ⁅x + y, x + y⁆ 38 | _ = ⁅x, x⁆ + ⁅x, y⁆ + ⁅y, x⁆ + ⁅y, y⁆ 39 | _ = ⁅x, y⁆ + ⁅y, x⁆ 40 | egg +lie [h] 41 | 42 | example /- L1' -/ : ⁅x, y⁆ = -⁅y, x⁆ := by 43 | egg +lie using ⁅x + y, x⁆ + ⁅x + y, y⁆ 44 | 45 | attribute [egg lie] L1' 46 | 47 | example /- neg_lie -/ : ⁅-x, m⁆ = -⁅x, m⁆ := by 48 | egg +lie 49 | -------------------------------------------------------------------------------- /Lean/Egg/Tests/Derived Guides.lean: -------------------------------------------------------------------------------- 1 | import Egg 2 | 3 | set_option egg.derivedGuides true 4 | set_option egg.builtins false 5 | set_option egg.tcProjs false 6 | 7 | set_option trace.egg.guides true 8 | set_option linter.unusedVariables false 9 | 10 | /-- 11 | trace: [egg.guides] Guides 12 | [egg.guides] ↣0!: 0 13 | [egg.guides] ↣1!: 0 = 0 14 | -/ 15 | #guard_msgs in 16 | example (h : 0 = 0) : 0 = 0 := by 17 | egg [h] 18 | 19 | /-- 20 | trace: [egg.guides] Guides 21 | [egg.guides] ↣0!: ∀ (x : Nat), x = 0 22 | [egg.guides] ↣1!: True 23 | [egg.guides] ↣2!: 0 24 | [egg.guides] ↣3!: 0 = 0 25 | -/ 26 | #guard_msgs in 27 | example (h : ∀ x : Nat, x = 0) : 0 = 0 := by 28 | egg [h] 29 | 30 | /-- 31 | trace: [egg.guides] Guides 32 | [egg.guides] ↣0!: ∀ (p q : Prop), p ∧ q 33 | [egg.guides] ↣1!: True 34 | [egg.guides] ↣2!: And 35 | [egg.guides] ↣3!: 0 = 0 36 | -/ 37 | #guard_msgs in 38 | example (h : ∀ p q : Prop, p ∧ q) : 0 = 0 := by 39 | egg [h] 40 | 41 | /-- 42 | trace: [egg.guides] Guides 43 | [egg.guides] ↣0!: ∀ (p q : Prop), 1 = 2 → p ∧ q 44 | [egg.guides] ↣1!: True 45 | [egg.guides] ↣2!: And 46 | [egg.guides] ↣3!: 1 = 2 47 | [egg.guides] ↣4!: 0 = 0 48 | -/ 49 | #guard_msgs in 50 | example (h : ∀ p q : Prop, (1 = 2) → p ∧ q) : 0 = 0 := by 51 | egg [h] 52 | 53 | /-- 54 | trace: [egg.guides] Guides 55 | [egg.guides] ↣0!: ∀ (p : Prop), p 56 | [egg.guides] ↣1!: 0 = 0 57 | -/ 58 | #guard_msgs in 59 | example (h : (∀ p : Prop, p) ↔ (∀ q : Prop, q)) : 0 = 0 := by 60 | egg [h] 61 | 62 | /-- 63 | trace: [egg.guides] Guides 64 | [egg.guides] ↣0!: ∀ (a : Prop), (∀ (p : Prop), a ∧ p) ↔ ∀ (q : Prop), q 65 | [egg.guides] ↣1!: Prop 66 | [egg.guides] ↣2!: ∀ (q : Prop), q 67 | [egg.guides] ↣3!: True 68 | [egg.guides] ↣4!: And 69 | [egg.guides] ↣5!: 0 = 0 70 | -/ 71 | #guard_msgs in 72 | example (h : ∀ a : Prop, (∀ p : Prop, a ∧ p) ↔ (∀ q : Prop, q)) : 0 = 0 := by 73 | egg [h] 74 | -------------------------------------------------------------------------------- /Lean/Egg/Tactic/Goal.lean: -------------------------------------------------------------------------------- 1 | import Egg.Core.Congr 2 | import Egg.Lean 3 | import Lean 4 | 5 | open Lean Meta Elab Tactic 6 | 7 | namespace Egg 8 | 9 | structure Goal extends Congr where private mk :: 10 | id : MVarId 11 | intros : Array (FVarId × Name) 12 | 13 | def Goal.gen (goal : MVarId) : TacticM Goal := 14 | goal.withContext do 15 | let goal ← getMainGoal 16 | let goalType ← withReducible goal.getType' 17 | let goalTypeType ← inferType goalType 18 | unless goalTypeType.isProp do 19 | throwError m!"goal type is not a proposition:\n {goalType} : {goalTypeType}" 20 | let fvars := (← getLCtx).getFVarIds 21 | withReducible do evalTactic <| ← `(tactic|repeat intro) 22 | let goal ← getMainGoal 23 | let (goal, intros) ← genIntros goal fvars 24 | goal.withContext do 25 | let goalType ← withReducible goal.getType' 26 | if let some cgr ← Congr.from? goalType then 27 | return { cgr with id := goal, intros } 28 | else 29 | let goalEqTrue ← mkEq goalType (.const ``True []) 30 | let mGoalEqTrue ← mkFreshExprMVar goalEqTrue 31 | let oldProof ← mkOfEqTrue mGoalEqTrue 32 | goal.assignIfDefeq' oldProof 33 | let cgr ← Congr.from! goalEqTrue 34 | return { cgr with id := mGoalEqTrue.mvarId!, intros } 35 | where 36 | genIntros (goal : MVarId) (previousFVars : Array FVarId) : MetaM (MVarId × Array (FVarId × Name)) := do 37 | goal.withContext do 38 | let mut goal := goal 39 | let mut intros := #[] 40 | let newFVars := (← getLCtx).getFVarIds.filter (!previousFVars.contains ·) 41 | for fvar in newFVars do 42 | let (g, name) ← goal.withContext do 43 | let userName := (← getLCtx).getUnusedName (← fvar.getUserName) 44 | pure (← goal.rename fvar userName, userName) 45 | goal := g 46 | intros := intros.push (fvar, name) 47 | return (goal, intros) 48 | -------------------------------------------------------------------------------- /Lean/Egg/Tests/WIP.lean: -------------------------------------------------------------------------------- 1 | import Egg 2 | 3 | -- The universe mvars (or universe params if you make this a theorem instead of an example) are 4 | -- different for the respective `α`s, so this doesn't hold by reflexivity. But `by rfl` can somehow 5 | -- prove this. 6 | example : (∀ α (l : List α), l.length = l.length) ↔ (∀ α (l : List α), l.length = l.length) := by 7 | set_option trace.egg true in egg 8 | 9 | -- For rewrites involving dependent arguments, we can easily get an incorrect motive. E.g. when 10 | -- rewriting the condition in ite without changing the type class instance: 11 | set_option trace.egg true in 12 | example : (if 0 = 0 then 0 else 1) = 0 := by 13 | have h1 : (0 = 0) = True := eq_self 0 14 | have h2 : 0 = 0 := rfl 15 | egg [h1, h2, ite_congr, if_true] 16 | 17 | -- For typeclass arguments we might be able to work around this by the following: 18 | -- When a rewrite is applied to a term containing a typeclass argument (which we might be able to 19 | -- track via e-class analysis), export that term, check if it's type correct, and if not, 20 | -- try to resynthesize any contained typeclass instances. If this works reintroduce the typecorrect 21 | -- term into the egraph. 22 | -- How do we prove that this new term is equivalent to the old one though? Changing typeclass 23 | -- instances isn't generally defeq. 24 | 25 | -- Could it be that it is usually the case that if it makes sense to rewrite a dependent argument 26 | -- by itself then its only dependents will be typeclass arguments (because otherwise the result 27 | -- would need to involve a cast or something like that)? 28 | 29 | -- Simp only somehow knows how to handle this: 30 | set_option pp.explicit true in 31 | theorem t : (if 0 = 0 then 0 else 1) = 0 := by 32 | have : (0 = 0) = True := eq_self 0 33 | simp only [this] 34 | sorry 35 | 36 | -- Where does it pull `ite_congr` from? Does it have something to do with the `congr` attribute? 37 | #print t 38 | #check ite_congr 39 | -------------------------------------------------------------------------------- /Rust/Egg/src/expl.rs: -------------------------------------------------------------------------------- 1 | use egg::*; 2 | use crate::analysis::*; 3 | use crate::lean_expr::*; 4 | 5 | #[derive(Clone, Copy)] 6 | pub enum ExplanationKind { 7 | None, 8 | SameEClass, 9 | EqTrue 10 | } 11 | 12 | impl ExplanationKind { 13 | pub fn to_c(self) -> u8 { 14 | match self { 15 | Self::None => 0, 16 | Self::SameEClass => 1, 17 | Self::EqTrue => 2, 18 | } 19 | } 20 | 21 | fn for_goal( 22 | egraph: &LeanEGraph, init_expr: RecExpr, goal_expr: RecExpr, init_id: Id, 23 | goal_id: Id 24 | ) -> (ExplanationKind, RecExpr, RecExpr) { 25 | if egraph.find(init_id) == egraph.find(goal_id) { 26 | return (ExplanationKind::SameEClass, init_expr, goal_expr) 27 | } 28 | 29 | let eq_expr = format!("(= {} {})", init_expr, goal_expr).parse().unwrap(); 30 | let true_expr = "(const \"True\")".parse().unwrap(); 31 | let true_id = egraph.lookup_expr(&true_expr).unwrap(); 32 | 33 | // Note: `lookup_expr` canonicalizes ids. 34 | if egraph.lookup_expr(&eq_expr) == Some(true_id) { 35 | return (ExplanationKind::EqTrue, eq_expr, true_expr) 36 | } 37 | 38 | (ExplanationKind::None, RecExpr::default(), RecExpr::default()) 39 | } 40 | } 41 | 42 | pub fn mk_explanation( 43 | egraph: &mut LeanEGraph, init_expr: RecExpr, goal_expr: RecExpr, init_id: Id, 44 | goal_id: Id 45 | ) -> (ExplanationKind, String) { 46 | let (kind, lhs, rhs) = ExplanationKind::for_goal(egraph, init_expr, goal_expr, init_id, goal_id); 47 | let expl: String; 48 | match kind { 49 | ExplanationKind::None => expl = "".to_string(), 50 | ExplanationKind::EqTrue | ExplanationKind::SameEClass => { 51 | expl = egraph.explain_equivalence(&lhs, &rhs).get_flat_string() 52 | } 53 | } 54 | 55 | (kind, expl) 56 | } 57 | -------------------------------------------------------------------------------- /Rust/Egg/src/util.rs: -------------------------------------------------------------------------------- 1 | use std::collections::HashSet; 2 | use std::hash::Hash; 3 | use egg::*; 4 | 5 | pub enum Either { 6 | Left(L), 7 | Right(R), 8 | } 9 | 10 | pub fn sub_expr(ast: &RecExpr, i: Id) -> RecExpr { 11 | let v: Vec<_> = ast.as_ref()[0..=usize::from(i)].iter().cloned().collect(); 12 | RecExpr::from(v) 13 | } 14 | 15 | pub fn union_sets(to: &mut HashSet, from: HashSet) -> DidMerge { 16 | let to_sub_from = to.is_subset(&from); 17 | let from_sub_to = from.is_subset(to); 18 | *to = &*to | &from; 19 | match (to_sub_from, from_sub_to) { 20 | (false, false) => DidMerge(true, true), 21 | (false, true) => DidMerge(false, true), 22 | (true, false) => DidMerge(true, false), 23 | (true, true) => DidMerge(false, false) 24 | } 25 | } 26 | 27 | pub fn intersect_sets(to: &mut HashSet, from: HashSet) -> DidMerge { 28 | let to_sub_from = to.is_subset(&from); 29 | let from_sub_to = from.is_subset(to); 30 | *to = &*to & &from; 31 | match (to_sub_from, from_sub_to) { 32 | (false, false) => DidMerge(true, true), 33 | (false, true) => DidMerge(true, false), 34 | (true, false) => DidMerge(false, true), 35 | (true, true) => DidMerge(false, false) 36 | } 37 | } 38 | 39 | // TODO: Figure out how to create the union of two hash sets properly. 40 | pub fn union_clone(fst: &HashSet, snd: &HashSet) -> HashSet { 41 | let mut result = fst.clone(); 42 | for elem in snd { result.insert(*elem); } 43 | return result 44 | } 45 | 46 | pub fn shift_down(indices: &HashSet) -> HashSet { 47 | let mut result = HashSet::with_capacity(indices.len()); 48 | for &idx in indices { 49 | if idx == 0 { 50 | continue 51 | } else { 52 | result.insert(idx - 1); 53 | } 54 | } 55 | return result 56 | } -------------------------------------------------------------------------------- /Lean/Egg/Tests/TC Proj Binders.lean: -------------------------------------------------------------------------------- 1 | 2 | import Egg 3 | 4 | set_option linter.unusedVariables false 5 | 6 | -- TODO: This should generate a type class projection reduction from HAdd.hAdd to Add.add. 7 | -- It doesn't because `α` and the `inst : Add α` are bvars. 8 | -- To fix this, the tc-proj generator should generate fvars for the given bvars, perform the 9 | -- reduction, and then over any remaining bvar-fvars. 10 | -- #guard_msgs in 11 | set_option egg.builtins false in 12 | example (h : (fun (α) [Mul α] (x y : α) => x * y) = a) : true = true := by 13 | sorry -- egg [h] 14 | 15 | -- TODO: This should generate a goal type specialization for `h`, but I think it doesn't for the 16 | -- same reason as outlined above. 17 | example (h : ∀ {α} [Add α] (a : α), a + a = a) : 1 + 1 = 1 := by 18 | sorry -- egg [h] 19 | 20 | -- This test ensures that projection reductions are produced for terms appearing in binder domains. 21 | /-- 22 | trace: [egg.rewrites.tcProj] Type Class Projections (4) 23 | [egg.rewrites.tcProj] #0[0?69632,0](⇒) 24 | [egg.rewrites.tcProj] HMul.hMul = Mul.mul 25 | [egg.rewrites.tcProj] LHS MVars 26 | [] 27 | [egg.rewrites.tcProj] RHS MVars 28 | [] 29 | [egg.rewrites.tcProj] #0[0?69632,0](⇐) 30 | [egg.rewrites.tcProj] Mul.mul = HMul.hMul 31 | [egg.rewrites.tcProj] LHS MVars 32 | [] 33 | [egg.rewrites.tcProj] RHS MVars 34 | [] 35 | [egg.rewrites.tcProj] #0[0?69632,1](⇒) 36 | [egg.rewrites.tcProj] Mul.mul = Nat.mul 37 | [egg.rewrites.tcProj] LHS MVars 38 | [] 39 | [egg.rewrites.tcProj] RHS MVars 40 | [] 41 | [egg.rewrites.tcProj] #0[0?69632,1](⇐) 42 | [egg.rewrites.tcProj] Nat.mul = Mul.mul 43 | [egg.rewrites.tcProj] LHS MVars 44 | [] 45 | [egg.rewrites.tcProj] RHS MVars 46 | [] 47 | -/ 48 | #guard_msgs in 49 | set_option trace.egg.rewrites.tcProj true in 50 | set_option egg.builtins false in 51 | example (x : Nat) (h : ∀ (_ : x * y = z), z = z) : x = x := by 52 | egg [h] 53 | -------------------------------------------------------------------------------- /Lean/Egg/Tests/EraseProofs.lean: -------------------------------------------------------------------------------- 1 | import Egg 2 | 3 | example (arr : Array α) (i : Nat) (h₁ h₂ : i < arr.size) : arr[i]'h₁ = arr[i]'h₂ := by 4 | egg 5 | 6 | -- Note: This and the next test case also demonstrate an interaction between proof erasure and 7 | -- conditional rewrites. Without proof erasure, the argument `g` of the rewrite is not 8 | -- considered a precondition, as it appears in the LHS of the equality. Naively, proof erasure 9 | -- implies that `g` does not appear in the equation (as it is erased), and thus becomes a pre- 10 | -- condition. Thus, we need to handle proof terms specially when determining preconditions. 11 | example (i : Nat) (h : i < 10) : (Fin.mk i h).val = i := by 12 | have : ∀ n m (g : n < m), (Fin.mk n g).val = n := by simp 13 | egg [this] 14 | 15 | example (i : Nat) (h : ∀ i : Nat, i < 10) : (Fin.mk i (h i)).val = i := by 16 | have : ∀ n m (g : n < m), (Fin.mk n g).val = n := by simp 17 | egg [this] 18 | 19 | -- The following test is an attempt to construct a rewrite where the mvar `x` does appear in the 20 | -- proof term on the rhs, but not in the *type* of the proof term. If this succeeded, proof erasure 21 | -- should cause a crash during variable substitution after e-matching, as proofs are encoded by 22 | -- their type, which is therefore missing the mvar for `x`. But obviously, the rewrite `h` does not 23 | -- achieve this yet. 24 | example (h : ∀ x : Nat, x = Exists.choose (Exists.intro x x.zero_le)) : True = True := by 25 | egg [h] 26 | 27 | -- TODO: We don't currently support rewriting the types of proof terms. I think this shouldn't be 28 | -- too difficult to support though. 29 | /-- 30 | error: egg received invalid explanation: step contains non-defeq type-level rewrite in proof or type class instance 31 | -/ 32 | #guard_msgs in 33 | example 34 | (f : (a b : Nat) → a > b → Nat) (g : Nat → Nat) (a₁ a₂ b₁ b₂ c d : Nat) (h₁ : a₁ > b₁) 35 | (h₂ : a₂ > b₂) (h₃ : a₁ = c) (h₄ : a₂ = c) (h₅ : b₁ = d) (h₆ : d = b₂) : 36 | g (g (f a₁ b₁ h₁)) = g (g (f a₂ b₂ h₂)) := by 37 | egg [*] 38 | -------------------------------------------------------------------------------- /Rust/Egg/src/eta.rs: -------------------------------------------------------------------------------- 1 | use egg::*; 2 | use crate::analysis::*; 3 | use crate::lean_expr::*; 4 | 5 | pub fn eta_reduction_rw() -> LeanRewrite { 6 | rewrite!("≡η"; "(λ ?t (app ?f (bvar 0)))" => { Eta { fun : "?f".parse().unwrap() }}) 7 | } 8 | 9 | pub fn eta_expansion_rw() -> LeanRewrite { 10 | rewrite!("≡η+"; "?e" => { EtaExpand { term : "?e".parse().unwrap() }}) 11 | } 12 | 13 | struct Eta { 14 | fun: Var 15 | } 16 | 17 | impl Applier for Eta { 18 | 19 | fn apply_one(&self, graph: &mut LeanEGraph, _: Id, subst: &Subst, ast: Option<&PatternAst>, rule: Symbol) -> Vec { 20 | let fun_bvars = &graph[subst[self.fun]].data.loose_bvars; 21 | if fun_bvars.contains(&0) { return vec![] } 22 | 23 | let new_fun: PatternAst = if fun_bvars.is_empty() { 24 | format!("{}", self.fun).parse().unwrap() 25 | } else { 26 | format!("(↑ - 1 0 {})", self.fun).parse().unwrap() 27 | }; 28 | 29 | let (id, did_union) = graph.union_instantiations(ast.unwrap(), &new_fun, subst, rule); 30 | if did_union { vec![id] } else { vec![] } 31 | } 32 | } 33 | 34 | struct EtaExpand { 35 | term: Var 36 | } 37 | 38 | impl Applier for EtaExpand { 39 | 40 | fn apply_one(&self, graph: &mut LeanEGraph, _: Id, subst: &Subst, ast: Option<&PatternAst>, rule: Symbol) -> Vec { 41 | // TODO: This does not respect primitive constructors. 42 | 43 | let term_bvars = &graph[subst[self.term]].data.loose_bvars; 44 | 45 | let new_term: PatternAst = if term_bvars.is_empty() { 46 | format!("{}", self.term).parse().unwrap() 47 | } else { 48 | format!("(↑ + 1 0 {})", self.term).parse().unwrap() 49 | }; 50 | 51 | let expanded = format!("(λ _ (app {} (bvar 0)))", new_term).parse().unwrap(); 52 | 53 | let (id, did_union) = graph.union_instantiations(ast.unwrap(), &expanded, subst, rule); 54 | if did_union { vec![id] } else { vec![] } 55 | } 56 | } -------------------------------------------------------------------------------- /Lean/Egg/Tests/Math/Math/Comparisons/Rotman Basic.lean: -------------------------------------------------------------------------------- 1 | import Mathlib 2 | 3 | open Nat 4 | 5 | notation:100 r "﹗" => Real.Gamma (r + 1) 6 | 7 | variable {n r : Nat} (h₁ : n + 1 ≥ r) (h₂ : r ≠ 0) (h₃ : r ≠ n + 1) 8 | 9 | example : n ! / ((r - 1)! * (n - r + 1)!) + n ! / (r ! * (n - r)!) = (n + 1)! / (r ! * (n + 1 - r)!) := 10 | have h₁ : (r - 1)! * (n - r + 1)! ∣ n ! := sorry 11 | have h₂ : ↑((r - 1)! * (n - r + 1)!) ≠ (0 : ℝ) := sorry 12 | have h₃ : r ! * (n - r)! ∣ n ! := sorry 13 | have h₄ : ↑(r ! * (n - r)!) ≠ (0 : ℝ) := sorry 14 | have h₅ : (n : ℝ) - r + 1 ≠ 0 := sorry 15 | have h₆ : (r : ℝ) - 1 + 1 ≠ 0 := sorry 16 | have h₇ : (r : ℝ) ≠ 0 := sorry 17 | have h₈ : (n : ℝ) + 1 ≠ 0 := sorry 18 | have h₉ : r ! * (n + 1 - r)! ∣ (n + 1)! := sorry 19 | have h₁₀ : ↑(r ! * (n + 1 - r)!) ≠ (0 : ℝ) := sorry 20 | 21 | cast_inj.mp <| calc 22 | _ = n﹗ / ((r - 1)﹗ * (n - r + 1)﹗) + n﹗ / (r﹗ * (n - r)﹗) := by 23 | rw [cast_add, cast_div h₁ h₂, cast_div h₃ h₄, cast_mul, cast_mul] 24 | repeat' rw [← Real.Gamma_nat_eq_factorial] 25 | rw [cast_sub (by omega), cast_add, cast_sub (by omega), cast_one] 26 | _ = n﹗ / ((r - 1)﹗ * (n - r)﹗) * (1 / (n - r + 1) + 1 / r) := by 27 | rw [Real.Gamma_add_one h₅, mul_comm (n - r + 1 : ℝ), 28 | ← mul_assoc, div_mul_eq_div_mul_one_div, ← sub_add_cancel ↑r (1 : ℝ), 29 | Real.Gamma_add_one h₆] 30 | ring_nf 31 | _ = n﹗ / ((r - 1)﹗ * (n - r)﹗) * ((r + (n - r + 1)) / (r * (n - r + 1))) := by 32 | rw [div_add_div _ _ h₅ h₇]; ring 33 | _ = n﹗ / ((r - 1)﹗ * (n - r)﹗) * ((n + 1) / (r * (n - r + 1))) := by ring 34 | _ = (n + 1)﹗ / (r﹗ * (n + 1 - r)﹗) := by 35 | rw [_root_.div_mul_div_comm, mul_comm, ← Real.Gamma_add_one h₈, mul_assoc, 36 | mul_comm ((n - r : ℝ)﹗), mul_assoc, ← mul_assoc, mul_comm ((r - 1 : ℝ)﹗), 37 | sub_add, sub_self, sub_zero, ← Real.Gamma_add_one h₇, ← Real.Gamma_add_one h₅] 38 | ring_nf 39 | _ = _ := by 40 | rw [cast_div h₉ h₁₀, cast_mul] 41 | repeat' rw [← Real.Gamma_nat_eq_factorial] 42 | rw [cast_add, cast_sub (by omega), cast_add, cast_one] 43 | -------------------------------------------------------------------------------- /Lean/Egg/Core/Config.lean: -------------------------------------------------------------------------------- 1 | namespace Egg.Config 2 | 3 | structure Normalization where 4 | betaReduceRws := true 5 | etaReduceRws := true 6 | natReduceRws := true 7 | deriving BEq 8 | 9 | def Normalization.noReduce : Normalization where 10 | betaReduceRws := false 11 | etaReduceRws := false 12 | natReduceRws := false 13 | 14 | structure Encoding extends Normalization where 15 | -- TODO: Currently, this option implicitly disables defeq rewrites as they can not handle shapes, yet. 16 | shapes := false 17 | deriving BEq 18 | 19 | -- TODO: The slotted backend is currently not supported. 20 | def Encoding.slotted (_ : Encoding) : Bool := 21 | false 22 | 23 | structure Gen where 24 | baskets := (#[] : Array Lean.Name) 25 | builtins := true 26 | structProjs := true 27 | tcProjs := true 28 | tcProjFusion := false 29 | goalTypeSpec := true 30 | groundEqs := true 31 | derivedGuides := true 32 | explosion := false 33 | deriving BEq 34 | 35 | structure DefEq where 36 | natLit := true 37 | eta := true 38 | etaExpand := false 39 | beta := true 40 | levels := true 41 | 42 | structure Backend where 43 | unionSemantics := true 44 | optimizeExpl := true 45 | timeLimit := 3 46 | nodeLimit := 1000000000000000000 47 | iterLimit := 1000000000000000000 48 | reporting := false 49 | flattenReports := false 50 | deriving BEq 51 | 52 | inductive Debug.ExitPoint 53 | | none 54 | | beforeEqSat 55 | | beforeProof 56 | deriving BEq, Inhabited 57 | 58 | structure Debug where 59 | exitPoint := Debug.ExitPoint.none 60 | proofFuel? := (none : Option Nat) 61 | vizPath := (none : Option String) 62 | deriving BEq 63 | 64 | structure _root_.Egg.Config extends Encoding, DefEq, Gen, Backend, Debug where 65 | retryWithShapes := false 66 | explLengthLimit := 200 67 | subgoals := false 68 | 69 | instance : Coe Encoding Normalization where 70 | coe := Encoding.toNormalization 71 | 72 | instance : Coe Config Encoding where 73 | coe := toEncoding 74 | 75 | instance : Coe Config Gen where 76 | coe := toGen 77 | -------------------------------------------------------------------------------- /Lean/Egg/Tests/DataflowRewriter.lean: -------------------------------------------------------------------------------- 1 | import Egg 2 | 3 | -- works 4 | example {α} {l : List α} (i : Nat) (h₁ h₂ : i < l.length) : 5 | l.get (Fin.mk i h₁) = l.get (Fin.mk i h₂) := by rfl 6 | 7 | -- works 8 | example {α} {l : List α} (i : Nat) (h₁ h₂ : i < l.length) : 9 | Fin.mk i h₁ = Fin.mk i h₂ := by rfl 10 | 11 | -- works 12 | example {α} {l : List α} (i : Nat) (h₁ h₂ : i < l.length) : 13 | l.get (Fin.mk i h₁) = l.get (Fin.mk i h₂) := by egg 14 | 15 | -- works 16 | example {α} {l : List α} (i : Nat) (h₁ h₂ : i < l.length) : 17 | Fin.mk i h₁ = Fin.mk i h₂ := by egg 18 | 19 | -- works 20 | example T x1 x2 val isLt : 21 | @Eq T 22 | (@List.get T x1 23 | (@Fin.mk (@List.length T x1) val 24 | isLt)) 25 | (@List.get T x1 26 | (@Fin.mk (@List.length T (@Prod.mk (List T) (List T) x1 x2).1) val 27 | isLt)) := by egg 28 | 29 | -- proof irrelevance works 30 | example (n val : Nat) (isLt₁ isLt₂ : val < n) : 31 | Eq (@Fin.mk n val isLt₁) 32 | (@Fin.mk n val isLt₂) := by egg 33 | 34 | -- structure projection rfl works 35 | example {val n isLt} : (@Fin.val n ⟨val, isLt⟩) = val := by egg 36 | 37 | -- proof irrelevance with value inside structure projection works 38 | example {val m n isLt} : ((@Fin.val n ⟨val, isLt⟩) < m) = (val < m) := by egg 39 | 40 | example {val n} (isLt : val < n) : 41 | (@LT.lt Nat instLTNat (↑(Fin.mk val isLt)) n) = (val < n) := by egg 42 | 43 | -- TODO: Turning around the goal to be `isLt' = isLt` fails during proof reconstruction in the sense 44 | -- that it produces a proof term containing unassigned mvars. I don't understand why this 45 | -- happens. Moreover, wrapping parts of the goal in `id` fixes the problem, even though the 46 | -- exact same explanation is produced. 47 | example {val n} (isLt : val < n) (isLt' : (@LT.lt Nat instLTNat (↑(Fin.mk val isLt)) n)) : 48 | isLt' = isLt := by 49 | egg [*] 50 | 51 | example {val T x1} (isLt : val < x1.length) : 52 | Eq (@Fin.val_lt_of_le (@List.length T x1) (@List.length T x1) (@Fin.mk (@List.length T x1) val isLt) (Nat.le_refl x1.length)) 53 | isLt := by 54 | egg [*, Fin.val_lt_of_le] 55 | -------------------------------------------------------------------------------- /Lean/Egg/Tests/batteries/mk_batteries.sh: -------------------------------------------------------------------------------- 1 | #!/usr/bin/env bash 2 | 3 | batteries_dir="$(realpath -s "$(dirname "$0")")" 4 | batteries_repo_dir="$batteries_dir/batteries" 5 | 6 | # Sets a variable corresponding to the `--main` flag. 7 | use_main_branch=false 8 | for arg in "$@" 9 | do 10 | if [[ "$arg" == "--main" ]]; then 11 | use_main_branch=true 12 | fi 13 | done 14 | 15 | # Fetches the Lean version used by egg (or uses `main` if the corresponding flag is given). 16 | if [[ "$use_main_branch" == true ]]; then 17 | lean_version="main" 18 | else 19 | lean_toolchain_file="$batteries_dir/../../../../lean-toolchain" 20 | lean_toolchain="$(cat "$lean_toolchain_file")" 21 | lean_version="${lean_toolchain#*:}" 22 | fi 23 | 24 | # Clones the previously determined version of batteries. 25 | cd "$batteries_dir" 26 | git clone -b "$lean_version" --single-branch --depth 1 "https://github.com/leanprover-community/batteries.git" || { exit 1; } 27 | cd "$batteries_repo_dir" || { exit 1; } 28 | git switch -c egg-tests || { exit 1; } 29 | git remote set-url origin "https://github.com/marcusrossel/batteries.git" || { exit 1; } 30 | 31 | # Adds egg as a dependency to the `lakefile.toml`. 32 | lakefile="$batteries_repo_dir/lakefile.toml" 33 | toml_dep="\n[[require]]\nname = \"egg\"\npath = \"../../../../..\"" 34 | echo -e "$toml_dep" >> "$lakefile" 35 | 36 | # Moves `SimpOnlyOverride.lean` to its intended destination. 37 | simp_only_override="$batteries_dir/SimpOnlyOverride.lean" 38 | simp_only_override_dst="$batteries_repo_dir/Batteries/Test/Egg" 39 | mkdir "$simp_only_override_dst" 40 | cp "$simp_only_override" "$simp_only_override_dst" 41 | 42 | # Adds an import statement for the simp only override to each `.lean` file, 43 | # except a few select ones. 44 | import_statement="import Batteries.Test.Egg.SimpOnlyOverride" 45 | cd "$batteries_repo_dir" 46 | targets="$(find Batteries -type f | grep -v -e 'SimpOnlyOverride.lean$' -e '.DS_Store')" || { exit 1; } 47 | 48 | while IFS= read -r file; do 49 | echo "$import_statement"$'\n'$"$(cat $file)" > "$file" 50 | done <<< "$targets" 51 | 52 | # Preparation for `lake build`. 53 | cd "$batteries_repo_dir" 54 | lake update -------------------------------------------------------------------------------- /Lean/Egg/Tests/Math/Math/Figures/Boolean Algebra.lean: -------------------------------------------------------------------------------- 1 | import Egg 2 | import Mathlib.Order.BooleanAlgebra.Basic 3 | 4 | attribute [egg slattice_sup] sup_idem sup_comm sup_assoc sup_left_right_swap sup_left_idem 5 | sup_right_idem sup_left_comm sup_right_comm sup_sup_sup_comm 6 | sup_sup_distrib_left sup_sup_distrib_right sup_congr_left 7 | sup_congr_right 8 | 9 | attribute [egg slattice_inf] inf_of_le_left inf_of_le_right inf_idem inf_comm inf_assoc 10 | inf_left_right_swap inf_left_idem inf_right_idem inf_left_comm 11 | inf_right_comm inf_inf_inf_comm inf_inf_distrib_left 12 | inf_inf_distrib_right inf_congr_left inf_congr_right 13 | 14 | egg_basket lattice extends slattice_sup, slattice_inf with 15 | inf_sup_self, sup_inf_self 16 | 17 | egg_basket distrib_lattice extends lattice with 18 | sup_inf_left, sup_inf_right, inf_sup_left, inf_sup_right, eq_of_inf_eq_sup_eq 19 | 20 | egg_basket bool extends distrib_lattice with 21 | sup_inf_sdiff, inf_inf_sdiff 22 | 23 | variable [GeneralizedBooleanAlgebra α] {x y z : α} 24 | 25 | example : y \ x ⊔ x = y ⊔ x := by 26 | egg +bool calc 27 | y \ x ⊔ x 28 | _ = y \ x ⊔ (x ⊔ x ⊓ y) 29 | _ = y ⊓ x ⊔ y \ x ⊔ x 30 | _ = y ⊔ x 31 | 32 | -- Note: These lemmas originate from `Mathlib.Order.BooleanAlgebra.Defs` and precede the proof of 33 | -- the following theorem. We can prove these lemmas (and others) lemmas using `egg +bool`. We 34 | -- just omit those proofs here for brevity. 35 | attribute [egg bool] inf_sdiff_right sdiff_sup sdiff_sdiff_right sdiff_sdiff_right' 36 | 37 | example : z \ (x \ y ⊔ y \ x) = z ⊓ (z \ x ⊔ y) ⊓ (z \ y ⊔ x) := by 38 | egg +bool calc 39 | z \ (x \ y ⊔ y \ x) 40 | _ = (z \ x ⊔ z ⊓ x ⊓ y) ⊓ (z \ y ⊔ z ⊓ y ⊓ x) 41 | _ = z ⊓ (z \ x ⊔ y) ⊓ (z \ y ⊔ z ⊓ y ⊓ x) 42 | _ = z ⊓ (z \ x ⊔ y) ⊓ (z ⊓ (z \ y ⊔ x)) 43 | _ = z ⊓ z ⊓ (z \ x ⊔ y) ⊓ (z \ y ⊔ x) 44 | _ = z ⊓ (z \ x ⊔ y) ⊓ (z \ y ⊔ x) 45 | 46 | example : z \ (x \ y ⊔ y \ x) = z ⊓ (z \ x ⊔ y) ⊓ (z \ y ⊔ x) := by 47 | egg +bool using z \ x ⊔ z ⊓ x ⊓ y 48 | -------------------------------------------------------------------------------- /Lean/Egg/Tests/Eta.lean: -------------------------------------------------------------------------------- 1 | import Egg 2 | 3 | -- Note: We disable β-reduction as it can also solve many of these cases without η-reduction. 4 | set_option egg.eta true 5 | set_option egg.beta false 6 | 7 | set_option egg.eta false in 8 | /-- error: egg failed to prove the goal (saturated) -/ 9 | #guard_msgs in 10 | example : (fun x => Nat.succ x) = Nat.succ := by 11 | egg 12 | 13 | example : (fun x => Nat.succ x) = Nat.succ := by 14 | egg 15 | 16 | example : id (fun x => Nat.succ x) = id Nat.succ := by 17 | egg 18 | 19 | example : (fun x => Nat.succ x) x = Nat.succ x := by 20 | egg 21 | 22 | example (f : Nat → Nat) (h : f = g) : (fun x : Nat => f x) y = g y := by 23 | egg [h] 24 | 25 | example (f : Nat → Nat) (h : f y = g) : (fun x : Nat => f x) y = g := by 26 | egg [h] 27 | 28 | elab "eta" n:num fn:ident ty:term : term => open Lean.Elab.Term in do 29 | let rec go (n : Nat) := 30 | if n = 0 31 | then elabTerm fn none 32 | else return .lam `x (← elabTerm ty none) (.app (← go <| n - 1) (.bvar 0)) .default 33 | go n.getNat 34 | 35 | example : (eta 2 Nat.succ Nat) = Nat.succ := by 36 | egg 37 | 38 | example : (eta 2 Nat.succ Nat) x = Nat.succ x := by 39 | egg 40 | 41 | example : id (eta 2 Nat.succ Nat) = id Nat.succ := by 42 | egg 43 | 44 | example : (eta 50 Nat.succ Nat) = Nat.succ := by 45 | egg 46 | 47 | set_option egg.eta false in 48 | /-- error: egg failed to prove the goal (saturated) -/ 49 | #guard_msgs in 50 | example (a : Nat) (h : ∀ b : Nat, b.succ.add a = 0) : (10 |> fun x => Nat.succ x).add a = 0 := by 51 | egg [h] 52 | 53 | example (a : Nat) (h : ∀ b : Nat, b.succ.add a = 0) : (10 |> fun x => Nat.succ x).add a = 0 := by 54 | egg [h] 55 | 56 | -- Note: This used to break when we were using direct e-class substitution instead of small-step 57 | -- substitution. 58 | -- Note: This relies on generated type class projection reductions. 59 | example : (fun x => x) = Add.add 0 := by 60 | egg [Nat.zero_add] 61 | 62 | -- This tests that we correctly handle explanations which explicitly contain the small-step 63 | -- substitition constructor. 64 | example (f : α → α → α → α) : (fun a b => (fun x => (f a b) x)) = (fun a b => f a b) := by 65 | egg 66 | -------------------------------------------------------------------------------- /Lean/Egg/Tests/Math/Math/Andrés/RingSupport.lean: -------------------------------------------------------------------------------- 1 | import Egg 2 | import Mathlib.RingTheory.PrimeSpectrum 3 | import Mathlib.Algebra.Module.LocalizedModule 4 | import Mathlib.RingTheory.Localization.AtPrime 5 | 6 | attribute [egg neg] Mathlib.Tactic.PushNeg.not_not_eq 7 | attribute [egg neg] Mathlib.Tactic.PushNeg.not_and_eq 8 | attribute [egg neg] Mathlib.Tactic.PushNeg.not_and_or_eq 9 | attribute [egg neg] Mathlib.Tactic.PushNeg.not_or_eq 10 | attribute [egg neg] Mathlib.Tactic.PushNeg.not_forall_eq 11 | attribute [egg neg] Mathlib.Tactic.PushNeg.not_exists_eq 12 | attribute [egg neg] Mathlib.Tactic.PushNeg.not_implies_eq 13 | attribute [egg neg] Mathlib.Tactic.PushNeg.not_ne_eq 14 | attribute [egg neg] Mathlib.Tactic.PushNeg.not_iff 15 | attribute [egg neg] Mathlib.Tactic.PushNeg.not_nonempty_eq 16 | attribute [egg neg] Mathlib.Tactic.PushNeg.ne_empty_eq_nonempty 17 | attribute [egg neg] Mathlib.Tactic.PushNeg.empty_ne_eq_nonempty 18 | 19 | variable {R M : Type*} [CommRing R] [AddCommGroup M] [Module R M] {p : PrimeSpectrum R} 20 | 21 | variable (R M) in 22 | /-- The support of a module, defined as the set of primes `p` such that `Mₚ ≠ 0`. -/ 23 | def Module.support : Set (PrimeSpectrum R) := 24 | { p | Nontrivial (LocalizedModule p.asIdeal.primeCompl M) } 25 | 26 | lemma Module.mem_support_iff : 27 | p ∈ Module.support R M ↔ Nontrivial (LocalizedModule p.asIdeal.primeCompl M) := Iff.rfl 28 | 29 | lemma Module.not_mem_support_iff : 30 | p ∉ Module.support R M ↔ Subsingleton (LocalizedModule p.asIdeal.primeCompl M) := 31 | not_nontrivial_iff_subsingleton 32 | 33 | lemma Module.not_mem_support_iff' : 34 | p ∉ Module.support R M ↔ ∀ m : M, ∃ r ∉ p.asIdeal, r • m = 0 := by 35 | egg [not_mem_support_iff, LocalizedModule.subsingleton_iff] 36 | --rw [not_mem_support_iff, LocalizedModule.subsingleton_iff] 37 | --rfl 38 | 39 | set_option egg.timeLimit 100 in 40 | set_option egg.slotted true in 41 | set_option egg.reporting true in 42 | lemma Module.mem_support_iff' : 43 | p ∈ Module.support R M ↔ ∃ m : M, ∀ r ∉ p.asIdeal, r • m ≠ 0 := by 44 | -- inifinite loop? 45 | egg neg [not_not , not_mem_support_iff'] 46 | --rw [← @not_not (_ ∈ _), not_mem_support_iff'] 47 | --push_neg 48 | --rfl 49 | -------------------------------------------------------------------------------- /Lean/Egg/Tests/Prune.lean: -------------------------------------------------------------------------------- 1 | import Egg 2 | 3 | -- These test check that redundant rewrites are pruned. 4 | 5 | set_option trace.egg.rewrites.pruned true 6 | set_option linter.unusedVariables false 7 | 8 | /-- 9 | trace: [egg.rewrites.pruned] Pruned (3) 10 | [egg.rewrites.pruned] #0(⇐) by #0 11 | [egg.rewrites.pruned] 0 = 0 12 | [egg.rewrites.pruned] LHS MVars 13 | [] 14 | [egg.rewrites.pruned] RHS MVars 15 | [] 16 | [egg.rewrites.pruned] #1(⇒) by #0 17 | [egg.rewrites.pruned] 0 = 0 18 | [egg.rewrites.pruned] LHS MVars 19 | [] 20 | [egg.rewrites.pruned] RHS MVars 21 | [] 22 | [egg.rewrites.pruned] #1(⇐) by #0 23 | [egg.rewrites.pruned] 0 = 0 24 | [egg.rewrites.pruned] LHS MVars 25 | [] 26 | [egg.rewrites.pruned] RHS MVars 27 | [] 28 | -/ 29 | #guard_msgs in 30 | example (h₁ h₂ : 0 = 0) : 0 = 0 := by 31 | egg [h₁, h₂] 32 | 33 | /-- 34 | trace: [egg.rewrites.pruned] Pruned (4) 35 | [egg.rewrites.pruned] #0(⇐) by #0 36 | [egg.rewrites.pruned] ?m + ?n = ?n + ?m 37 | [egg.rewrites.pruned] LHS MVars 38 | [?n: [unconditionallyVisible], ?m: [unconditionallyVisible]] 39 | [egg.rewrites.pruned] RHS MVars 40 | [?n: [unconditionallyVisible], ?m: [unconditionallyVisible]] 41 | [egg.rewrites.pruned] #1(⇒) by #0 42 | [egg.rewrites.pruned] ?a + ?b = ?b + ?a 43 | [egg.rewrites.pruned] LHS MVars 44 | [?a: [unconditionallyVisible], ?b: [unconditionallyVisible]] 45 | [egg.rewrites.pruned] RHS MVars 46 | [?a: [unconditionallyVisible], ?b: [unconditionallyVisible]] 47 | [egg.rewrites.pruned] #1(⇐) by #0 48 | [egg.rewrites.pruned] ?b + ?a = ?a + ?b 49 | [egg.rewrites.pruned] LHS MVars 50 | [?a: [unconditionallyVisible], ?b: [unconditionallyVisible]] 51 | [egg.rewrites.pruned] RHS MVars 52 | [?a: [unconditionallyVisible], ?b: [unconditionallyVisible]] 53 | [egg.rewrites.pruned] #1↓(⇒) by #0↓ 54 | [egg.rewrites.pruned] ∀ (a b : Nat), a + b = b + a = True 55 | [egg.rewrites.pruned] LHS MVars 56 | [] 57 | [egg.rewrites.pruned] RHS MVars 58 | [] 59 | -/ 60 | #guard_msgs in 61 | example (h : ∀ a b : Nat, a + b = b + a) : 0 = 0 := by 62 | egg [Nat.add_comm, h] 63 | -------------------------------------------------------------------------------- /Lean/Egg/Tactic/Config/Option.lean: -------------------------------------------------------------------------------- 1 | import Egg.Core.Config 2 | import Lean 3 | open Lean 4 | 5 | namespace Egg 6 | 7 | declare_syntax_cat _egg_opt 8 | syntax ident (" : " term:max)? (str)? : _egg_opt 9 | 10 | private def parseOpt : TSyntax `_egg_opt → MacroM (Ident × Term × String) 11 | | `(_egg_opt| $name:ident $[: $ty]? $[$descr]?) => do 12 | let ty := ty.getD <| ← `(Bool) 13 | let descr := descr.map (·.getString) |>.getD "" 14 | return (name, ty, descr) 15 | | _ => unreachable! 16 | 17 | local macro "register_egg_options" opts:_egg_opt* : command => do 18 | let mut regs := #[] 19 | let mut names := #[] 20 | let mut defs := #[] 21 | for opt in opts do 22 | let (name, ty, descr) ← parseOpt opt 23 | let regName := mkIdentFrom name (`egg ++ name.getId) 24 | regs := regs.push <| ← `( 25 | register_option $regName : $ty := { 26 | defValue := ({} : Config).$name 27 | descr := $(quote descr) 28 | } 29 | ) 30 | names := names.push name 31 | defs := defs.push <| mkIdentFrom name (`egg ++ name.getId ++ `get) 32 | let cfgFromOpts ← `( 33 | def $(mkIdent `Config.fromOptions) : MetaM Config := return { 34 | $[$names:ident := $defs (← getOptions)]* 35 | toDebug := {} 36 | } 37 | ) 38 | let cmds := regs.push cfgFromOpts 39 | return ⟨mkNullNode cmds⟩ 40 | 41 | register_egg_options 42 | shapes 43 | betaReduceRws 44 | etaReduceRws 45 | natReduceRws 46 | builtins 47 | structProjs 48 | goalTypeSpec 49 | tcProjs 50 | tcProjFusion 51 | natLit 52 | eta 53 | etaExpand 54 | beta 55 | levels 56 | explosion 57 | derivedGuides 58 | groundEqs 59 | unionSemantics 60 | subgoals 61 | optimizeExpl 62 | timeLimit : Nat "The number of seconds allotted to equality saturation before it aborts. Note 63 | that the total invocation time of `egg` can exceed this time limit as it only 64 | applies to the equality saturation step, and not other steps like equation 65 | generation and proof reconstruction." 66 | nodeLimit : Nat 67 | iterLimit : Nat 68 | reporting 69 | flattenReports 70 | retryWithShapes "When proof reconstruction fails, try running again with `egg.shapes := true`." 71 | explLengthLimit : Nat 72 | -------------------------------------------------------------------------------- /Lean/Egg/Tactic/Rewrite/Rules/Gen/Basic.lean: -------------------------------------------------------------------------------- 1 | import Egg.Core.Gen.Builtins 2 | import Egg.Core.Gen.Tagged 3 | import Egg.Core.Gen.GoalTypeSpecialization 4 | import Egg.Core.Gen.Explosion 5 | import Egg.Core.Gen.TcProjs 6 | import Egg.Core.Gen.StructProjs 7 | import Egg.Tactic.Goal 8 | import Egg.Tactic.Baskets 9 | import Egg.Tactic.Rewrite.Rules.Parse 10 | import Egg.Tactic.Rewrite.Rules.Gen.GenM 11 | import Egg.Tactic.Rewrite.Rules.Gen.Intros 12 | import Egg.Tactic.Trace 13 | 14 | open Lean Meta Elab Tactic 15 | 16 | namespace Egg.Rewrite.Rules 17 | 18 | def gen (goal : Goal) (args : TSyntax `egg_args) (guides : Guides) (cfg : Config) : TacticM Rules := do 19 | withTraceNode `egg.rewrites (fun _ => return "Rewrites") do 20 | let { all, pruned } ← GenM.run core 21 | withTraceNode `egg.rewrites.pruned (fun _ => return m!"Pruned ({pruned.rules.rws.size})") do 22 | pruned.rules.tracePruned pruned.reasons `egg.rewrites.pruned cfg.subgoals 23 | return all 24 | where 25 | genBasic : GenM Rules := do 26 | let basic ← Rules.elab cfg cfg.groundEqs args 27 | for (_, stx) in basic.stxs do 28 | let `($name:ident) := stx | continue 29 | for basket in cfg.baskets do 30 | if ← extension.basketContains basket name.getId then 31 | logWarningAt name m!"This theorem already appears in the egg basket '{basket}'" 32 | return basic 33 | genTcProjs : GenM Rules := open GenM in do 34 | let rws ← allExceptGeneratedGroundEqs 35 | let targets := rws.tcProjTargets ++ goal.tcProjTargets .goal ++ guides.tcProjTargets 36 | genTcProjReductions targets cfg 37 | core : GenM Unit := open GenM in do 38 | generate .intros cfg do genIntros goal.intros.unzip.fst cfg 39 | generate .basic cfg do genBasic 40 | generate .tagged cfg do genTagged cfg 41 | generate .builtins cfg do genBuiltins cfg 42 | -- Note: the order of the following derived rewrites is chosen explicitly. 43 | generate .goalTypeSpec cfg do genGoalTypeSpecializations (← allExceptGeneratedGroundEqs) goal.toCongr cfg.subgoals 44 | generate .explosion cfg do genExplosions (← allExceptGeneratedGroundEqs) cfg.subgoals 45 | generate .tcProj cfg do genTcProjs 46 | generate .structProj cfg do genStructProjRws goal.toCongr (← all) guides cfg 47 | -------------------------------------------------------------------------------- /.github/workflows/release.yml: -------------------------------------------------------------------------------- 1 | name: Create Release 2 | on: 3 | # Allows this workflow to be run manually from the "Actions" tab. 4 | workflow_dispatch: 5 | inputs: 6 | release-tag: 7 | type: string 8 | description: Release Tag 9 | 10 | # https://docs.github.com/en/actions/writing-workflows/choosing-what-your-workflow-does/controlling-permissions-for-github_token 11 | permissions: 12 | contents: write 13 | 14 | # For each step that uses `gh`, we need to set the `GITHUB_TOKEN`: 15 | # https://docs.github.com/en/actions/writing-workflows/choosing-what-your-workflow-does/using-github-cli-in-workflows 16 | 17 | jobs: 18 | create-release: 19 | name: Create Release 20 | runs-on: ubuntu-latest 21 | env: 22 | GITHUB_TOKEN: ${{ secrets.GITHUB_TOKEN }} 23 | TAG: "${{ github.event.inputs.release-tag }}" 24 | steps: 25 | - name: Checkout Repository 26 | uses: actions/checkout@v3 27 | - name: Create Tag and Release 28 | run: | 29 | git tag "$TAG" 30 | git push origin "$TAG" 31 | gh release create "$TAG" --repo="$GITHUB_REPOSITORY" --title="$TAG" --prerelease 32 | upload-artifacts: 33 | # https://docs.github.com/en/actions/writing-workflows/choosing-what-your-workflow-does/running-variations-of-jobs-in-a-workflow#about-matrix-strategies 34 | strategy: 35 | matrix: 36 | # https://docs.github.com/en/actions/writing-workflows/workflow-syntax-for-github-actions#standard-github-hosted-runners-for-public-repositories 37 | os: [macos-latest, macos-13, ubuntu-24.04-arm, ubuntu-latest] # TODO: windows-latest 38 | runs-on: ${{ matrix.os }} 39 | env: 40 | GITHUB_TOKEN: ${{ secrets.GITHUB_TOKEN }} 41 | steps: 42 | - name: Checkout Repository 43 | uses: actions/checkout@v3 44 | - name: Install elan 45 | run: | 46 | set -o pipefail 47 | curl https://raw.githubusercontent.com/leanprover/elan/master/elan-init.sh -sSf | sh -s -- --default-toolchain none -y 48 | "$HOME/.elan/bin/lean" --version 49 | echo "$HOME/.elan/bin" >> $GITHUB_PATH 50 | - name: Build Library 51 | run: lake build 52 | - name: Upload Artifacts 53 | run: lake upload "${{ github.event.inputs.release-tag }}" -------------------------------------------------------------------------------- /Lean/Egg/Tests/mathlib4/mk_mathlib.sh: -------------------------------------------------------------------------------- 1 | #!/usr/bin/env bash 2 | 3 | mathlib_dir="$(realpath -s "$(dirname "$0")")" 4 | mathlib_repo_dir="$mathlib_dir/mathlib4" 5 | 6 | # Sets a variable corresponding to the `--main` flag. 7 | use_main_branch=false 8 | for arg in "$@" 9 | do 10 | if [[ "$arg" == "--main" ]]; then 11 | use_main_branch=true 12 | fi 13 | done 14 | 15 | # Fetches the Lean version used by egg (or uses `master` if the corresponding flag is given). 16 | if [[ "$use_main_branch" == true ]]; then 17 | lean_version="master" 18 | else 19 | lean_toolchain_file="$mathlib_dir/../../../../lean-toolchain" 20 | lean_toolchain="$(cat "$lean_toolchain_file")" 21 | lean_version="${lean_toolchain#*:}" 22 | fi 23 | 24 | # Clones the previously determined version of mathlib. 25 | cd "$mathlib_dir" 26 | git clone -b "$lean_version" --single-branch --depth 1 "https://github.com/leanprover-community/mathlib4.git" || { exit 1; } 27 | cd "$mathlib_repo_dir" || { exit 1; } 28 | git switch -c egg-tests || { exit 1; } 29 | git remote set-url origin "https://github.com/marcusrossel/mathlib4.git" || { exit 1; } 30 | 31 | # Adds egg as a dependency to the `lakefile.lean`. 32 | lakefile="$mathlib_repo_dir/lakefile.lean" 33 | dep="\nrequire egg from \"../../../../..\"" 34 | echo -e "$dep" >> "$lakefile" 35 | 36 | # Sets `linter.style.header` to `false` in `lakefile.lean`. 37 | header_option_true='⟨`linter.style.header, true⟩' 38 | header_option_false='⟨`linter.style.header, false⟩' 39 | sed -i -e "s/$header_option_true/$header_option_false/" "$lakefile" 40 | 41 | # Moves `SimpOnlyOverride.lean` to its intended destination. 42 | simp_only_override="$mathlib_dir/SimpOnlyOverride.lean" 43 | simp_only_override_dst="$mathlib_repo_dir/Mathlib/Testing/Egg" 44 | mkdir "$simp_only_override_dst" 45 | cp "$simp_only_override" "$simp_only_override_dst" 46 | 47 | # Adds an import statement for the simp only override to each `.lean` file, 48 | # except a few select ones. 49 | import_statement="import Mathlib.Testing.Egg.SimpOnlyOverride" 50 | cd "$mathlib_repo_dir" 51 | targets="$(find Mathlib -type f | grep -v -e 'SimpOnlyOverride.lean$' -e '.DS_Store')" || { exit 1; } 52 | 53 | while IFS= read -r file; do 54 | echo "$import_statement"$'\n'$"$(cat $file)" > "$file" 55 | done <<< "$targets" 56 | 57 | # Preparation for `lake build`. 58 | cd "$mathlib_repo_dir" 59 | lake update -------------------------------------------------------------------------------- /Lean/Egg/Core/MVars/Subst.lean: -------------------------------------------------------------------------------- 1 | import Egg.Core.MVars.Basic 2 | import Lean 3 | 4 | open Lean Std Meta 5 | 6 | namespace Egg.MVars 7 | 8 | structure Subst where 9 | expr : HashMap MVarId MVarId 10 | lvl : HashMap LMVarId LMVarId 11 | 12 | instance : EmptyCollection Subst where 13 | emptyCollection := { expr := ∅, lvl := ∅ } 14 | 15 | namespace Subst 16 | 17 | def insertExpr (subst : Subst) (m m' : MVarId) : Subst := 18 | { subst with expr := subst.expr.insert m m' } 19 | 20 | def insertLevel (subst : Subst) (m m' : LMVarId) : Subst := 21 | { subst with lvl := subst.lvl.insert m m' } 22 | 23 | def apply (subst : Subst) (e : Expr) : Expr := 24 | e.replace replaceExpr 25 | where 26 | replaceExpr : Expr → Option Expr 27 | | .mvar id => subst.expr[id]? >>= (Expr.mvar ·) 28 | | .sort lvl => Expr.sort <| lvl.replace replaceLvl 29 | | .const name lvls => Expr.const name <| lvls.map (·.replace replaceLvl) 30 | | _ => none 31 | replaceLvl : Level → Option Level 32 | | .mvar id => subst.lvl[id]? >>= (Level.mvar ·) 33 | | _ => none 34 | 35 | end Subst 36 | 37 | def fresh (mvars : MVars) (init : Subst := {}) : MetaM (MVars × Subst) := do 38 | let mut subst := init 39 | let mut mvars' : MVars := {} 40 | -- Assigns fresh expression mvars. 41 | for (m, ps) in mvars.expr do 42 | if let some m' := subst.expr[m]? then 43 | mvars' := mvars'.insertExpr m' ps 44 | else 45 | let m' := (← mkFreshExprMVar none).mvarId! 46 | subst := subst.insertExpr m m' 47 | mvars' := mvars'.insertExpr m' ps 48 | -- Assigns fresh level mvars. 49 | for (m, ps) in mvars.lvl do 50 | if let some m' := subst.lvl[m]? then 51 | mvars' := mvars'.insertLevel m' ps 52 | else 53 | let m' := (← mkFreshLevelMVar).mvarId! 54 | subst := subst.insertLevel m m' 55 | mvars' := mvars'.insertLevel m' ps 56 | -- As the type of an mvar may also contain mvars, we also have to replace those type-level mvars 57 | -- with their fresh counterpart. We can only do this once we know the fresh counterpart for each 58 | -- mvar, so we postpone the type assignment until now (that is, the type given for 59 | -- `mkFreshExprMVar` above was `none`). 60 | -- TODO: What about mvars contained in the type of the type, etc.? 61 | for (m, _) in mvars.expr do 62 | subst.expr[m]!.setType <| subst.apply (← m.getType) 63 | return (mvars', subst) 64 | -------------------------------------------------------------------------------- /Lean/Egg/Core/Encode/EncodeM.lean: -------------------------------------------------------------------------------- 1 | import Egg.Core.Config 2 | import Egg.Core.Source 3 | 4 | open Lean Std 5 | 6 | namespace Egg 7 | 8 | abbrev Expression := String 9 | 10 | structure EncodeM.State where 11 | config : Config.Encoding 12 | bvars : List FVarId := [] 13 | cache : HashMap Expr Expression := ∅ 14 | 15 | abbrev EncodeM := StateT EncodeM.State MetaM 16 | 17 | namespace EncodeM 18 | 19 | def config : EncodeM Config.Encoding := 20 | State.config <$> get 21 | 22 | -- Note: This only works as intended if `m` does not add any additional bvars (permanently). 23 | def withInstantiatedBVar (ty body : Expr) (m : String → Expr → EncodeM α) : EncodeM α := do 24 | Meta.withLocalDecl .anonymous .default ty fun fvar => do 25 | let s ← get 26 | set { s with bvars := fvar.fvarId! :: s.bvars } 27 | let bvarName := if (← config).slotted then s!"${fvar.fvarId!.uniqueIdx!} " else "" 28 | let a ← m bvarName (body.instantiate #[fvar]) 29 | set { s with bvars := s.bvars } 30 | return a 31 | 32 | def bvarName? (id : FVarId) : EncodeM (Option String) := do 33 | let some bvarIdx := (← get).bvars.idxOf? id | return none 34 | return if (← config).slotted then s!"${id.uniqueIdx!}" else s!"{bvarIdx}" 35 | 36 | open Meta 37 | 38 | def needsProofErasure (e : Expr) : EncodeM Bool := do 39 | isProof e 40 | 41 | def needsInstErasure? (e : Expr) : EncodeM (Option Expr) := do 42 | let ty ← inferType e 43 | -- Note: `isClass?` does not require the type operator to be fully applied. That is, if `ty` is 44 | -- `Inhabited` instead of `Inhabited Nat`, `isClass?` will still succeed. This shouldn't 45 | -- pose a problem though as long as we traverse terms top-down during encoding. That is, as 46 | -- long as we always encounter `Inhabited Nat` before we would encounter just `Inhabited`. 47 | unless (← isClass? ty).isSome do return none 48 | return ty 49 | 50 | def withCache (e : Expr) (m : EncodeM Expression) : EncodeM Expression := do 51 | let s ← get 52 | if let some enc := s.cache[e]? then 53 | return enc 54 | else 55 | let enc ← m 56 | set { s with cache := s.cache.insert e enc } 57 | return enc 58 | 59 | def withoutShapes (m : EncodeM Expression) : EncodeM Expression := do 60 | let s ← get 61 | let shapes := s.config.shapes 62 | set { s with config.shapes := false } 63 | let enc ← m 64 | set { s with config.shapes := shapes } 65 | return enc 66 | -------------------------------------------------------------------------------- /Lean/Egg/Tests/Conditional.lean: -------------------------------------------------------------------------------- 1 | import Egg 2 | 3 | /-- error: egg failed to prove the goal (saturated) -/ 4 | #guard_msgs in 5 | example (h : x ∧ y → 1 = 2) : 1 = 2 := by 6 | egg [h] 7 | 8 | example (h₁ : x ∧ y) (h₂ : x ∧ y → 1 = 2) : 1 = 2 := by 9 | egg [h₁, h₂] 10 | 11 | example (h₁ : x ∧ y) (h₂ : x ∧ y → 1 = 2) : 1 = 2 := by 12 | egg [*] 13 | 14 | example (h₁ : x ∧ y) (h₂ : y ∧ x → 1 = 2) : 1 = 2 := by 15 | egg [and_comm.mp h₁, h₂] 16 | 17 | example (h₁ : x ∧ y) (h₂ : y ∧ x → 1 = 2) : 1 = 2 := by 18 | have h₁' := and_comm.mp h₁ 19 | egg [h₁', h₂] 20 | 21 | example (h₁ : x ∧ y) (h₂ : y ∧ x → 1 = 2) : 1 = 2 := by 22 | have := and_comm.mp h₁ 23 | egg [*] 24 | 25 | example (h₁ : ∀ n, n > 2 → n = x) (h₂ : 3 > 2) : 3 = x := by 26 | egg [h₁, h₂] 27 | 28 | -- This tests that we can handle multiple conditions. 29 | example (h₁ : ∀ n, n > 2 → n > 3 → n = x) (h₂ : 4 > 2) (h₃ : 4 > 3) : 4 = x := by 30 | egg [h₁, h₂, h₃] 31 | 32 | -- This tests that conditions and facts do not need to be in order for proof reconstruction to work. 33 | example (h₁ : ∀ n, n > 2 → n > 3 → n = x) (h₃ : 4 > 3) (h₂ : 4 > 2) : 4 = x := by 34 | egg [h₁, h₂, h₃] 35 | 36 | example {a : Nat} (h : a < b) : a % b = a := by 37 | egg [Nat.mod_eq_of_lt, h] 38 | 39 | -- This test, and the following two, check that adding redundant rewrites and changing their order 40 | -- does not affect the outcome of the tactic. 41 | example {x : Nat} (h₁ : x = y) (h₂ : x = y → 1 = 2) : 1 = 2 := by 42 | egg [h₁, h₂] 43 | 44 | example {x : Nat} (h₁ : x = y) (h₂ : x = y → 1 = 2) : 1 = 2 := by 45 | egg [h₁, h₁, h₂] 46 | 47 | example {x : Nat} (h₁ : x = y) (h₂ : x = y → 1 = 2) : 1 = 2 := by 48 | egg [h₁, h₂, h₁] 49 | 50 | -- NOTE: This requires recovery of weak vars. 51 | example (h : ∀ p : Prop, p → 1 = id 1) : 1 = id 1 := by 52 | egg [h] 53 | 54 | class Fix (α : Type) where 55 | fix : ∀ (f : α → α) (a : α), f a = a 56 | 57 | -- NOTE: This example relies of `egg.goalTypeSpec`. 58 | example [inst : Fix Nat] (f : Nat → Nat) (a : Nat) : f a = a := by 59 | egg [Fix.fix] 60 | 61 | -- This test checks that rewriting of facts is handled correctly. 62 | example {p q r : Prop} (h₁ : p) (h₂ : p ↔ q) (h₃ : q → (p ↔ r)) : p ↔ r := by 63 | egg [h₁, h₂, h₃] 64 | 65 | -- This test checks that multiple rewriting of facts is handled correctly. 66 | example {p q r : Prop} (h₁ : p) (h₂ : p ↔ q) (h₃ : q ↔ r) (h₄ : r → (p ↔ s)) : p ↔ s := by 67 | egg [h₁, h₂, h₃, h₄] 68 | -------------------------------------------------------------------------------- /Lean/Egg/Tests/Explosion.lean: -------------------------------------------------------------------------------- 1 | import Egg 2 | 3 | set_option egg.tcProjs false 4 | set_option egg.builtins false 5 | set_option egg.explosion true 6 | 7 | variable (f : Nat → Nat → Nat) 8 | 9 | set_option trace.egg.rewrites.explosion true 10 | 11 | -- This should not generate exploded rewrites. 12 | /-- trace: [egg.rewrites.explosion] Explosion (0) -/ 13 | #guard_msgs in 14 | example : true = true := by 15 | egg 16 | 17 | -- This should not generate exploded rewrites. 18 | /-- trace: [egg.rewrites.explosion] Explosion (0) -/ 19 | #guard_msgs in 20 | example (h : true = false) : true = false := by 21 | egg [h] 22 | 23 | -- This should not generate exploded rewrites. 24 | /-- trace: [egg.rewrites.explosion] Explosion (0) -/ 25 | #guard_msgs in 26 | example (h : ∀ x y : Nat, f x y = f y x) : f 1 2 = f 2 1 := by 27 | egg [h] 28 | 29 | -- This should not generate exploded rewrites. 30 | /-- trace: [egg.rewrites.explosion] Explosion (0) -/ 31 | #guard_msgs in 32 | example (a b : Nat) (h : ∀ x y : Nat, f x y = f y x) : f a b = f b a := by 33 | egg [h] 34 | 35 | -- This should generate two explosions of `h` - one for `a` and one for `b`. 36 | /-- 37 | trace: [egg.rewrites.explosion] Explosion (2) 38 | [egg.rewrites.explosion] #0💥[3](⇒) 39 | [egg.rewrites.explosion] f ?m.15 ?m.15 = f a ?m.15 40 | [egg.rewrites.explosion] LHS MVars 41 | [?m.15: [unconditionallyVisible]] 42 | [egg.rewrites.explosion] RHS MVars 43 | [?m.15: [unconditionallyVisible]] 44 | [egg.rewrites.explosion] #0💥[4](⇒) 45 | [egg.rewrites.explosion] f ?m.21 ?m.21 = f b ?m.21 46 | [egg.rewrites.explosion] LHS MVars 47 | [?m.21: [unconditionallyVisible]] 48 | [egg.rewrites.explosion] RHS MVars 49 | [?m.21: [unconditionallyVisible]] 50 | -/ 51 | #guard_msgs in 52 | example (a b : Nat) (h : ∀ x y : Nat, f x x = f y x) : f a a = f b a := by 53 | egg [h] 54 | 55 | set_option trace.egg.rewrites.explosion false 56 | 57 | -- TODO: Egg finds a a broken proof paths: by rewriting `f #0 #0` with both `h₁` and `h₂` which 58 | -- establishes `0 = f #0 #0 = 1`. Is there any sensible way to fix this? 59 | 60 | set_option egg.explosion false in 61 | example (a : Nat) (h₁ : ∀ x : Nat, f x x = 0) (h₂ : ∀ x : Nat, f x x = 1) : 0 = 1 := by 62 | sorry -- egg [h₁, h₂] 63 | 64 | set_option egg.explosion true in 65 | example (a : Nat) (h₁ : ∀ x : Nat, f x x = 0) (h₂ : ∀ x : Nat, f x x = 1) : 0 = 1 := by 66 | egg [h₁, h₂] 67 | -------------------------------------------------------------------------------- /Lean/Egg/Tests/Functional.lean: -------------------------------------------------------------------------------- 1 | import Egg 2 | 3 | -- Tests involving basic functional programming. 4 | namespace Functional 5 | 6 | inductive List α 7 | | nil : List α 8 | | cons : α → List α → List α 9 | 10 | def append : List α → List α → List α 11 | | .nil, bs => bs 12 | | .cons a as, bs => .cons a (append as bs) 13 | 14 | theorem append_nil (as : List α) : append as .nil = as := by 15 | induction as with 16 | | nil => egg [append] 17 | | cons _ _ ih => egg [ih, append] 18 | 19 | -- Note: This test contains universe parameters. 20 | theorem append_assoc (as bs cs : List α) : append (append as bs) cs = append as (append bs cs) := by 21 | induction as with 22 | | nil => egg [append] 23 | | cons _ _ ih => egg [ih, append] 24 | 25 | def reverseAux : List α → List α → List α 26 | | .nil, r => r 27 | | .cons a l, r => reverseAux l (.cons a r) 28 | 29 | def reverse (as : List α) : List α := 30 | reverseAux as .nil 31 | 32 | theorem reverseAux_eq_append (as bs : List α) : 33 | reverseAux as bs = append (reverseAux as .nil) bs := by 34 | induction as generalizing bs with 35 | | nil => egg [reverseAux, append] 36 | | cons _ _ ih => egg [reverseAux, append_assoc, ih, append] 37 | 38 | theorem reverse_nil : reverse (.nil : List α) = .nil := by 39 | egg [reverse, reverseAux] 40 | 41 | theorem reverse_cons (a : α) (as : List α) : 42 | reverse (.cons a as) = append (reverse as) (.cons a .nil) := by 43 | egg [reverse, reverseAux, reverseAux_eq_append] 44 | 45 | theorem reverse_append (as bs : List α) : 46 | reverse (append as bs) = append (reverse bs) (reverse as) := by 47 | induction as generalizing bs with 48 | | nil => egg [reverse_nil, append_nil, append] 49 | | cons a as ih => egg [ih, append_assoc, reverse_cons, append] 50 | 51 | def map (f : α → β) : List α → List β 52 | | .nil => .nil 53 | | .cons a as => .cons (f a) (map f as) 54 | 55 | def foldr (f : α → β → β) (init : β) : List α → β 56 | | .nil => init 57 | | .cons a l => f a (foldr f init l) 58 | 59 | def all (p : α → Bool) (xs : List α) : Bool := 60 | foldr and true (map p xs) 61 | 62 | def all' (p : α → Bool) : List α → Bool 63 | | .nil => true 64 | | .cons x xs => (p x) && (all' p xs) 65 | 66 | theorem all_deforestation (p : α → Bool) (xs : List α) : all p xs = all' p xs := by 67 | induction xs with 68 | | nil => egg [all, all', foldr, map] 69 | | cons _ _ ih => egg [all, all', foldr, map, ih] 70 | -------------------------------------------------------------------------------- /Lean/Egg/Tests/Freshman.lean: -------------------------------------------------------------------------------- 1 | import Egg 2 | 3 | class CommRing (α) extends Zero α, One α, Add α, Sub α, Mul α, Div α, Pow α Nat, Inv α, Neg α where 4 | comm_add (a b : α) : a + b = b + a 5 | comm_mul (a b : α) : a * b = b * a 6 | add_assoc (a b c : α) : a + (b + c) = (a + b) + c 7 | mul_assoc (a b c : α) : a * (b * c) = (a * b) * c 8 | sub_canon (a b : α) : a - b = a + -b 9 | neg_add (a : α) : a + -a = 0 10 | div_canon (a b : α) : a / b = a * b⁻¹ 11 | zero_add (a : α) : a + 0 = a 12 | zero_mul (a : α) : a * 0 = 0 13 | one_mul (a : α) : a * 1 = a 14 | distrib (a b c : α) : a * (b + c) = (a * b) + (a * c) 15 | pow_zero (a : α) : a ^ 0 = 1 16 | pow_one (a : α) : a ^ 1 = a 17 | pow_two (a : α) : a ^ 2 = (a ^ 1) * a 18 | pow_three (a : α) : a ^ 3 = (a ^ 2) * a 19 | 20 | class CharTwoRing (α) extends CommRing α where 21 | char_two (a : α) : a + a = 0 22 | 23 | open CommRing CharTwoRing Egg.Guides Egg.Config.Modifier in 24 | macro "char_two_ring" mod:egg_cfg_mod guides:(egg_guides)? : tactic => `(tactic| 25 | egg $mod:egg_cfg_mod [comm_add, comm_mul, add_assoc, mul_assoc, sub_canon, neg_add, 26 | div_canon, zero_add, zero_mul, one_mul, distrib, pow_zero, pow_one, pow_two, pow_three, char_two] 27 | $[$guides]? 28 | ) 29 | 30 | variable [CharTwoRing α] (x y : α) 31 | 32 | theorem freshmans_dream₂ : (x + y) ^ 2 = (x ^ 2) + (y ^ 2) := by 33 | calc (x + y) ^ 2 34 | _ = (x + y) * (x + y) := by char_two_ring 35 | _ = x * (x + y) + y * (x + y) := by char_two_ring 36 | _ = x ^ 2 + x * y + y * x + y ^ 2 := by char_two_ring 37 | _ = x ^ 2 + y ^ 2 := by char_two_ring 38 | 39 | theorem freshmans_dream₂' : (x + y) ^ 2 = (x ^ 2) + (y ^ 2) := by 40 | char_two_ring 41 | 42 | theorem freshmans_dream₃ : (x + y) ^ 3 = x ^ 3 + x * y ^ 2 + x ^ 2 * y + y ^ 3 := by 43 | calc (x + y) ^ 3 44 | _ = (x + y) * (x + y) * (x + y) := by char_two_ring 45 | _ = (x + y) * (x * (x + y) + y * (x + y)) := by char_two_ring 46 | _ = (x + y) * (x ^ 2 + x * y + y * x + y ^ 2) := by char_two_ring 47 | _ = (x + y) * (x ^ 2 + y ^ 2) := by char_two_ring 48 | _ = x * (x ^ 2 + y ^ 2) + y * (x ^ 2 + y ^ 2) := by char_two_ring 49 | _ = (x * x ^ 2) + x * y ^ 2 + y * x ^ 2 + y * y ^ 2 := by char_two_ring 50 | _ = x ^ 3 + x * y ^ 2 + x ^ 2 * y + y ^ 3 := by char_two_ring 51 | 52 | -- NOTE: The explanation is too long. 53 | -- theorem freshmans_dream₃' : (x + y) ^ 3 = x ^ 3 + x * y ^ 2 + x ^ 2 * y + y ^ 3 := by 54 | -- char_two_ring 55 | -------------------------------------------------------------------------------- /lakefile.lean: -------------------------------------------------------------------------------- 1 | import Lake 2 | open Lake DSL 3 | 4 | require "nomeata" / "calcify" @ git "master" 5 | 6 | -- Cf. https://github.com/lurk-lab/RustFFI.lean 7 | 8 | package egg where 9 | srcDir := "Lean" 10 | -- See https://github.com/leanprover/lean4/tree/master/src/lake#github-release-builds 11 | preferReleaseBuild := true 12 | releaseRepo? := none 13 | buildArchive? := none 14 | 15 | @[default_target] 16 | lean_lib Egg where 17 | -- This enables the interpreter to run functions marked `@[extern]`. 18 | precompileModules := true 19 | 20 | target importTarget pkg : System.FilePath := 21 | pkg.afterBuildCacheAsync do 22 | let oFile := pkg.buildDir / "c" / "ffi.o" 23 | let srcJob ← inputTextFile <| pkg.dir / "C" / "ffi.c" 24 | buildFileAfterDep oFile srcJob fun srcFile => do 25 | let flags := #["-I", toString (← getLeanIncludeDir), "-fPIC"] 26 | compileO oFile srcFile flags 27 | 28 | extern_lib ffi pkg := do 29 | let job ← fetch <| pkg.target ``importTarget 30 | let libFile := pkg.sharedLibDir / nameToStaticLib "ffi" 31 | buildStaticLib libFile #[job] 32 | 33 | extern_lib egg_for_lean pkg := do 34 | pkg.afterBuildCacheAsync do 35 | let name := nameToStaticLib "egg_for_lean" 36 | let srcPath := pkg.dir / "Rust" / "Egg" / "target" / "release" / name 37 | let tgtPath := pkg.sharedLibDir / name 38 | let traceFile := pkg.buildDir / "rust" / "egg.trace" 39 | let _ ← buildUnlessUpToDate? traceFile (← getTrace) traceFile do 40 | proc { 41 | cmd := "cargo", 42 | args := #["rustc", "--release", "--", "-C", "relocation-model=pic"], 43 | cwd := pkg.dir / "Rust" / "Egg" 44 | } 45 | IO.FS.createDirAll pkg.sharedLibDir 46 | IO.FS.writeBinFile tgtPath (← IO.FS.readBinFile srcPath) 47 | return pure tgtPath 48 | 49 | extern_lib slotted_for_lean pkg := do 50 | pkg.afterBuildCacheAsync do 51 | let name := nameToStaticLib "slotted_for_lean" 52 | let srcPath := pkg.dir / "Rust" / "Slotted" / "target" / "release" / name 53 | let tgtPath := pkg.sharedLibDir / name 54 | let traceFile := pkg.buildDir / "rust" / "slotted.trace" 55 | let _ ← buildUnlessUpToDate? traceFile (← getTrace) traceFile do 56 | proc { 57 | cmd := "cargo", 58 | args := #["rustc", "--release", "--", "-C", "relocation-model=pic"], 59 | cwd := pkg.dir / "Rust" / "Slotted" 60 | } 61 | IO.FS.createDirAll pkg.sharedLibDir 62 | IO.FS.writeBinFile tgtPath (← IO.FS.readBinFile srcPath) 63 | return pure tgtPath 64 | 65 | @[test_driver] 66 | lean_exe TestDriver where 67 | srcDir := "Egg/Tests" 68 | -------------------------------------------------------------------------------- /Rust/Slotted/src/rewrite.rs: -------------------------------------------------------------------------------- 1 | use std::collections::HashSet; 2 | use slotted_egraphs::*; 3 | use crate::result::*; 4 | use crate::lean_expr::*; 5 | use crate::analysis::*; 6 | 7 | pub struct RewriteTemplate { 8 | pub name: String, 9 | pub lhs: Pattern, 10 | pub rhs: Pattern, 11 | pub _conds: Vec> 12 | } 13 | 14 | fn slots_for_node(e: &LeanExpr) -> HashSet { 15 | match e { 16 | LeanExpr::BVar(s) | LeanExpr::Lam(s, _, _) | LeanExpr::Forall(s, _, _) => HashSet::from([*s]), 17 | _ => Default::default() 18 | } 19 | } 20 | 21 | fn private_slots(p: &Pattern) -> HashSet { 22 | let mut result: HashSet = Default::default(); 23 | let mut queue: Vec<&Pattern> = vec![p]; 24 | while let Some(next) = queue.pop() { 25 | match next { 26 | Pattern::ENode(node, children) => { 27 | result.extend(slots_for_node(node)); 28 | queue.extend(children); 29 | }, 30 | _ => continue 31 | } 32 | } 33 | result 34 | } 35 | 36 | fn subst_is_valid(subst: &Subst, illegal_slots: &HashSet) -> bool { 37 | for i in subst.values() { 38 | for s in i.slots() { 39 | if illegal_slots.contains(&s) { return false } 40 | } 41 | } 42 | true 43 | } 44 | 45 | pub fn templates_to_rewrites(templates: Vec, _subgoals: bool) -> Res> { 46 | let mut result: Vec = vec![]; 47 | for template in templates { 48 | let lhs_search = template.lhs.clone(); 49 | let lhs_apply = template.lhs.clone(); 50 | 51 | let mut illegal_slots = private_slots(&template.lhs); 52 | illegal_slots.extend(&private_slots(&template.rhs)); 53 | 54 | let rw = RewriteT { 55 | searcher: Box::new(move |graph| { ematch_all(graph, &lhs_search) }), 56 | applier: Box::new(move |substs, graph| { 57 | for subst in substs { 58 | if !subst_is_valid(&subst, &illegal_slots) { continue; } 59 | 60 | let lhs = pattern_subst(graph, &lhs_apply, &subst); 61 | let analysis: &LeanAnalysis = graph.analysis_data(lhs.id); 62 | 63 | // Disallows rewriting on primitive e-nodes. 64 | if analysis.is_primitive { continue } 65 | 66 | graph.union_instantiations(&template.lhs, &template.rhs, &subst, Some(template.name.clone())); 67 | } 68 | }), 69 | }; 70 | result.push(rw.into()); 71 | } 72 | Ok(result) 73 | } -------------------------------------------------------------------------------- /Lean/Egg/Core/Rewrite/Condition.lean: -------------------------------------------------------------------------------- 1 | import Egg.Core.MVars.Subst 2 | import Egg.Core.MVars.Collect 3 | open Lean Meta 4 | 5 | namespace Egg.Rewrite 6 | 7 | protected structure MVars where 8 | lhs : MVars 9 | rhs : MVars 10 | deriving Inhabited 11 | 12 | namespace Condition 13 | 14 | inductive Kind where 15 | | proof 16 | | tcInst 17 | deriving BEq 18 | 19 | def Kind.isProof : Kind → Bool 20 | | proof => true 21 | | tcInst => false 22 | 23 | def Kind.isTcInst : Kind → Bool 24 | | proof => false 25 | | tcInst => true 26 | 27 | structure _root_.Egg.Rewrite.Condition where 28 | kind : Kind 29 | mvar : MVarId 30 | type : Expr 31 | 32 | inductive Result where 33 | | none 34 | | synthesized 35 | | unsynthesizable 36 | | some (cond : Condition) 37 | 38 | def from? (mvar : MVarId) (lhs : MVars) : MetaM Result := do 39 | let some kind ← kind? | return .none 40 | -- Note: It seems `inferType` does not instantiate mvars. 41 | let type ← instantiateMVars (← inferType <| .mvar mvar) 42 | -- This is a small optimization. If a type class can already be synthesized, we do it immediately 43 | -- and don't generate a condition to be synthesized during eqsat. 44 | match ← trySynthesizeTcInst type kind with 45 | | .undef => return .some { kind, mvar, type } 46 | | .true => return .synthesized 47 | | .false => return .unsynthesizable 48 | where 49 | -- If the mvar appears in the LHS, then it's a condition only if it's a nested instance or proof. 50 | -- If it does not appear in the LHS, it's a condition immediately if it's an instance or proof. 51 | kind? : MetaM <| Option Kind := do 52 | if let some ps := lhs.expr[mvar]? then 53 | if ps.contains .isTcInst && ps.contains .inTcInstTerm then 54 | return some .tcInst 55 | else if ps.contains .isProof && ps.contains .inProofTerm then 56 | return some .proof 57 | else if ← isTCInstance (.mvar mvar) then 58 | return some .tcInst 59 | else if ← isProof (.mvar mvar) then 60 | return some .proof 61 | return none 62 | trySynthesizeTcInst (type : Expr) : Kind → MetaM LBool 63 | | .proof => return .undef 64 | | .tcInst => do 65 | if type.hasMVar then return .undef 66 | let some inst ← synthInstance? type | return .false 67 | unless ← isDefEq (.mvar mvar) inst do 68 | throwError "Internal error in 'Egg.Rewrite.Condition.from?.synthesizeTcInst?'" 69 | return .true 70 | 71 | def fresh (cond : Condition) (init : MVars.Subst) : MetaM (Condition × MVars.Subst) := do 72 | let (_, subst) ← (← MVars.collect <| .mvar cond.mvar).fresh init 73 | let fresh := { cond with 74 | mvar := subst.expr[cond.mvar]! 75 | type := subst.apply cond.type 76 | } 77 | return (fresh, subst) 78 | -------------------------------------------------------------------------------- /Lean/Egg/Tests/Freshman Tags.lean: -------------------------------------------------------------------------------- 1 | import Egg 2 | 3 | class CommRing (α) extends Zero α, One α, Add α, Sub α, Mul α, Div α, Pow α Nat, Inv α, Neg α where 4 | comm_add (a b : α) : a + b = b + a 5 | comm_mul (a b : α) : a * b = b * a 6 | add_assoc (a b c : α) : a + (b + c) = (a + b) + c 7 | mul_assoc (a b c : α) : a * (b * c) = (a * b) * c 8 | sub_canon (a b : α) : a - b = a + -b 9 | neg_add (a : α) : a + -a = 0 10 | div_canon (a b : α) : a / b = a * b⁻¹ 11 | zero_add (a : α) : a + 0 = a 12 | zero_mul (a : α) : a * 0 = 0 13 | one_mul (a : α) : a * 1 = a 14 | distrib (a b c : α) : a * (b + c) = (a * b) + (a * c) 15 | pow_zero (a : α) : a ^ 0 = 1 16 | pow_one (a : α) : a ^ 1 = a 17 | pow_two (a : α) : a ^ 2 = (a ^ 1) * a 18 | pow_three (a : α) : a ^ 3 = (a ^ 2) * a 19 | 20 | attribute [egg comm_ring] CommRing.comm_add 21 | attribute [egg comm_ring] CommRing.comm_mul 22 | attribute [egg comm_ring] CommRing.add_assoc 23 | attribute [egg comm_ring] CommRing.mul_assoc 24 | attribute [egg comm_ring] CommRing.sub_canon 25 | attribute [egg comm_ring] CommRing.neg_add 26 | attribute [egg comm_ring] CommRing.div_canon 27 | attribute [egg comm_ring] CommRing.zero_add 28 | attribute [egg comm_ring] CommRing.zero_mul 29 | attribute [egg comm_ring] CommRing.one_mul 30 | attribute [egg comm_ring] CommRing.distrib 31 | attribute [egg comm_ring] CommRing.pow_zero 32 | attribute [egg comm_ring] CommRing.pow_one 33 | attribute [egg comm_ring] CommRing.pow_two 34 | attribute [egg comm_ring] CommRing.pow_three 35 | 36 | class CharTwoRing (α) extends CommRing α where 37 | char_two (a : α) : a + a = 0 38 | 39 | variable [CharTwoRing α] (x y : α) 40 | 41 | theorem freshmans_dream₂ : (x + y) ^ 2 = (x ^ 2) + (y ^ 2) := by 42 | egg +comm_ring calc 43 | _ = (x + y) * (x + y) 44 | _ = x * (x + y) + y * (x + y) 45 | _ = x ^ 2 + x * y + y * x + y ^ 2 46 | _ = x ^ 2 + y ^ 2 with [CharTwoRing.char_two] 47 | 48 | theorem freshmans_dream₂' : (x + y) ^ 2 = (x ^ 2) + (y ^ 2) := by 49 | egg +comm_ring [CharTwoRing.char_two] 50 | 51 | theorem freshmans_dream₃ : (x + y) ^ 3 = x ^ 3 + x * y ^ 2 + x ^ 2 * y + y ^ 3 := by 52 | egg +comm_ring calc [CharTwoRing.char_two] 53 | _ = (x + y) * (x + y) * (x + y) 54 | _ = (x + y) * (x * (x + y) + y * (x + y)) 55 | _ = (x + y) * (x ^ 2 + x * y + y * x + y ^ 2) 56 | _ = (x + y) * (x ^ 2 + y ^ 2) 57 | _ = x * (x ^ 2 + y ^ 2) + y * (x ^ 2 + y ^ 2) 58 | _ = (x * x ^ 2) + x * y ^ 2 + y * x ^ 2 + y * y ^ 2 59 | _ = x ^ 3 + x * y ^ 2 + x ^ 2 * y + y ^ 3 60 | 61 | -- NOTE: The explanation is too long. 62 | -- theorem freshmans_dream₃' : (x + y) ^ 3 = x ^ 3 + x * y ^ 2 + x ^ 2 * y + y ^ 3 := by 63 | -- egg +comm_ring [CharTwoRing.char_two] 64 | -------------------------------------------------------------------------------- /Lean/Egg/Core/Gen/Explosion.lean: -------------------------------------------------------------------------------- 1 | import Egg.Core.Rewrite.Rule 2 | import Lean 3 | open Lean Meta 4 | 5 | private def Lean.Meta.findLocalDeclWithTypeMinIdx? (type : Expr) (minIdx : Nat) : 6 | MetaM <| Option (FVarId × Nat) := do 7 | let decls := (← getLCtx).decls 8 | for h : idx in [minIdx:decls.size] do 9 | let some decl := decls[idx] | continue 10 | if ← isDefEq type decl.type then return some (decl.fvarId, idx) 11 | return none 12 | 13 | namespace Egg 14 | 15 | private partial def exprReferencesMVar (e : Expr) (m : MVarId) : MetaM Bool := do 16 | if e.isSort then 17 | return false 18 | else 19 | match e with 20 | | .mvar id => if id == m then return true 21 | | .app e₁ e₂ | .lam _ e₁ e₂ _ | .forallE _ e₁ e₂ _ => 22 | if ← exprReferencesMVar e₁ m <||> exprReferencesMVar e₂ m then return true 23 | | _ => pure () 24 | exprReferencesMVar (← inferType e) m 25 | 26 | private partial def genExplosionsForRw 27 | (rw : Rewrite) (id : Rewrite.Rule.Id) (subgoals : Bool) : MetaM Rewrite.Rules := do 28 | let mut missing := rw.mvars.rhs.visibleExpr.eraseMany rw.mvars.lhs.visibleExpr 29 | for violation in ← rw.violations subgoals do 30 | if let .lhsSingleMVar m := violation then 31 | missing := missing.insert m 32 | break 33 | if missing.isEmpty then return ∅ 34 | let ordered ← missing.toList.qsortM fun m₁ m₂ => return !(← exprReferencesMVar (.mvar m₁) m₂) 35 | core rw ordered [] 36 | where 37 | core (rw : Rewrite) (missing : List MVarId) (loc : List Nat) : MetaM Rewrite.Rules := do 38 | let m :: miss := missing | 39 | let id := { id with src := .explosion id.src loc } 40 | return {⟨id, ← rw.fresh⟩} 41 | let mut minIdx := 0 42 | let mut explosions := ∅ 43 | while minIdx < (← getLCtx).decls.size do 44 | let (fresh, subst) ← rw.freshWithSubst 45 | let m := subst.expr[m]! 46 | let some (fvar, idx) ← findLocalDeclWithTypeMinIdx? (← m.getType) minIdx | break 47 | minIdx := idx + 1 48 | unless ← isDefEq (.mvar m) (.fvar fvar) do throwError "egg: internal error in explosion gen" 49 | let fresh ← fresh.instantiateMVars 50 | unless fresh.conds.unsynthesizable.isEmpty do continue 51 | let miss := miss.filterMap fun i => 52 | let i' := subst.expr[i]! 53 | if fresh.mvars.lhs.expr.contains i' || fresh.mvars.rhs.expr.contains i' then i' else none 54 | explosions := explosions ∪ (← core fresh miss <| loc ++ [minIdx]) 55 | return explosions 56 | 57 | def genExplosions (rules : Rewrite.Rules) (subgoals : Bool) : MetaM Rewrite.Rules := do 58 | let mut explosions := ∅ 59 | for rule in rules.entries do 60 | let new ← genExplosionsForRw rule.rw rule.id subgoals 61 | explosions := explosions ∪ new 62 | return explosions 63 | -------------------------------------------------------------------------------- /Lean/Egg/Tests/NatLit.lean: -------------------------------------------------------------------------------- 1 | import Egg 2 | 3 | -- Tests involving conversions between `Nat.zero` and `Nat.succ _` and `.lit (.natVal _)`. 4 | 5 | set_option egg.natLit true 6 | 7 | example : 0 = Nat.zero := by 8 | egg 9 | 10 | example : 1 = Nat.succ 0 := by 11 | egg 12 | 13 | example : Nat.succ 1 = Nat.succ (Nat.succ Nat.zero) := by 14 | egg 15 | 16 | example : Int.ofNat (Nat.succ 1) = Int.ofNat (Nat.succ (Nat.succ Nat.zero)) := by 17 | egg 18 | 19 | example (h : ∀ n, Nat.succ n = n + 1) : 1 = Nat.zero + 1 := by 20 | egg [h] 21 | 22 | example : 1 = Nat.zero + 1 := by 23 | egg [Nat.succ_eq_add_one] 24 | 25 | elab "app" n:num fn:ident arg:term : term => open Lean.Elab.Term in do 26 | let fn ← elabTerm fn none 27 | let rec go (n : Nat) := if n = 0 then elabTerm arg none else return .app fn <| ← go (n - 1) 28 | go n.getNat 29 | 30 | set_option egg.timeLimit 10 in 31 | set_option egg.builtins false in 32 | example : (app 70 Nat.succ (nat_lit 0)) = (nat_lit 70) := by egg 33 | 34 | -- Note: This produces a gigantic proof. 35 | example (f : Nat → Nat) (h : ∀ x, f x = x.succ) : 30 = app 30 f 0 := by 36 | egg [h] 37 | 38 | -- TODO: If nat-lit is enabled, then all symbols contained in nat-lit rewrites should always be active. 39 | -- However, this means that activations determined in the backend are not respected. 40 | -- Should we determine activations in the frontend afterall? 41 | 42 | example : 12345 + 67890 = 80235 := by 43 | egg 44 | 45 | example : 12345 - 67890 = 0 := by 46 | egg 47 | 48 | example : 67890 - 12345 = 55545 := by 49 | egg 50 | 51 | example : 12345 * 67890 = 838102050 := by 52 | egg 53 | 54 | example : 1234 ^ 5 = 2861381721051424 := by 55 | egg 56 | 57 | example : 12345 / 67890 = 0 := by 58 | egg 59 | 60 | example : 67890 / 12345 = 5 := by 61 | egg 62 | 63 | example : 12345 / 0 = 0 := by 64 | egg 65 | 66 | example : 67890 % 12345 = 6165 := by 67 | egg 68 | 69 | example : 12345 % 67890 = 12345 := by 70 | egg 71 | 72 | example : 12345 % 0 = 12345 := by 73 | egg 74 | 75 | set_option egg.natLit false in 76 | set_option egg.natReduceRws false in 77 | /-- error: egg failed to prove the goal (saturated) -/ 78 | #guard_msgs in 79 | example (h : ∀ f : Nat → Nat, f (1 + 1) = x) : id 2 = x := by 80 | egg [h] 81 | 82 | example (h : ∀ f : Nat → Nat, f (1 + 1) = x) : id 2 = x := by 83 | egg [h] 84 | 85 | example (h : ∀ f : Nat → Nat, f (3 - 2) = x) : id 1 = x := by 86 | egg [h] 87 | 88 | example (h : ∀ f : Nat → Nat, f (2 * 3) = x) : id 6 = x := by 89 | egg [h] 90 | 91 | example (h : ∀ f : Nat → Nat, f (4 / 2) = x) : id 2 = x := by 92 | egg [h] 93 | 94 | example (h : ∀ f : Nat → Nat, f (5 % 3) = x) : id 2 = x := by 95 | egg [h] 96 | 97 | example (h : ∀ f : Nat → Nat, f (2 ^ 3) = x) : id 8 = x := by 98 | egg [h] 99 | -------------------------------------------------------------------------------- /Lean/Egg/Tests/Freshman Pow.lean: -------------------------------------------------------------------------------- 1 | import Egg 2 | 3 | class CommRing (α) extends Zero α, One α, Add α, Sub α, Mul α, Div α, Pow α Nat, Inv α, Neg α where 4 | comm_add (a b : α) : a + b = b + a 5 | comm_mul (a b : α) : a * b = b * a 6 | add_assoc (a b c : α) : a + (b + c) = (a + b) + c 7 | mul_assoc (a b c : α) : a * (b * c) = (a * b) * c 8 | sub_canon (a b : α) : a - b = a + -b 9 | neg_add (a : α) : a + -a = 0 10 | div_canon (a b : α) : a / b = a * b⁻¹ 11 | zero_add (a : α) : a + 0 = a 12 | zero_mul (a : α) : a * 0 = 0 13 | one_mul (a : α) : a * 1 = a 14 | distrib (a b c : α) : a * (b + c) = (a * b) + (a * c) 15 | pow_zero (a : α) : a ^ 0 = 1 16 | pow_succ (a : α) : a ^ (n + 1) = a * (a ^ n) 17 | 18 | class CharTwoRing (α) extends CommRing α where 19 | char_two (a : α) : a + a = 0 20 | 21 | open CommRing CharTwoRing Egg.Guides Egg.Config.Modifier in 22 | macro "char_two_ring" mod:egg_cfg_mod guides:(egg_guides)? : tactic => `(tactic| 23 | egg $mod:egg_cfg_mod [comm_add, comm_mul, add_assoc, mul_assoc, sub_canon, neg_add, 24 | div_canon, zero_add, zero_mul, one_mul, distrib, pow_zero, pow_succ, char_two, 25 | Nat.succ_eq_add_one] $[$guides]? 26 | ) 27 | 28 | variable [CharTwoRing α] (x y : α) 29 | 30 | theorem freshmans_dream₂ : (x + y) ^ 2 = x ^ 2 + y ^ 2 := by 31 | calc (x + y) ^ 2 32 | _ = (x + y) * (x + y) := by char_two_ring 33 | _ = x * (x + y) + y * (x + y) := by char_two_ring 34 | -- _ = x ^ 2 + x * y + y * x + y ^ 2 := by char_two_ring 35 | _ = x ^ 2 + y ^ 2 := by char_two_ring 36 | 37 | theorem freshmans_dream₂' : (x + y) ^ 2 = x ^ 2 + y ^ 2 := by 38 | char_two_ring 39 | 40 | set_option egg.timeLimit 10 in 41 | set_option egg.explLengthLimit 300 in 42 | theorem freshmans_dream₃ : (x + y) ^ 3 = x ^ 3 + x * y ^ 2 + x ^ 2 * y + y ^ 3 := by 43 | calc (x + y) ^ 3 44 | _ = (x + y) * (x + y) * (x + y) := by char_two_ring 45 | -- _ = (x + y) * (x * (x + y) + y * (x + y)) := by char_two_ring 46 | _ = (x + y) * (x ^ 2 + x * y + y * x + y ^ 2) := by char_two_ring 47 | _ = (x + y) * (x ^ 2 + y ^ 2) := by char_two_ring 48 | _ = x * (x ^ 2 + y ^ 2) + y * (x ^ 2 + y ^ 2) := by char_two_ring 49 | _ = (x * x ^ 2) + x * y ^ 2 + y * x ^ 2 + y * y ^ 2 := by char_two_ring 50 | _ = x ^ 3 + x * y ^ 2 + x ^ 2 * y + y ^ 3 := by char_two_ring 51 | 52 | -- NOTE: The explanation is too long. 53 | -- theorem freshmans_dream₃' : (x + y) ^ 3 = x ^ 3 + x * y ^ 2 + x ^ 2 * y + y ^ 3 := by 54 | -- char_two_ring using (x + y) * (x + y) 55 | 56 | -- NOTE: The explanation is too long. 57 | -- theorem freshmans_dream₃'' : (x + y) ^ 3 = x ^ 3 + x * y ^ 2 + x ^ 2 * y + y ^ 3 := by 58 | -- char_two_ring 59 | -------------------------------------------------------------------------------- /Lean/Egg/Tests/Math/Math/Rotman.lean: -------------------------------------------------------------------------------- 1 | import Mathlib 2 | import Egg 3 | 4 | open Nat 5 | 6 | attribute [egg cast] cast_zero cast_one cast_succ cast_pred cast_add_one cast_add cast_sub cast_mul 7 | cast_div cast_pow cast_dvd_cast cast_id cast_ite 8 | 9 | attribute [egg real] 10 | /- + -/ add_comm add_assoc add_zero 11 | /- - -/ sub_zero zero_sub sub_self 12 | /- * -/ mul_comm mul_assoc mul_zero mul_one 13 | /- / -/ div_one zero_div 14 | /- + / -/ add_div 15 | /- * / -/ mul_div_mul_left mul_div_mul_right mul_div_mul_comm div_mul_div_cancel 16 | _root_.div_mul_div_comm div_mul_eq_div_mul_one_div 17 | /- + - -/ sub_sub sub_add add_sub add_comm_sub sub_add_cancel sub_add_eq_add_sub add_sub_assoc 18 | /- + * / -/ left_distrib right_distrib add_div_eq_mul_add_div 19 | 20 | -- TODO: Pruning of rewrites seems to not be working properly. 21 | set_option maxHeartbeats 300000 22 | set_option egg.subgoals true 23 | set_option egg.timeLimit 10 24 | 25 | notation:100 r "﹗" => Real.Gamma (r + 1) 26 | 27 | theorem proposition_1_15 {n r : Nat} (h : n ≥ r) : n.choose r = (n !) / (r ! * (n - r)!) := by 28 | induction n generalizing r 29 | case zero => egg [choose, factorial, le_zero.mp h] 30 | case succ n ih => 31 | by_cases hr : r = 0 <;> try subst hr 32 | all_goals try by_cases hn : r = n + 1 <;> try subst hn 33 | all_goals try simp; rw [Nat.div_self <| factorial_pos _] 34 | have ho : (n - r + 1) = n - (r - 1) := by omega 35 | 36 | calc 37 | _ = (n !) / ((r - 1)! * (n - r + 1)!) + (n !) / (r ! * (n - r)!) := by egg [choose_succ_left, ih, ho] <;> omega 38 | _ = _ := cast_inj (R := Real) |>.mp ?_ 39 | 40 | egg +cast +real calc [Real.Gamma_nat_eq_factorial, Real.Gamma_add_one] 41 | ↑(n ! / ((r - 1)! * (n - r + 1)!) + n ! / (r ! * (n - r)!)) 42 | _ = n﹗ / ((r - 1)﹗ * (n - r + 1)﹗) + n﹗ / (r﹗ * (n - r)﹗) 43 | _ = n﹗ / ((r - 1)﹗ * (n - r)﹗) * (1 / (n - r + 1) + 1 / r) 44 | _ = n﹗ / ((r - 1)﹗ * (n - r)﹗) * ((r + (n - r + 1)) / (r * (n - r + 1))) 45 | _ = n﹗ / ((r - 1)﹗ * (n - r)﹗) * ((n + 1) / (r * (n - r + 1))) 46 | _ = (n + 1)﹗ / (r﹗ * (n + 1 - r)﹗) 47 | _ = ↑((n + 1)! / (r ! * (n + 1 - r)!)) 48 | 49 | all_goals try first | (norm_cast; done) | (norm_cast; omega) 50 | · rw [cast_ne_zero]; exact mul_ne_zero (factorial_ne_zero _) (factorial_ne_zero _) 51 | · exact factorial_mul_factorial_dvd_factorial h 52 | · rw [←cast_one, ←cast_sub (by omega : r ≤ n), ←cast_add, cast_ne_zero]; omega 53 | · rw [←cast_sub (by omega : r ≤ n), Real.Gamma_nat_eq_factorial, Real.Gamma_nat_eq_factorial, ←cast_mul, cast_ne_zero]; exact mul_ne_zero (factorial_ne_zero _) (factorial_ne_zero _) 54 | · rw [cast_ne_zero]; exact mul_ne_zero (factorial_ne_zero _) (factorial_ne_zero _) 55 | · exact ho ▸ factorial_mul_factorial_dvd_factorial (sub_le_of_le_add h) 56 | · rw [cast_ne_zero]; exact mul_ne_zero (factorial_ne_zero _) (factorial_ne_zero _) 57 | · exact factorial_mul_factorial_dvd_factorial <| by omega 58 | -------------------------------------------------------------------------------- /Lean/Egg/Tests/TestDriver.lean: -------------------------------------------------------------------------------- 1 | import Lake 2 | open IO 3 | 4 | -- See: https://leanprover.zulipchat.com/#narrow/channel/217875-Is-there-code-for-X.3F/topic/String.2EisInfixOf 5 | def String.isInfixOf (needle : String) (hey : String) := Id.run do 6 | let mut i := hey.mkIterator 7 | while not i.atEnd do 8 | if needle.isPrefixOf i.remainingToString 9 | then return true 10 | else i := i.next 11 | return false 12 | 13 | def ciFlag := "ci" 14 | 15 | def testName (file : FS.DirEntry) : String := 16 | file.fileName.stripSuffix ".lean" 17 | 18 | def excludeFromTests (file : FS.DirEntry) : Bool := 19 | testName file == "TestDriver" || 20 | file.fileName.startsWith "WIP" 21 | 22 | def getTests : IO (Array FS.DirEntry) := do 23 | let testsDir := (← currentDir) / "Lean" / "Egg" / "Tests" 24 | let entries ← testsDir.readDir 25 | let files ← entries.filterM fun e => return !(← e.path.isDir) 26 | let leanFiles := files.filter (·.fileName.endsWith ".lean") 27 | let tests := leanFiles.filter (!excludeFromTests ·) 28 | return tests.insertionSort (·.fileName < ·.fileName) 29 | 30 | inductive TestResult where 31 | | success 32 | | sorry 33 | | failure (msg : String) 34 | 35 | def fileContainsSorry (file : FS.DirEntry) : IO Bool := 36 | return "sorry".isInfixOf (← FS.readFile file.path) 37 | 38 | def runFile (file : FS.DirEntry) : IO TestResult := do 39 | if let some err ← buildResult then 40 | return .failure err 41 | else if ← fileContainsSorry file then 42 | return .sorry 43 | else 44 | return .success 45 | where 46 | buildResult : IO (Option String) := do 47 | let output ← Process.output { 48 | stdin := .null, stdout := .null, stderr := .null, 49 | cmd := "lake", args := #["build", s!"Egg.Tests.«{testName file}»"] 50 | } 51 | if output.exitCode = 0 then 52 | return none 53 | else 54 | return output.stdout 55 | 56 | def runTest (test : FS.DirEntry) (printErr : Bool) : IO Bool := do 57 | match ← runFile test with 58 | | .success => 59 | println s!"✅ {testName test}" 60 | return true 61 | | .sorry => 62 | println s!"🟡 {testName test}" 63 | return true 64 | | .failure msg => 65 | println s!"❌ {testName test}{if printErr then s!"\n{msg}" else ""}" 66 | return false 67 | 68 | def runAllTests (printErr : Bool) : IO Bool := do 69 | let mut overallSuccess := true 70 | for test in ← getTests do 71 | let success ← runTest test printErr 72 | overallSuccess := overallSuccess && success 73 | return overallSuccess 74 | 75 | def runSingleTest (name : String) : IO Bool := do 76 | runTest (printErr := true) { 77 | root := (← currentDir) / "Lean" / "Egg" / "Tests", 78 | fileName := s!"{name}.lean" 79 | } 80 | 81 | def main (args : List String) : IO Unit := do 82 | assert! args.length ≤ 1 83 | let mut success := true 84 | match args[0]? with 85 | | none => success ← runAllTests (printErr := false) 86 | | some "ci" => success ← runAllTests (printErr := true) 87 | | some test => success ← runSingleTest test 88 | unless success do Process.exit 1 89 | -------------------------------------------------------------------------------- /Lean/Egg/Core/Gen/GoalTypeSpecialization.lean: -------------------------------------------------------------------------------- 1 | import Egg.Core.Rewrite.Rule 2 | open Lean Meta 3 | 4 | namespace Egg 5 | 6 | /- 7 | For goal type specialization, we consider the following shapes of goals and rewrites: 8 | 9 | (1) . = . where . is not a proposition 10 | (2) (_ = _) ~ . where . is a proposition (necessarily) 11 | (3) . ~ (_ = _) where . is a proposition (necessarily) 12 | 13 | First note that a goal or rewrite can have shapes (2) and (3) at the same time. 14 | In that case, we need to try multiple specializations. 15 | 16 | From these 3 shapes, we get the following combinations for what should be unified with eachother: 17 | 18 | ⊢ Rw Unify 19 | 1 1 (EQ_TYPE ⊢) with (EQ_TYPE Rw) 20 | 1 2 (EQ_TYPE ⊢) with (EQ_TYPE (LHS Rw)) 21 | 1 3 (EQ_TYPE ⊢) with (EQ_TYPE (RHS Rw)) 22 | 2 1 (EQ_TYPE (LHS ⊢)) with (EQ_TYPE Rw) 23 | 2 2 (EQ_TYPE (LHS ⊢)) with (EQ_TYPE (LHS Rw)) 24 | 2 3 (EQ_TYPE (LHS ⊢)) with (EQ_TYPE (RHS Rw)) 25 | 3 1 (EQ_TYPE (RHS ⊢)) with (EQ_TYPE Rw) 26 | 3 2 (EQ_TYPE (RHS ⊢)) with (EQ_TYPE (LHS Rw)) 27 | 3 3 (EQ_TYPE (RHS ⊢)) with (EQ_TYPE (RHS Rw)) 28 | 29 | That is, the goal's unification target is independent of the rewrite's unification target and vice 30 | versa. It can be determined completely from the shape. 31 | -/ 32 | private def Congr.unificationTargets (cgr : Congr) : MetaM (List Expr) := do 33 | let cgrType ← cgr.type 34 | -- Shape 1: 35 | if cgrType.isProp then 36 | -- Shapes 2 & 3: 37 | let mut targets := [] 38 | if let some (target, _, _) := cgr.lhs.eq? then targets := target :: targets 39 | if let some (target, _, _) := cgr.rhs.eq? then targets := target :: targets 40 | return targets 41 | else 42 | return [cgrType] 43 | 44 | private def specializeForTargets (rw : Rewrite) (goalUnif rwUnif : Expr) (subgoals : Bool) : 45 | MetaM (Option Rewrite) := do 46 | let mut (spec, subst) ← rw.freshWithSubst 47 | unless ← isDefEq goalUnif (subst.apply rwUnif) do return none 48 | spec ← spec.instantiateMVars 49 | unless ← spec.isValid subgoals do return none 50 | return spec 51 | 52 | private def specialize (rule : Rewrite.Rule) (goal : Congr) (subgoals : Bool) : MetaM Rewrite.Rules := do 53 | -- Computes the unification targets. 54 | let goalTargets ← goal.unificationTargets 55 | let rwTargets ← rule.rw.unificationTargets 56 | -- (Potentially) generates a rewrite for each pair of unification targets. 57 | let mut result := ∅ 58 | let mut idx := 0 59 | for goalTarget in goalTargets do 60 | for rwTarget in rwTargets do 61 | let some spec ← specializeForTargets rule.rw goalTarget rwTarget subgoals | continue 62 | result := result.insert <| spec.with (.goalTypeSpec rule.src idx) rule.dir 63 | idx := idx + 1 64 | return result 65 | 66 | def genGoalTypeSpecializations (rules : Rewrite.Rules) (goal : Congr) (subgoals : Bool) : 67 | MetaM Rewrite.Rules := do 68 | let mut result := ∅ 69 | for rule in rules.entries do 70 | if ← rule.rw.isValid subgoals then continue 71 | let specs ← specialize rule goal subgoals 72 | result := result ∪ specs 73 | return result 74 | -------------------------------------------------------------------------------- /Lean/Egg/Tests/MVars Properties.lean: -------------------------------------------------------------------------------- 1 | import Egg 2 | 3 | set_option egg.groundEqs false 4 | set_option trace.egg.rewrites.explicit true 5 | set_option linter.unusedVariables false 6 | 7 | /-- 8 | trace: [egg.rewrites.explicit] Basic (1) 9 | [egg.rewrites.explicit] #0(⇒)❌[lhsSingleMVar]: h 10 | [egg.rewrites.explicit] ?x = ?x 11 | [egg.rewrites.explicit] LHS MVars 12 | [?x: [unconditionallyVisible]] 13 | [egg.rewrites.explicit] RHS MVars 14 | [?x: [unconditionallyVisible]] 15 | -/ 16 | #guard_msgs in 17 | example (h : ∀ x : Nat, x = x) : true = true := by 18 | egg [h] 19 | 20 | /-- 21 | trace: [egg.rewrites.explicit] Basic (2) 22 | [egg.rewrites.explicit] #0(⇒): h 23 | [egg.rewrites.explicit] f ?x = ?x 24 | [egg.rewrites.explicit] LHS MVars 25 | [?α: [unconditionallyVisible], ?x: [unconditionallyVisible]] 26 | [egg.rewrites.explicit] RHS MVars 27 | [?x: [unconditionallyVisible]] 28 | [egg.rewrites.explicit] #0(⇐)❌[rhsMVarInclusion: [?α], lhsSingleMVar]: h 29 | [egg.rewrites.explicit] ?x = f ?x 30 | [egg.rewrites.explicit] LHS MVars 31 | [?x: [unconditionallyVisible]] 32 | [egg.rewrites.explicit] RHS MVars 33 | [?α: [unconditionallyVisible], ?x: [unconditionallyVisible]] 34 | -/ 35 | #guard_msgs in 36 | example (f : {α : Type} → α → α) (h : ∀ α (x : α), f x = x) : true = true := by 37 | egg [h] 38 | 39 | /-- 40 | trace: [egg.rewrites.explicit] Basic (2) 41 | [egg.rewrites.explicit] #0(⇒): h 42 | [egg.rewrites.explicit] ?x + ?x = ?x 43 | [egg.rewrites.explicit] Conditions 44 | [egg.rewrites.explicit] Add ?α 45 | [egg.rewrites.explicit] LHS MVars 46 | [?α: [inErasedTcInst, inTcInstTerm, unconditionallyVisible], 47 | ?inst✝: [inTcInstTerm, isTcInst], 48 | ?x: [unconditionallyVisible]] 49 | [egg.rewrites.explicit] RHS MVars 50 | [?x: [unconditionallyVisible]] 51 | [egg.rewrites.explicit] #0(⇐)❌[tcMVarInclusion: [?α], rhsMVarInclusion: [?α], lhsSingleMVar]: h 52 | [egg.rewrites.explicit] ?x = ?x + ?x 53 | [egg.rewrites.explicit] Conditions 54 | [egg.rewrites.explicit] Add ?α 55 | [egg.rewrites.explicit] LHS MVars 56 | [?x: [unconditionallyVisible]] 57 | [egg.rewrites.explicit] RHS MVars 58 | [?α: [inErasedTcInst, inTcInstTerm, unconditionallyVisible], 59 | ?inst✝: [inTcInstTerm, isTcInst], 60 | ?x: [unconditionallyVisible]] 61 | -/ 62 | #guard_msgs in 63 | example (h : ∀ {α} [Add α] (x : α), x + x = x) : true = true := by 64 | egg [h] 65 | 66 | /-- 67 | trace: [egg.rewrites.explicit] Basic (2) 68 | [egg.rewrites.explicit] #0(⇒): h 69 | [egg.rewrites.explicit] f ?n ⋯ = ?n 70 | [egg.rewrites.explicit] LHS MVars 71 | [?n: [inErasedProof, inProofTerm, unconditionallyVisible]] 72 | [egg.rewrites.explicit] RHS MVars 73 | [?n: [unconditionallyVisible]] 74 | [egg.rewrites.explicit] #0(⇐)❌[lhsSingleMVar]: h 75 | [egg.rewrites.explicit] ?n = f ?n ⋯ 76 | [egg.rewrites.explicit] LHS MVars 77 | [?n: [unconditionallyVisible]] 78 | [egg.rewrites.explicit] RHS MVars 79 | [?n: [inErasedProof, inProofTerm, unconditionallyVisible]] 80 | -/ 81 | #guard_msgs in 82 | example (f : (n : Nat) → (0 < n + 1) → Nat) (h : ∀ n, f n (Nat.zero_lt_succ n) = n) : 0 = 0 := by 83 | egg [h] 84 | -------------------------------------------------------------------------------- /Rust/Slotted/src/basic.rs: -------------------------------------------------------------------------------- 1 | use slotted_egraphs::*; 2 | use crate::result::*; 3 | use crate::analysis::*; 4 | use crate::beta::*; 5 | use crate::eta::*; 6 | use crate::levels::*; 7 | use crate::nat_lit::*; 8 | use crate::rewrite::*; 9 | 10 | #[repr(C)] 11 | pub struct Config { 12 | optimize_expl: bool, 13 | time_limit: usize, 14 | node_limit: usize, 15 | iter_limit: usize, 16 | nat_lit: bool, 17 | eta: bool, 18 | eta_expand: bool, 19 | beta: bool, 20 | levels: bool, 21 | shapes: bool, 22 | union_semantics: bool, 23 | subgoals: bool 24 | } 25 | 26 | pub fn explain_congr( 27 | init: String, goal: String, rw_templates: Vec, guides: Vec, 28 | cfg: Config, _viz_path: Option 29 | ) -> Result<(String, LeanEGraph, Report), Error> { 30 | let mut egraph: LeanEGraph = EGraph::new(); 31 | 32 | let init_expr = RecExpr::parse(&init).map_err(|err| { 33 | Error::Init(format!("Failed to parse lhs of goal: {:?}\n\n {}", err, init).to_string()) 34 | })?; 35 | let goal_expr = RecExpr::parse(&goal).map_err(|err| { 36 | Error::Goal(format!("Failed to parse rhs of goal: {:?}\n\n {}", err, goal).to_string()) 37 | })?; 38 | let init_id = egraph.add_expr(init_expr.clone()); 39 | let goal_id = egraph.add_expr(goal_expr.clone()); 40 | 41 | for guide in guides { 42 | let expr = RecExpr::parse(&guide).map_err(|err| { 43 | Error::Guide(format!("Failed to parse guide term: {:?}\n\n {}", err, guide).to_string()) 44 | })?; 45 | egraph.add_expr(expr); 46 | } 47 | 48 | let mut rws; 49 | match templates_to_rewrites(rw_templates, cfg.subgoals) { 50 | Ok(r) => rws = r, 51 | Err(err) => return Err(Error::Rewrite(err.to_string())) 52 | } 53 | if cfg.nat_lit { rws.append(&mut nat_lit_rws()) } 54 | if cfg.eta { rws.push(eta_reduction_rw()) } 55 | if cfg.eta_expand { rws.push(eta_expansion_rw()) } 56 | if cfg.beta { rws.append(&mut beta_reduction_rws(/*small_step:*/false)) } 57 | if cfg.levels { rws.append(&mut level_rws()) } 58 | 59 | let i = init_id.clone(); 60 | let g = goal_id.clone(); 61 | let report = run_eqsat(&mut egraph, rws, cfg.iter_limit, cfg.time_limit, move |egraph| { 62 | if egraph.eq(&i, &g) { 63 | Err("proved goal".to_string()) 64 | } else { 65 | Ok(()) 66 | } 67 | }); 68 | 69 | if egraph.eq(&init_id, &goal_id) { 70 | let expl = egraph.explain_equivalence(init_expr, goal_expr); 71 | 72 | // HACK: For debugging purposes we're abusing the `optimize_expl` option to indicate whether 73 | // the tree explanation should be returned instead of a flat explanation. A value of 74 | // `false` indicates that the tree explanation should be returned. 75 | if !cfg.optimize_expl { 76 | let tree_expl = expl.to_string(&egraph); 77 | Ok((tree_expl, egraph, report)) 78 | } else { 79 | let flat_expl = expl.to_flat_string(&egraph); 80 | Ok((flat_expl, egraph, report)) 81 | } 82 | } else { 83 | Ok(("".to_string(), egraph, report)) 84 | } 85 | } -------------------------------------------------------------------------------- /Lean/Egg/Core/Encode/Basic.lean: -------------------------------------------------------------------------------- 1 | import Egg.Lean 2 | import Egg.Core.Normalize 3 | import Egg.Core.Encode.EncodeM 4 | import Egg.Core.Encode.Shapes 5 | import Lean 6 | open Lean 7 | 8 | namespace Egg 9 | 10 | open EncodeM 11 | 12 | private def encodeLevel : Level → EncodeM Expression 13 | | .zero => return "0" 14 | | .succ l => return s!"(succ {← encodeLevel l})" 15 | | .max l₁ l₂ => return s!"(max {← encodeLevel l₁} {← encodeLevel l₂})" 16 | | .imax l₁ l₂ => return s!"(imax {← encodeLevel l₁} {← encodeLevel l₂})" 17 | | .param name => return s!"(param \"{name}\")" 18 | | .mvar id => do 19 | if ← isLevelMVarAssignable id 20 | then return s!"?{id.uniqueIdx!}" 21 | else return s!"(uvar {id.uniqueIdx!})" 22 | 23 | -- Note: This function expects its input expression to be normalized (cf. `Egg.normalize`). 24 | partial def encode (e : Expr) (cfg : Config.Encoding) : MetaM Expression := do 25 | Prod.fst <$> (go e).run { config := cfg } 26 | where 27 | go (e : Expr) : EncodeM Expression := 28 | withCache e do 29 | let basic ← 30 | if ← needsProofErasure e then 31 | let prop ← normalize (← Meta.inferType e) .noReduce 32 | let enc ← withoutShapes do go prop 33 | pure s!"(proof {enc})" 34 | else if let some cls ← needsInstErasure? e then 35 | let cls ← normalize cls .noReduce 36 | let enc ← withoutShapes do go cls 37 | pure s!"(inst {enc})" 38 | else if let some (lhs, rhs) := e.eqOrIff? then 39 | pure s!"(= {← go lhs} {← go rhs})" 40 | else 41 | core e 42 | if (← config).shapes then 43 | let shape := shape (← Meta.inferType e) 44 | return s!"(◇ {shape} {basic})" 45 | else 46 | return basic 47 | 48 | core : Expr → EncodeM Expression 49 | | .fvar id => encodeFVar id 50 | | .mvar id => encodeMVar id 51 | | .sort lvl => return s!"(sort {← encodeLevel lvl})" 52 | | .const name lvls => return s!"(const \"{name}\"{← encodeConstLvls lvls})" 53 | | .app fn arg => return s!"(app {← go fn} {← go arg})" 54 | | .lam _ ty b _ => encodeLambda ty b 55 | | .forallE _ ty b _ => encodeForall ty b 56 | | .lit (.strVal l) => return s!"(lit {Json.renderString l})" 57 | | .lit (.natVal l) => return s!"(lit {l})" 58 | | .bvar _ => panic! "'Egg.encode.core' found loose bvar" 59 | | _ => panic! "'Egg.encode.core' received non-normalized expression" 60 | 61 | encodeFVar (id : FVarId) : EncodeM Expression := do 62 | if let some bvarName ← bvarName? id then 63 | return s!"(bvar {bvarName})" 64 | else return s!"(fvar {id.uniqueIdx!})" 65 | 66 | encodeMVar (id : MVarId) : EncodeM Expression := do 67 | if ← id.isAssignable 68 | then return s!"?{id.uniqueIdx!}" 69 | else return s!"(mvar {id.uniqueIdx!})" 70 | 71 | encodeConstLvls (lvls : List Level) : EncodeM Expression := 72 | lvls.foldlM (init := "") (return s!"{·} {← encodeLevel ·}") 73 | 74 | encodeLambda (ty b : Expr) : EncodeM Expression := do 75 | -- It's critical that we encode `ty` outside of the `withInstantiatedBVar` block, as otherwise 76 | -- the bvars in `encTy` are incorrectly shifted by 1. 77 | let encTy ← go ty 78 | withInstantiatedBVar ty b fun var? body => return s!"(λ {var?}{encTy} {← go body})" 79 | 80 | encodeForall (ty b : Expr) : EncodeM Expression := do 81 | -- It's critical that we encode `ty` outside of the `withInstantiatedBVar` block, as otherwise 82 | -- the bvars in `encTy` are incorrectly shifted by 1. 83 | let encTy ← go ty 84 | withInstantiatedBVar ty b fun var? body => return s!"(∀ {var?}{encTy} {← go body})" 85 | -------------------------------------------------------------------------------- /Lean/Egg/Tests/FunctionalTc.lean: -------------------------------------------------------------------------------- 1 | import Egg 2 | 3 | -- Tests involving auto-generated rewrite rules for reducing type class projections. 4 | 5 | namespace TcProj 6 | 7 | inductive List (α) where 8 | | nil : List α 9 | | cons : α → List α → List α 10 | 11 | notation "[]" => List.nil 12 | infixr:50 " :: " => List.cons 13 | 14 | def append : List α → List α → List α 15 | | [], bs => bs 16 | | a :: as, bs => a :: (append as bs) 17 | 18 | instance {α : Type _} : Append (List α) where 19 | append := append 20 | 21 | theorem append_nil (as : List α) : as ++ [] = as := by 22 | induction as with 23 | | nil => egg [append] 24 | | cons _ _ ih => egg [ih, append] 25 | 26 | theorem append_assoc (as bs cs : List α) : (as ++ bs) ++ cs = as ++ (bs ++ cs) := by 27 | induction as with 28 | | nil => egg [append] 29 | | cons _ _ ih => egg [ih, append] 30 | 31 | def reverseAux : List α → List α → List α 32 | | [], r => r 33 | | a :: l, r => reverseAux l (a :: r) 34 | 35 | def reverse (as : List α) : List α := 36 | reverseAux as [] 37 | 38 | theorem reverseAux_eq_append (as bs : List α) : reverseAux as bs = (reverseAux as []) ++ bs := by 39 | induction as generalizing bs with 40 | | nil => egg [reverseAux, append] 41 | | cons _ _ ih => egg [reverseAux, append_assoc, ih, append] 42 | 43 | theorem reverse_nil : reverse ([] : List α) = [] := by 44 | egg [reverse, reverseAux] 45 | 46 | theorem reverse_cons (a : α) (as : List α) : reverse (a :: as) = (reverse as) ++ (a :: []) := by 47 | egg [reverse, reverseAux, reverseAux_eq_append] 48 | 49 | theorem reverse_append (as bs : List α) : reverse (as ++ bs) = (reverse bs) ++ (reverse as) := by 50 | induction as generalizing bs with 51 | | nil => egg [reverse_nil, append_nil, append] 52 | | cons a as ih => egg [ih, append_assoc, reverse_cons, append] 53 | 54 | -- EGRAPHS Example: 55 | 56 | example (as bs : List α) : reverse (as ++ bs) = (reverse bs) ++ (reverse as) := by 57 | induction as generalizing bs with 58 | | nil => 59 | calc reverse ([] ++ bs) 60 | _ = reverse bs := by simp only [(· ++ ·), Append.append, append] 61 | _ = reverse bs ++ [] := by rw [append_nil] 62 | _ = reverse bs ++ reverse [] := by rw [reverse_nil] 63 | | cons a as ih => 64 | calc reverse ((a :: as) ++ bs) 65 | _ = reverse (a :: (as ++ bs)) := by simp only [(· ++ ·), Append.append, append] 66 | _ = reverse (as ++ bs) ++ (a :: []) := by rw [reverse_cons] 67 | _ = reverse bs ++ reverse as ++ (a :: []) := by rw [ih] 68 | _ = reverse bs ++ (reverse as ++ (a :: [])) := by rw [append_assoc] 69 | _ = reverse bs ++ reverse (a :: as) := by rw [reverse_cons] 70 | 71 | example (as bs : List α) : reverse (as ++ bs) = (reverse bs) ++ (reverse as) := by 72 | induction as generalizing bs with 73 | | nil => 74 | calc reverse ([] ++ bs) 75 | _ = reverse bs := by simp only [(· ++ ·), Append.append, append] 76 | _ = reverse bs ++ [] := by rw [append_nil] 77 | _ = reverse bs ++ reverse [] := by rw [reverse_nil] 78 | | cons a as ih => 79 | calc reverse ((a :: as) ++ bs) 80 | _ = reverse (as ++ bs) ++ (a :: []) := by simp only [(· ++ ·), Append.append, append, reverse_cons] 81 | _ = reverse bs ++ reverse (a :: as) := by rw [ih, append_assoc, reverse_cons] 82 | 83 | syntax "lists" egg_args : tactic 84 | macro_rules 85 | | `(tactic| lists $[[ $rws?,* ]]?) => open List in 86 | `(tactic| egg [reverse_nil, reverse_cons, append, append_nil, append_assoc, $(rws?.getD ⟨#[]⟩),*]) 87 | 88 | example (as bs : List α) : reverse (as ++ bs) = (reverse bs) ++ (reverse as) := by 89 | induction as generalizing bs with 90 | | nil => lists 91 | | cons _ _ ih => lists [ih] 92 | 93 | example (as bs : List α) : reverse (as ++ bs) = (reverse bs) ++ (reverse as) := by 94 | induction as generalizing bs <;> lists [*] 95 | -------------------------------------------------------------------------------- /Lean/Egg/Core/Gen/StructProjs.lean: -------------------------------------------------------------------------------- 1 | import Egg.Core.Rewrite.Rule 2 | import Egg.Core.Guides 3 | 4 | open Lean Std Meta Elab Term 5 | 6 | namespace Egg 7 | 8 | private structure StructInfo where 9 | projs : HashSet Nat := ∅ 10 | ctor : Option ConstructorVal := none 11 | 12 | private abbrev StructInfos := HashMap Name StructInfo 13 | 14 | private def StructInfos.ofExpr (e : Expr) (init : StructInfos := ∅) : MetaM StructInfos := do 15 | match e with 16 | | .const c _ => ofConst c init 17 | | .app e₁ e₂ | .lam _ e₁ e₂ _ | .forallE _ e₁ e₂ _ => ofExpr e₂ (← ofExpr e₁ init) 18 | | .mdata .. | .letE .. | .proj .. => throwError "egg: internal error: 'Egg.StructInfo.ofExpr' received non-normalized expression" 19 | | _ => return init 20 | where 21 | ofConst (const : Name) (infos : StructInfos) : MetaM StructInfos := do 22 | let env ← getEnv 23 | if let some projInfo ← getProjectionFnInfo? const then 24 | if let some (.ctorInfo { induct, .. }) := env.find? projInfo.ctorName then 25 | return infos.alterD induct {} fun info => 26 | { info with projs := info.projs.insert projInfo.i } 27 | else if let some (.ctorInfo ctorInfo) := env.find? const then 28 | if isStructure env ctorInfo.induct then 29 | return infos.alterD ctorInfo.induct {} fun info => 30 | { info with ctor := some ctorInfo } 31 | return infos 32 | 33 | private def StructInfos.rules (infos : StructInfos) (cfg : Config.Normalization) : MetaM Rewrite.Rules := do 34 | let mut rules := ∅ 35 | for (structName, info) in infos do 36 | let some ctor := info.ctor | continue 37 | for projIdx in info.projs do 38 | let ls ← mkFreshLevelMVarsFor (.ctorInfo ctor) 39 | let (mvars, _) ← forallMetaTelescope (ctor.type.instantiateLevelParams ctor.levelParams ls) 40 | let appCtor := mkAppN (.const ctor.name ls) mvars 41 | let some field := (getStructureFields (← getEnv) structName)[projIdx]? 42 | | throwError "Internal error in 'Egg.StructInfos.rws'" 43 | let lhs ← mkProjection appCtor field 44 | let rhs := mvars[ctor.numParams + projIdx]! 45 | let eq ← mkForallFVars mvars (← mkEq lhs rhs) 46 | let proof ← mkLambdaFVars mvars (← mkEqRefl lhs) 47 | let some rules' ← rules.add? (.structProj rules.rws.size) proof eq cfg (.dir .forward) 48 | | throwError "Internal error in 'Egg.StructInfos.rws'" 49 | rules := rules' 50 | return rules 51 | 52 | private def Congr.structInfos (cgr : Congr) (init : StructInfos := ∅) : MetaM StructInfos := do 53 | StructInfos.ofExpr cgr.rhs (init := ← StructInfos.ofExpr cgr.lhs init) 54 | 55 | private def Rewrite.Condition.structInfos (cond : Rewrite.Condition) (init : StructInfos := ∅) : 56 | MetaM StructInfos := 57 | StructInfos.ofExpr cond.type init 58 | 59 | private def Rewrite.structInfos (rw : Rewrite) (init : StructInfos := ∅) : MetaM StructInfos := do 60 | let mut infos ← rw.toCongr.structInfos init 61 | for cond in rw.conds.active do infos ← cond.structInfos (init := infos) 62 | return infos 63 | 64 | private def Rewrite.Rules.structInfos (rules : Rewrite.Rules) (init : StructInfos := ∅) : MetaM StructInfos := do 65 | let mut infos := init 66 | for rule in rules.entries do infos ← rule.rw.structInfos (init := infos) 67 | return infos 68 | 69 | private def Guide.structInfos (guide : Guide) (init : StructInfos := ∅) : MetaM StructInfos := do 70 | StructInfos.ofExpr guide.expr init 71 | 72 | private def Guides.structInfos (guides : Guides) (init : StructInfos := ∅) : MetaM StructInfos := do 73 | let mut infos := init 74 | for guide in guides do infos ← guide.structInfos (init := infos) 75 | return infos 76 | 77 | def genStructProjRws 78 | (goal : Congr) (rules : Rewrite.Rules) (guides : Guides) (cfg : Config.Normalization) : 79 | MetaM Rewrite.Rules := do 80 | (← goal.structInfos <| ← rules.structInfos <| ← guides.structInfos).rules cfg 81 | --------------------------------------------------------------------------------