From 01a23b338a66d255126cd3cfb85f429b136b7648 Mon Sep 17 00:00:00 2001 From: Leni Aniva Date: Thu, 28 Mar 2024 18:56:42 -0700 Subject: [PATCH 1/5] feat: Unfold aux lemmas when printing root expr --- Pantograph/Library.lean | 12 ++++++++++-- Test/Common.lean | 3 +++ Test/Environment.lean | 3 ++- Test/Main.lean | 4 ++-- Test/{Holes.lean => Metavar.lean} | 6 +++--- Test/Serial.lean | 3 ++- 6 files changed, 22 insertions(+), 9 deletions(-) rename Test/{Holes.lean => Metavar.lean} (98%) diff --git a/Pantograph/Library.lean b/Pantograph/Library.lean index 078ca0f..c40322a 100644 --- a/Pantograph/Library.lean +++ b/Pantograph/Library.lean @@ -165,13 +165,21 @@ def goalResume (target: GoalState) (goals: Array String): Except String GoalStat def goalSerialize (state: GoalState) (options: @&Protocol.Options): Lean.CoreM (Array Protocol.Goal) := runMetaM <| state.serializeGoals (parent := .none) options +def Lean.Name.isAuxLemma (n : Lean.Name) : Bool := n matches .num (.str _ "_auxLemma") _ + +/-- Unfold all lemmas created by `Lean.Meta.mkAuxLemma`. These end in `_auxLemma.nn` where `nn` is a number. -/ +def unfoldAuxLemmas (e : Lean.Expr) : Lean.MetaM Lean.Expr := do + Lean.Meta.deltaExpand e Lean.Name.isAuxLemma + @[export pantograph_goal_print_m] def goalPrint (state: GoalState) (options: @&Protocol.Options): Lean.CoreM Protocol.GoalPrintResult := do let metaM := do state.restoreMetaM return { - root? := ← state.rootExpr?.mapM (λ expr => serialize_expression options expr), - parent? := ← state.parentExpr?.mapM (λ expr => serialize_expression options expr), + root? := ← state.rootExpr?.mapM (λ expr => do + serialize_expression options (← unfoldAuxLemmas expr)), + parent? := ← state.parentExpr?.mapM (λ expr => do + serialize_expression options (← unfoldAuxLemmas expr)), } runMetaM metaM diff --git a/Test/Common.lean b/Test/Common.lean index 7c5b6e5..2e7149d 100644 --- a/Test/Common.lean +++ b/Test/Common.lean @@ -1,8 +1,11 @@ import Pantograph.Goal import Pantograph.Library import Pantograph.Protocol +import Lean import LSpec +open Lean + namespace Pantograph namespace Protocol diff --git a/Test/Environment.lean b/Test/Environment.lean index 9e3bd70..ba99380 100644 --- a/Test/Environment.lean +++ b/Test/Environment.lean @@ -2,11 +2,12 @@ import LSpec import Pantograph.Serial import Pantograph.Environment import Test.Common +import Lean +open Lean namespace Pantograph.Test.Environment open Pantograph -open Lean deriving instance DecidableEq, Repr for Protocol.InductInfo deriving instance DecidableEq, Repr for Protocol.ConstructorInfo diff --git a/Test/Main.lean b/Test/Main.lean index 4a8ab1f..d24f45f 100644 --- a/Test/Main.lean +++ b/Test/Main.lean @@ -1,6 +1,6 @@ import LSpec import Test.Environment -import Test.Holes +import Test.Metavar import Test.Integration import Test.Proofs import Test.Serial @@ -11,7 +11,7 @@ def main := do Lean.initSearchPath (← Lean.findSysroot) let suites := [ - Holes.suite, + Metavar.suite, Integration.suite, Proofs.suite, Serial.suite, diff --git a/Test/Holes.lean b/Test/Metavar.lean similarity index 98% rename from Test/Holes.lean rename to Test/Metavar.lean index af322e9..cbbfc81 100644 --- a/Test/Holes.lean +++ b/Test/Metavar.lean @@ -3,7 +3,7 @@ import Pantograph.Goal import Pantograph.Serial import Test.Common -namespace Pantograph.Test.Holes +namespace Pantograph.Test.Metavar open Pantograph open Lean @@ -204,6 +204,6 @@ def suite: IO LSpec.TestSeq := do let tests ← proofRunner env tests return acc ++ (LSpec.group name tests)) LSpec.TestSeq.done - return LSpec.group "Holes" tests + return LSpec.group "Metavar" tests -end Pantograph.Test.Holes +end Pantograph.Test.Metavar diff --git a/Test/Serial.lean b/Test/Serial.lean index 490b538..c2810c8 100644 --- a/Test/Serial.lean +++ b/Test/Serial.lean @@ -1,11 +1,12 @@ import LSpec import Pantograph.Serial import Test.Common +import Lean +open Lean namespace Pantograph.Test.Serial open Pantograph -open Lean deriving instance Repr, DecidableEq for Protocol.BoundExpression From d9af064888f5068f61bc8f6b4c7864ed506e11e2 Mon Sep 17 00:00:00 2001 From: Leni Aniva Date: Thu, 28 Mar 2024 19:27:45 -0700 Subject: [PATCH 2/5] test: Elimination of aux lemmas --- Test/Metavar.lean | 58 +++++++++++++++++++++++++++++++++++++++++++++++ 1 file changed, 58 insertions(+) diff --git a/Test/Metavar.lean b/Test/Metavar.lean index cbbfc81..734c1d9 100644 --- a/Test/Metavar.lean +++ b/Test/Metavar.lean @@ -2,6 +2,7 @@ import LSpec import Pantograph.Goal import Pantograph.Serial import Test.Common +import Lean namespace Pantograph.Test.Metavar open Pantograph @@ -85,6 +86,62 @@ def test_m_couple: TestM Unit := do addTest $ LSpec.check "exact 3" ((← state1b.serializeGoals (options := ← read)).map (·.target.pp?) = #[.some "2 ≤ 3", .some "3 ≤ 5"]) addTest $ LSpec.test "(2 root)" state1b.rootExpr?.isNone + +def test_m_couple_simp: TestM Unit := do + let state? ← startProof "(2: Nat) ≤ 5" + let state0 ← match state? with + | .some state => pure state + | .none => do + addTest $ assertUnreachable "Goal could not parse" + return () + + let state1 ← match ← state0.execute (goalId := 0) (tactic := "apply Nat.le_trans") with + | .success state => pure state + | other => do + addTest $ assertUnreachable $ other.toString + return () + addTest $ LSpec.check "apply Nat.le_trans" ((← state1.serializeGoals (options := ← read)).map (·.target.pp?) = + #[.some "2 ≤ ?m", .some "?m ≤ 5", .some "Nat"]) + let state2 ← match ← state1.execute (goalId := 2) (tactic := "exact 2") with + | .success state => pure state + | other => do + addTest $ assertUnreachable $ other.toString + return () + addTest $ LSpec.test "(1b root)" state2.rootExpr?.isNone + let state1b ← match state2.continue state1 with + | .error msg => do + addTest $ assertUnreachable $ msg + return () + | .ok state => pure state + addTest $ LSpec.check "exact 2" ((← state1b.serializeGoals (options := ← read)).map (·.target.pp?) = + #[.some "2 ≤ 2", .some "2 ≤ 5"]) + addTest $ LSpec.test "(2 root)" state1b.rootExpr?.isNone + let state3 ← match ← state1b.execute (goalId := 0) (tactic := "simp") with + | .success state => pure state + | other => do + addTest $ assertUnreachable $ other.toString + return () + let state4 ← match state3.continue state1b with + | .error msg => do + addTest $ assertUnreachable $ msg + return () + | .ok state => pure state + let state5 ← match ← state4.execute (goalId := 0) (tactic := "simp") with + | .success state => pure state + | other => do + addTest $ assertUnreachable $ other.toString + return () + + state5.restoreMetaM + let root ← match state5.rootExpr? with + | .some e => pure e + | .none => + addTest $ assertUnreachable "(5 root)" + return () + let rootStr: String := toString (← Lean.Meta.ppExpr root) + addTest $ LSpec.check "(5 root)" (rootStr = "Nat.le_trans (of_eq_true (Init.Data.Nat.Basic._auxLemma.4 2)) (of_eq_true (eq_true_of_decide (Eq.refl true)))") + let rootStr: String := toString (← Lean.Meta.ppExpr (← unfoldAuxLemmas root)) + addTest $ LSpec.check "(5 root)" (rootStr = "Nat.le_trans (of_eq_true (eq_true (Nat.le_refl 2))) (of_eq_true (eq_true_of_decide (Eq.refl true)))") return () def test_proposition_generation: TestM Unit := do @@ -196,6 +253,7 @@ def suite: IO LSpec.TestSeq := do (trustLevel := 1) let tests := [ ("2 < 5", test_m_couple), + ("2 < 5", test_m_couple_simp), ("Proposition Generation", test_proposition_generation), ("Partial Continuation", test_partial_continuation) ] From 35c4ea693d22ca5695fb840a7fa824cde6c7f57b Mon Sep 17 00:00:00 2001 From: Leni Aniva Date: Thu, 28 Mar 2024 19:49:44 -0700 Subject: [PATCH 3/5] feat: Stop cataloging internal/detail dependencies --- Pantograph/Environment.lean | 25 ++++++++++++++++--------- Pantograph/Library.lean | 6 ------ Pantograph/Serial.lean | 17 ++++++++++------- Test/Environment.lean | 11 +++++------ 4 files changed, 31 insertions(+), 28 deletions(-) diff --git a/Pantograph/Environment.lean b/Pantograph/Environment.lean index ecae517..a85342a 100644 --- a/Pantograph/Environment.lean +++ b/Pantograph/Environment.lean @@ -7,14 +7,15 @@ open Pantograph namespace Pantograph.Environment -def is_symbol_unsafe_or_internal (n: Lean.Name) (info: Lean.ConstantInfo): Bool := - isLeanSymbol n ∨ (Lean.privateToUserName? n |>.map isLeanSymbol |>.getD false) ∨ info.isUnsafe +def isNameInternal (n: Lean.Name): Bool := + -- Returns true if the name is an implementation detail which should not be shown to the user. + isLeanSymbol n ∨ (Lean.privateToUserName? n |>.map isLeanSymbol |>.getD false) ∨ n.isAuxLemma where isLeanSymbol (name: Lean.Name): Bool := match name.getRoot with | .str _ name => name == "Lean" | _ => true -def to_compact_symbol_name (n: Lean.Name) (info: Lean.ConstantInfo): String := +def toCompactSymbolName (n: Lean.Name) (info: Lean.ConstantInfo): String := let pref := match info with | .axiomInfo _ => "a" | .defnInfo _ => "d" @@ -26,14 +27,14 @@ def to_compact_symbol_name (n: Lean.Name) (info: Lean.ConstantInfo): String := | .recInfo _ => "r" s!"{pref}{toString n}" -def to_filtered_symbol (n: Lean.Name) (info: Lean.ConstantInfo): Option String := - if is_symbol_unsafe_or_internal n info +def toFilteredSymbol (n: Lean.Name) (info: Lean.ConstantInfo): Option String := + if isNameInternal n || info.isUnsafe then Option.none - else Option.some <| to_compact_symbol_name n info + else Option.some <| toCompactSymbolName n info def catalog (_: Protocol.EnvCatalog): CoreM Protocol.EnvCatalogResult := do let env ← Lean.MonadEnv.getEnv let names := env.constants.fold (init := #[]) (λ acc name info => - match to_filtered_symbol name info with + match toFilteredSymbol name info with | .some x => acc.push x | .none => acc) return { symbols := names } @@ -58,8 +59,14 @@ def inspect (args: Protocol.EnvInspect) (options: @&Protocol.Options): CoreM (Pr value? := ← value?.mapM (λ v => serialize_expression options v |>.run'), publicName? := Lean.privateToUserName? name |>.map (·.toString), -- BUG: Warning: getUsedConstants here will not include projections. This is a known bug. - typeDependency? := if args.dependency?.getD false then .some <| info.type.getUsedConstants.map (λ n => name_to_ast n) else .none, - valueDependency? := if args.dependency?.getD false then info.value?.map (·.getUsedConstants.map (λ n => name_to_ast n)) else .none, + typeDependency? := if args.dependency?.getD false + then .some <| info.type.getUsedConstants.map (λ n => name_to_ast n) + else .none, + valueDependency? := ← if args.dependency?.getD false + then info.value?.mapM (λ e => do + let e ← (unfoldAuxLemmas e).run' + pure $ e.getUsedConstants.filter (!isNameInternal ·) |>.map (λ n => name_to_ast n) ) + else pure (.none), module? := module? } let result := match info with diff --git a/Pantograph/Library.lean b/Pantograph/Library.lean index c40322a..7fda4d9 100644 --- a/Pantograph/Library.lean +++ b/Pantograph/Library.lean @@ -165,12 +165,6 @@ def goalResume (target: GoalState) (goals: Array String): Except String GoalStat def goalSerialize (state: GoalState) (options: @&Protocol.Options): Lean.CoreM (Array Protocol.Goal) := runMetaM <| state.serializeGoals (parent := .none) options -def Lean.Name.isAuxLemma (n : Lean.Name) : Bool := n matches .num (.str _ "_auxLemma") _ - -/-- Unfold all lemmas created by `Lean.Meta.mkAuxLemma`. These end in `_auxLemma.nn` where `nn` is a number. -/ -def unfoldAuxLemmas (e : Lean.Expr) : Lean.MetaM Lean.Expr := do - Lean.Meta.deltaExpand e Lean.Name.isAuxLemma - @[export pantograph_goal_print_m] def goalPrint (state: GoalState) (options: @&Protocol.Options): Lean.CoreM Protocol.GoalPrintResult := do let metaM := do diff --git a/Pantograph/Serial.lean b/Pantograph/Serial.lean index 213ae6d..bf79314 100644 --- a/Pantograph/Serial.lean +++ b/Pantograph/Serial.lean @@ -6,17 +6,20 @@ import Lean import Pantograph.Protocol import Pantograph.Goal -namespace Pantograph open Lean +-- Symbol processing functions -- + +def Lean.Name.isAuxLemma (n : Lean.Name) : Bool := n matches .num (.str _ "_auxLemma") _ + +namespace Pantograph + +/-- Unfold all lemmas created by `Lean.Meta.mkAuxLemma`. These end in `_auxLemma.nn` where `nn` is a number. -/ +def unfoldAuxLemmas (e : Lean.Expr) : Lean.MetaM Lean.Expr := do + Lean.Meta.deltaExpand e Lean.Name.isAuxLemma + --- Input Functions --- - -/-- Read a theorem from the environment -/ -def expr_from_const (env: Environment) (name: Name): Except String Lean.Expr := - match env.find? name with - | none => throw s!"Symbol not found: {name}" - | some cInfo => return cInfo.type /-- Read syntax object from string -/ def syntax_from_str (env: Environment) (s: String): Except String Syntax := Parser.runParserCategory diff --git a/Test/Environment.lean b/Test/Environment.lean index ba99380..7014584 100644 --- a/Test/Environment.lean +++ b/Test/Environment.lean @@ -14,14 +14,14 @@ deriving instance DecidableEq, Repr for Protocol.ConstructorInfo deriving instance DecidableEq, Repr for Protocol.RecursorInfo deriving instance DecidableEq, Repr for Protocol.EnvInspectResult -def test_symbol_visibility (env: Environment): IO LSpec.TestSeq := do +def test_symbol_visibility: IO LSpec.TestSeq := do let entries: List (Name × Bool) := [ ("Nat.add_comm".toName, false), - ("Lean.Name".toName, true) + ("Lean.Name".toName, true), + ("Init.Data.Nat.Basic._auxLemma.4".toName, true), ] let suite := entries.foldl (λ suites (symbol, target) => - let constant := env.constants.find! symbol - let test := LSpec.check symbol.toString ((Environment.is_symbol_unsafe_or_internal symbol constant) == target) + let test := LSpec.check symbol.toString ((Environment.isNameInternal symbol) == target) LSpec.TestSeq.append suites test) LSpec.TestSeq.done return suite @@ -78,12 +78,11 @@ def test_inspect (env: Environment): IO LSpec.TestSeq := do def suite: IO LSpec.TestSeq := do let env: Environment ← importModules (imports := #[`Init]) - --(imports := #["Prelude"].map (λ str => { module := str.toName, runtimeOnly := false })) (opts := {}) (trustLevel := 1) return LSpec.group "Environment" $ - (LSpec.group "Symbol visibility" (← test_symbol_visibility env)) ++ + (LSpec.group "Symbol visibility" (← test_symbol_visibility)) ++ (LSpec.group "Inspect" (← test_inspect env)) end Pantograph.Test.Environment From 78a3b240ba6de1826d879801ceed5a4d61df4940 Mon Sep 17 00:00:00 2001 From: Leni Aniva Date: Thu, 28 Mar 2024 19:56:03 -0700 Subject: [PATCH 4/5] test: Catalog has no numeric symbols --- Pantograph/Environment.lean | 2 +- Test/Environment.lean | 29 ++++++++++++++++++++++------- 2 files changed, 23 insertions(+), 8 deletions(-) diff --git a/Pantograph/Environment.lean b/Pantograph/Environment.lean index a85342a..19c3249 100644 --- a/Pantograph/Environment.lean +++ b/Pantograph/Environment.lean @@ -9,7 +9,7 @@ namespace Pantograph.Environment def isNameInternal (n: Lean.Name): Bool := -- Returns true if the name is an implementation detail which should not be shown to the user. - isLeanSymbol n ∨ (Lean.privateToUserName? n |>.map isLeanSymbol |>.getD false) ∨ n.isAuxLemma + isLeanSymbol n ∨ (Lean.privateToUserName? n |>.map isLeanSymbol |>.getD false) ∨ n.isAuxLemma ∨ n.hasMacroScopes where isLeanSymbol (name: Lean.Name): Bool := match name.getRoot with | .str _ name => name == "Lean" diff --git a/Test/Environment.lean b/Test/Environment.lean index 7014584..7a398da 100644 --- a/Test/Environment.lean +++ b/Test/Environment.lean @@ -14,6 +14,21 @@ deriving instance DecidableEq, Repr for Protocol.ConstructorInfo deriving instance DecidableEq, Repr for Protocol.RecursorInfo deriving instance DecidableEq, Repr for Protocol.EnvInspectResult +def test_catalog: IO LSpec.TestSeq := do + let env: Environment ← importModules + (imports := #[`Init]) + (opts := {}) + (trustLevel := 1) + let inner: CoreM LSpec.TestSeq := do + let catalogResult ← Environment.catalog {} + let symbolsWithNum := env.constants.fold (init := #[]) (λ acc name info => + match (Environment.toFilteredSymbol name info).isSome && (name matches .num _ _) with + | false => acc + | true => acc.push name + ) + return LSpec.check "No num symbols" (symbolsWithNum.size == 0) + runCoreMSeq env inner + def test_symbol_visibility: IO LSpec.TestSeq := do let entries: List (Name × Bool) := [ ("Nat.add_comm".toName, false), @@ -30,7 +45,11 @@ inductive ConstantCat where | ctor (info: Protocol.ConstructorInfo) | recursor (info: Protocol.RecursorInfo) -def test_inspect (env: Environment): IO LSpec.TestSeq := do +def test_inspect: IO LSpec.TestSeq := do + let env: Environment ← importModules + (imports := #[`Init]) + (opts := {}) + (trustLevel := 1) let testCases: List (String × ConstantCat) := [ ("Or", ConstantCat.induct { numParams := 2, @@ -76,13 +95,9 @@ def test_inspect (env: Environment): IO LSpec.TestSeq := do runCoreMSeq env inner def suite: IO LSpec.TestSeq := do - let env: Environment ← importModules - (imports := #[`Init]) - (opts := {}) - (trustLevel := 1) - return LSpec.group "Environment" $ + (LSpec.group "Catalog" (← test_catalog)) ++ (LSpec.group "Symbol visibility" (← test_symbol_visibility)) ++ - (LSpec.group "Inspect" (← test_inspect env)) + (LSpec.group "Inspect" (← test_inspect)) end Pantograph.Test.Environment From d853cb8cc2ad3f523d9e1de9e24ddad2ae83ba7e Mon Sep 17 00:00:00 2001 From: Leni Aniva Date: Thu, 28 Mar 2024 22:08:22 -0700 Subject: [PATCH 5/5] chore: Version bump --- Pantograph/Version.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/Pantograph/Version.lean b/Pantograph/Version.lean index e23d886..688fc60 100644 --- a/Pantograph/Version.lean +++ b/Pantograph/Version.lean @@ -1,6 +1,6 @@ namespace Pantograph @[export pantograph_version] -def version := "0.2.13" +def version := "0.2.14" end Pantograph