├── lean-toolchain ├── .gitignore ├── Test ├── Tactic.lean ├── Frontend.lean ├── Parser.lean ├── Tactic │ ├── Assign.lean │ ├── Prograde.lean │ └── Fragment.lean ├── Library.lean ├── Main.lean ├── Frontend │ ├── Refactor.lean │ └── Collect.lean ├── Environment.lean ├── Serial.lean ├── Common.lean └── Delate.lean ├── Pantograph ├── Tactic.lean ├── Version.lean ├── Frontend.lean ├── Elab.lean ├── Tactic │ ├── Assign.lean │ ├── Prograde.lean │ └── Fragment.lean ├── Serial.lean ├── Frontend │ ├── MetaTranslate.lean │ ├── Basic.lean │ ├── InfoTree.lean │ ├── Distil.lean │ └── Refactor.lean ├── Library.lean └── Environment.lean ├── Pantograph.lean ├── lake-manifest.json ├── lakefile.lean ├── .pre-commit-config.yaml ├── doc ├── contributing.md ├── rationale.md ├── icon.svg └── repl.md ├── README.md ├── Tomograph.lean ├── Main.lean ├── flake.lock ├── flake.nix └── LICENSE /lean-toolchain: -------------------------------------------------------------------------------- 1 | leanprover/lean4:v4.26.0 2 | -------------------------------------------------------------------------------- /.gitignore: -------------------------------------------------------------------------------- 1 | .* 2 | !.gitignore 3 | !.pre-commit-config.yaml 4 | 5 | *.[io]lean 6 | /result 7 | -------------------------------------------------------------------------------- /Test/Tactic.lean: -------------------------------------------------------------------------------- 1 | import Test.Tactic.Assign 2 | import Test.Tactic.Fragment 3 | import Test.Tactic.Prograde 4 | -------------------------------------------------------------------------------- /Test/Frontend.lean: -------------------------------------------------------------------------------- 1 | import Test.Frontend.Collect 2 | import Test.Frontend.Distil 3 | import Test.Frontend.Refactor 4 | -------------------------------------------------------------------------------- /Pantograph/Tactic.lean: -------------------------------------------------------------------------------- 1 | import Pantograph.Tactic.Assign 2 | import Pantograph.Tactic.Fragment 3 | import Pantograph.Tactic.Prograde 4 | -------------------------------------------------------------------------------- /Pantograph/Version.lean: -------------------------------------------------------------------------------- 1 | namespace Pantograph 2 | 3 | @[export pantograph_version] 4 | def version := "0.3.12" 5 | 6 | end Pantograph 7 | -------------------------------------------------------------------------------- /Pantograph/Frontend.lean: -------------------------------------------------------------------------------- 1 | import Pantograph.Frontend.Basic 2 | import Pantograph.Frontend.Distil 3 | import Pantograph.Frontend.InfoTree 4 | import Pantograph.Frontend.MetaTranslate 5 | import Pantograph.Frontend.Refactor 6 | -------------------------------------------------------------------------------- /Pantograph.lean: -------------------------------------------------------------------------------- 1 | import Pantograph.Delate 2 | import Pantograph.Elab 3 | import Pantograph.Environment 4 | import Pantograph.Frontend 5 | import Pantograph.Goal 6 | import Pantograph.Library 7 | import Pantograph.Protocol 8 | import Pantograph.Serial 9 | import Pantograph.Version 10 | -------------------------------------------------------------------------------- /lake-manifest.json: -------------------------------------------------------------------------------- 1 | {"version": "1.1.0", 2 | "packagesDir": ".lake/packages", 3 | "packages": 4 | [{"url": "https://github.com/argumentcomputer/LSpec.git", 5 | "type": "git", 6 | "subDir": null, 7 | "scope": "", 8 | "rev": "db76512cd5266f0c576d561d8c69e2dc4890bea5", 9 | "name": "LSpec", 10 | "manifestFile": "lake-manifest.json", 11 | "inputRev": "db76512cd5266f0c576d561d8c69e2dc4890bea5", 12 | "inherited": false, 13 | "configFile": "lakefile.toml"}], 14 | "name": "pantograph", 15 | "lakeDir": ".lake"} 16 | -------------------------------------------------------------------------------- /Test/Parser.lean: -------------------------------------------------------------------------------- 1 | import LSpec 2 | import Pantograph.Elab 3 | import Test.Common 4 | 5 | open Lean 6 | namespace Pantograph.Test.Parser 7 | 8 | abbrev TestM := TestT $ Elab.TermElabM 9 | 10 | def test_runParserCategory : TestM Unit := do 11 | let .ok (stx, pos) := runParserCategory' (← getEnv) `tactic "intro n\ncases n" 12 | | fail "Tactic failed to parse" 13 | checkEq "stx" (toString stx.prettyPrint) "intro n" 14 | checkEq "pos" pos.byteIdx 8 15 | 16 | def suite (env : Environment) : List (String × IO LSpec.TestSeq) := 17 | [ 18 | ("runParserCategory", test_runParserCategory), 19 | ] |>.map (λ (name, t) => (name, runTestTermElabM env t)) 20 | -------------------------------------------------------------------------------- /lakefile.lean: -------------------------------------------------------------------------------- 1 | import Lake 2 | open Lake DSL 3 | 4 | package pantograph 5 | 6 | lean_lib Pantograph { 7 | roots := #[`Pantograph] 8 | defaultFacets := #[LeanLib.sharedFacet] 9 | } 10 | 11 | lean_lib Repl { 12 | } 13 | 14 | @[default_target] 15 | lean_exe repl { 16 | root := `Main 17 | -- Solves the native symbol not found problem 18 | supportInterpreter := true 19 | } 20 | 21 | lean_exe tomograph { 22 | root := `Tomograph 23 | -- Solves the native symbol not found problem 24 | supportInterpreter := true 25 | } 26 | 27 | require LSpec from git 28 | "https://github.com/argumentcomputer/LSpec.git" @ "db76512cd5266f0c576d561d8c69e2dc4890bea5" 29 | lean_lib Test { 30 | } 31 | @[test_driver] 32 | lean_exe test { 33 | root := `Test.Main 34 | -- Solves the native symbol not found problem 35 | supportInterpreter := true 36 | } 37 | -------------------------------------------------------------------------------- /.pre-commit-config.yaml: -------------------------------------------------------------------------------- 1 | default_install_hook_types: 2 | - pre-commit 3 | - commit-msg 4 | repos: 5 | - repo: https://github.com/pre-commit/pre-commit-hooks 6 | rev: v6.0.0 7 | hooks: 8 | - id: check-yaml 9 | - id: end-of-file-fixer 10 | - id: trailing-whitespace 11 | - id: check-symlinks 12 | - id: check-added-large-files 13 | - id: check-merge-conflict 14 | - id: detect-private-key 15 | - id: forbid-submodules 16 | - id: no-commit-to-branch 17 | args: [ --branch, main, --branch, dev ] 18 | - repo: https://github.com/compilerla/conventional-pre-commit 19 | rev: v4.3.0 20 | hooks: 21 | - id: conventional-pre-commit 22 | stages: [commit-msg] 23 | args: [feat, fix, doc, chore, test, build, merge, refactor] 24 | - repo: local 25 | hooks: 26 | - id: format nix 27 | name: Format Nix Code 28 | language: system 29 | entry: nix fmt 30 | files: .*\.nix$ 31 | -------------------------------------------------------------------------------- /doc/contributing.md: -------------------------------------------------------------------------------- 1 | # Contributing 2 | 3 | A Lean development shell is provided in the Nix flake. Nix usage is optional. 4 | Any contribution has to pass the pre-commit hooks: 5 | ```sh 6 | pre-commit install --install-hooks 7 | ``` 8 | 9 | All commit messages must conform to the Conventional Commits specification. 10 | 11 | ## Testing 12 | 13 | The tests are based on `LSpec`. To run tests, use either 14 | 15 | ``` sh 16 | nix flake check 17 | ``` 18 | or 19 | ``` sh 20 | lake test 21 | ``` 22 | 23 | You can run an individual test by specifying a prefix 24 | 25 | ``` sh 26 | lake test -- Frontend/Collect 27 | ``` 28 | 29 | ## Formatting 30 | 31 | When writing Lean code, follow the guidelines 32 | 33 | - Functions should be in `camelCase` 34 | - Theorems and tests should be in `snake_case` 35 | - Write the `|` in a pattern-matching `let` on the next line. This is for visual 36 | distinction with long function arguments. 37 | ```lean 38 | let .some result := function 39 | | fail "incorrect" 40 | ``` 41 | - Each test should be pinpointed and devolatilized. 42 | -------------------------------------------------------------------------------- /Test/Tactic/Assign.lean: -------------------------------------------------------------------------------- 1 | import Lean.Meta 2 | import Lean.Elab 3 | import LSpec 4 | import Test.Common 5 | 6 | open Lean 7 | 8 | namespace Pantograph.Test.Tactic.Assign 9 | 10 | def test_draft : TestT Elab.TermElabM Unit := do 11 | let expr := "forall (p : Prop), (p ∨ p) ∨ p" 12 | let skeleton := "by\nhave a : p ∨ p := sorry\nsorry" 13 | let expr ← parseSentence expr 14 | Meta.forallTelescope expr $ λ _ body => do 15 | let skeleton' ← match Parser.runParserCategory 16 | (env := ← MonadEnv.getEnv) 17 | (catName := `term) 18 | (input := skeleton) 19 | (fileName := ← getFileName) with 20 | | .ok syn => pure syn 21 | | .error error => throwError "Failed to parse: {error}" 22 | -- Apply the tactic 23 | let target ← Meta.mkFreshExprSyntheticOpaqueMVar body 24 | let tactic := Tactic.evalDraft skeleton' 25 | let newGoals ← runTacticOnMVar tactic target.mvarId! 26 | addTest $ LSpec.check "goals" ((← newGoals.mapM (λ g => do exprToStr (← g.getType))) = ["p ∨ p", "(p ∨ p) ∨ p"]) 27 | 28 | def suite (env: Environment): List (String × IO LSpec.TestSeq) := 29 | [ 30 | ("draft", test_draft), 31 | ] |>.map (λ (name, t) => (name, runTestTermElabM env t)) 32 | -------------------------------------------------------------------------------- /README.md: -------------------------------------------------------------------------------- 1 | # Pantograph 2 | 3 | A Machine-to-Machine interaction system for Lean 4. 4 | 5 | ![Pantograph](doc/icon.svg) 6 | 7 | Pantograph provides interfaces to interact with Lean's frontend, execute proofs, 8 | construct expressions, and examine the Lean environment. 9 | 10 | See [documentations](doc/rationale.md) for design rationale and references. 11 | 12 | ## Installation 13 | 14 | For Nix users, run 15 | ``` sh 16 | nix build .#{sharedLib,executable} 17 | ``` 18 | to build either the shared library or executable. 19 | 20 | For non-Nix users, install `lake` and `lean` fixed to the version of the 21 | `lean-toolchain` file, and run 22 | 23 | ``` sh 24 | lake build 25 | ``` 26 | This builds the executable in `.lake/build/bin/repl`. 27 | 28 | ### Executable Usage 29 | 30 | The default build target is a Read-Eval-Print-Loop (REPL). See [REPL 31 | Documentation](./doc/repl.md) 32 | 33 | Another executable is the `tomograph`, which processes a Lean file and displays 34 | syntax or elaboration level data. 35 | 36 | ### Library Usage 37 | 38 | `Pantograph/Library.lean` exposes a series of interfaces which allow FFI call 39 | with `Pantograph` which mirrors the REPL commands above. Note that there isn't a 40 | 1-1 correspondence between executable (REPL) commands and library functions. 41 | 42 | ## Development 43 | 44 | See [Contributing](./doc/contributing.md). 45 | -------------------------------------------------------------------------------- /Tomograph.lean: -------------------------------------------------------------------------------- 1 | /- A tool for analysing Lean source code. -/ 2 | import Pantograph.Frontend 3 | import Pantograph.Library 4 | 5 | open Lean 6 | 7 | namespace Pantograph 8 | 9 | def fail (s : String) : IO UInt32 := do 10 | IO.eprintln s 11 | return 2 12 | 13 | def dissect (args : List String) : IO UInt32 := do 14 | let fileName :: _args := args | fail s!"Must supply a file name" 15 | let file ← IO.FS.readFile fileName 16 | let (context, state) ← do Frontend.createContextStateFromFile file fileName (env? := .none) {} 17 | let frontendM: Frontend.FrontendM _ := 18 | Frontend.mapCompilationSteps λ step => do 19 | IO.println s!"{step.stx}" 20 | IO.println s!"🐈 {step.stx.getKind.toString}" 21 | for (tree, i) in step.trees.zipIdx do 22 | IO.println s!"🌲[{i}] {← tree.toString}" 23 | for (msg, i) in step.msgs.zipIdx do 24 | IO.println s!"🔈[{i}] {← msg.toString}" 25 | let (_, _) ← frontendM.run {} |>.run context |>.run state 26 | return 0 27 | 28 | end Pantograph 29 | 30 | open Pantograph 31 | 32 | def help : IO UInt32 := do 33 | IO.println "Usage: tomograph dissect FILE_NAME" 34 | return 1 35 | 36 | def main (args : List String) : IO UInt32 := do 37 | let command :: args := args | help 38 | unsafe do 39 | Pantograph.initSearch "" 40 | match command with 41 | | "dissect" => dissect args 42 | | _ => fail s!"Unknown command {command}" 43 | -------------------------------------------------------------------------------- /Pantograph/Elab.lean: -------------------------------------------------------------------------------- 1 | import Lean.Elab 2 | import Lean.Parser 3 | 4 | namespace Pantograph 5 | 6 | open Lean 7 | 8 | -- Functions for creating contexts and states 9 | @[export pantograph_default_elab_context] 10 | def defaultElabContext: Elab.Term.Context := { 11 | declName? := .some `mystery, 12 | errToSorry := false, 13 | } 14 | 15 | /-- Read syntax object from string -/ 16 | def parseTerm (env: Environment) (s: String): Except String Syntax := 17 | Parser.runParserCategory 18 | (env := env) 19 | (catName := `term) 20 | (input := s) 21 | (fileName := "") 22 | 23 | def parseTermM [Monad m] [MonadEnv m] (s: String): m (Except String Syntax) := do 24 | return Parser.runParserCategory 25 | (env := ← MonadEnv.getEnv) 26 | (catName := `term) 27 | (input := s) 28 | (fileName := "") 29 | 30 | /-- Parse a syntax object. May generate additional metavariables! -/ 31 | def elabType (syn: Syntax): Elab.TermElabM (Except String Expr) := do 32 | try 33 | let expr ← Elab.Term.elabType syn 34 | return .ok expr 35 | catch ex => return .error (← ex.toMessageData.toString) 36 | def elabTerm (syn: Syntax) (expectedType? : Option Expr := .none): Elab.TermElabM (Except String Expr) := do 37 | try 38 | let expr ← Elab.Term.elabTerm (stx := syn) expectedType? 39 | return .ok expr 40 | catch ex => return .error (← ex.toMessageData.toString) 41 | 42 | open Parser in 43 | def runParserCategory' (env : Environment) (catName : Name) (input : String) (fileName := "") : Except String (Syntax × String.Pos.Raw) := 44 | let p := adaptCacheableContextFn ({ · with savedPos? := .some 0 }) (categoryParserFnImpl catName) 45 | let ictx := mkInputContext input fileName 46 | let s := p.run ictx { env, options := {} } (getTokenTable env) (mkParserState input) 47 | if s.allErrors.isEmpty then 48 | Except.ok (s.stxStack.back, s.pos) 49 | else 50 | Except.error (s.toErrorMsg ictx) 51 | 52 | open Parser in 53 | def runParser (env : Environment) (parser : Parser) (input : String) (fileName := "") : Except String (Syntax × String.Pos.Raw) := 54 | let pfn := (withPosition parser).fn 55 | let ictx := mkInputContext input fileName 56 | let s := pfn.run ictx { env, options := {} } (getTokenTable env) (mkParserState input) 57 | if s.allErrors.isEmpty then 58 | Except.ok (s.stxStack.back, s.pos) 59 | else 60 | Except.error (s.toErrorMsg ictx) 61 | -------------------------------------------------------------------------------- /Test/Library.lean: -------------------------------------------------------------------------------- 1 | import Pantograph.Library 2 | import Test.Common 3 | 4 | namespace Pantograph.Test.Library 5 | 6 | open Lean 7 | 8 | def runTermElabM { α } (termElabM: Elab.TermElabM α): CoreM α := 9 | termElabM.run' (ctx := defaultElabContext) |>.run' 10 | 11 | def test_expr_echo (env: Environment): IO LSpec.TestSeq := do 12 | let inner: CoreM LSpec.TestSeq := do 13 | let prop_and_proof := "⟨∀ (x: Prop), x → x, λ (x: Prop) (h: x) => h⟩" 14 | let tests := LSpec.TestSeq.done 15 | let echoResult ← runTermElabM $ exprEcho prop_and_proof (options := {}) 16 | let tests := tests.append (LSpec.test "fail" (echoResult.toOption == .some { 17 | type := { pp? := "?m.1" }, expr := { pp? := "?m.2" } 18 | })) 19 | let echoResult ← runTermElabM $ exprEcho prop_and_proof (expectedType? := .some "Σ' p:Prop, p") (options := { printExprAST := true }) 20 | let tests := tests.append (LSpec.test "fail" (echoResult.toOption == .some { 21 | type := { 22 | pp? := "(p : Prop) ×' p", 23 | sexp? := "((:c PSigma) (:sort 0) (:lambda p (:sort 0) 0))", 24 | }, 25 | expr := { 26 | pp? := "⟨∀ (x : Prop), x → x, fun x h => h⟩", 27 | sexp? := "((:c PSigma.mk) (:sort 0) (:lambda p (:sort 0) 0) (:forall x (:sort 0) (:forall a 0 1)) (:lambda x (:sort 0) (:lambda h 0 0)))", 28 | } 29 | })) 30 | return tests 31 | runCoreMSeq env (options := #["pp.proofs.threshold=100"]) inner 32 | 33 | def test_core_context (env : Environment) : IO LSpec.TestSeq := runTest do 34 | let coreM := runTermElabM $ runTest do 35 | let goal := (← Meta.mkFreshExprSyntheticOpaqueMVar (.const `Nat [])).mvarId! 36 | let (_, _) ← (Tactic.collatz 27).run { elaborator := .anonymous} |>.run { goals := [goal] } 37 | fail "should fail" 38 | let coreContext ← createCoreContext (options := #["maxRecDepth=10"]) 39 | match ← (coreM.run' coreContext { env }).toBaseIO with 40 | | .error exception => 41 | let message ← exception.toMessageData.toString 42 | checkEq "exception" message "maximum recursion depth has been reached\nuse `set_option maxRecDepth ` to increase limit\nuse `set_option diagnostics true` to get diagnostic information" 43 | | .ok _ => 44 | fail "macRecDepth set should fail" 45 | let coreContext ← createCoreContext (options := #["maxRecDepth=200"]) 46 | match ← (coreM.run' coreContext { env }).toBaseIO with 47 | | .error exception => 48 | let message ← exception.toMessageData.toString 49 | fail s!"Exception: {message}" 50 | | .ok _ => 51 | pure () 52 | 53 | def suite (env: Environment): List (String × IO LSpec.TestSeq) := 54 | [ 55 | ("expr_echo", test_expr_echo env), 56 | ("core.context", test_core_context env) 57 | ] 58 | -------------------------------------------------------------------------------- /Test/Main.lean: -------------------------------------------------------------------------------- 1 | import Test.Delate 2 | import Test.Environment 3 | import Test.Frontend 4 | import Test.Integration 5 | import Test.Library 6 | import Test.Metavar 7 | import Test.Parser 8 | import Test.Proofs 9 | import Test.Serial 10 | import Test.Tactic 11 | 12 | import LSpec 13 | 14 | -- Test running infrastructure 15 | 16 | namespace Pantograph.Test 17 | 18 | def addPrefix (pref: String) (tests: List (String × α)): List (String × α) := 19 | tests.map (λ (name, x) => (pref ++ "/" ++ name, x)) 20 | 21 | abbrev TestTask := Task (Except IO.Error LSpec.TestSeq) 22 | def filterTestGroup (tests : List (String × IO LSpec.TestSeq)) (nameFilter? : Option String) 23 | : IO (List (String × TestTask)) := do 24 | let tests : List (String × IO LSpec.TestSeq) := match nameFilter? with 25 | | .some nameFilter => tests.filter (λ (name, _) => nameFilter.isPrefixOf name) 26 | | .none => tests 27 | tests.mapM λ (name, t) => return (name, ← IO.asTask t) 28 | 29 | /-- Runs test in parallel. Filters test name if given -/ 30 | def runTestTask (t : (String × TestTask)) : IO LSpec.TestSeq := do 31 | let (name, task) := t 32 | let v: Except IO.Error LSpec.TestSeq := task.get 33 | return match v with 34 | | .ok case => LSpec.group name case 35 | | .error e => expectationFailure name e.toString 36 | 37 | end Pantograph.Test 38 | 39 | open Pantograph.Test 40 | 41 | /-- Main entry of tests; Provide an argument to filter tests by prefix -/ 42 | def main (args: List String) := do 43 | let nameFilter? := args.head? 44 | Lean.initSearchPath (← Lean.findSysroot) 45 | let env_default : Lean.Environment ← Lean.importModules 46 | (imports := #[`Init]) 47 | (opts := {}) 48 | (trustLevel := 1) 49 | (loadExts := true) 50 | 51 | let suites: List (String × (Lean.Environment → List (String × IO LSpec.TestSeq))) := [ 52 | ("Environment", Environment.suite), 53 | ("Frontend/Collect", Frontend.Collect.suite), 54 | ("Frontend/Distil", Frontend.Distil.suite), 55 | ("Frontend/Refactor", Frontend.Refactor.suite), 56 | ("Integration", Integration.suite), 57 | ("Library", Library.suite), 58 | ("Metavar", Metavar.suite), 59 | ("Parser", Parser.suite), 60 | ("Proofs", Proofs.suite), 61 | ("Delate", Delate.suite), 62 | ("Serial", Serial.suite), 63 | ("Tactic/Assign", Tactic.Assign.suite), 64 | ("Tactic/Fragment", Tactic.Fragment.suite), 65 | ("Tactic/Prograde", Tactic.Prograde.suite), 66 | ] 67 | let suiterunner (f : Lean.Environment → List (String × IO LSpec.TestSeq)) := 68 | f env_default 69 | let tests : List (String × IO LSpec.TestSeq) := suites.foldl (init := []) λ acc (name, suite) => 70 | acc ++ (addPrefix name $ suiterunner suite) 71 | LSpec.lspecEachIO (← filterTestGroup tests nameFilter?) runTestTask 72 | -------------------------------------------------------------------------------- /Main.lean: -------------------------------------------------------------------------------- 1 | import Lean.Data.Json 2 | import Lean.Environment 3 | 4 | import Pantograph 5 | import Repl 6 | 7 | -- Main IO functions 8 | open Pantograph.Repl 9 | open Pantograph.Protocol 10 | 11 | /-- Print a string to stdout without buffering -/ 12 | def printImmediate (s : String) : IO Unit := do 13 | let stdout ← IO.getStdout 14 | stdout.putStr (s ++ "\n") 15 | stdout.flush 16 | 17 | /-- Parse a command either in `{ "cmd": ..., "payload": ... }` form or `cmd { ... }` form. -/ 18 | def parseCommand (s: String): Except String Command := do 19 | match String.Pos.Raw.get? s.trim 0 with 20 | | .some '{' => 21 | -- Parse in Json mode 22 | Lean.fromJson? (← Lean.Json.parse s) 23 | | .some _ => 24 | -- Parse in line mode 25 | let offset := s.posOf ' ' |> s.offsetOfPos 26 | if offset = s.length then 27 | return { cmd := s.take offset, payload := Lean.Json.null } 28 | else 29 | let payload ← s.drop offset |> Lean.Json.parse 30 | return { cmd := s.take offset, payload := payload } 31 | | .none => 32 | throw "Command is empty" 33 | 34 | partial def loop : MainM Unit := do repeat do 35 | let state ← get 36 | let command ← (← IO.getStdin).getLine 37 | -- Halt the program if empty line is given 38 | if command.trim.length = 0 then break 39 | match parseCommand command with 40 | | .error error => 41 | let error := Lean.toJson ({ error := "command", desc := error }: InteractionError) 42 | -- Using `Lean.Json.compress` here to prevent newline 43 | printImmediate error.compress 44 | | .ok command => 45 | try 46 | let ret ← execute command 47 | let str := match state.options.printJsonPretty with 48 | | true => ret.pretty 49 | | false => ret.compress 50 | printImmediate str 51 | catch e => 52 | let message := e.toString 53 | let error := Lean.toJson ({ error := "main", desc := message }: InteractionError) 54 | printImmediate error.compress 55 | 56 | def main (args: List String): IO Unit := do 57 | -- NOTE: A more sophisticated scheme of command line argument handling is needed. 58 | if args == ["--version"] then do 59 | IO.println s!"{Pantograph.version}" 60 | return 61 | 62 | unsafe do 63 | Pantograph.initSearch 64 | 65 | -- Separate imports and options 66 | let (options, imports) := args.partition (·.startsWith "--") 67 | let coreContext ← options.map (·.drop 2) |>.toArray |> Pantograph.createCoreContext 68 | let env ← Lean.importModules 69 | (imports := imports.toArray.map ({ module := ·.toName })) 70 | (opts := {}) 71 | (trustLevel := 1) 72 | (loadExts := true) 73 | try 74 | let mainM := loop.run { coreContext } |>.run' { env } 75 | printImmediate "ready." 76 | mainM 77 | catch ex => 78 | let message := ex.toString 79 | let error := Lean.toJson ({ error := "io", desc := message }: InteractionError) 80 | IO.println error.compress 81 | -------------------------------------------------------------------------------- /Pantograph/Tactic/Assign.lean: -------------------------------------------------------------------------------- 1 | import Lean.Elab 2 | import Lean.Meta 3 | 4 | open Lean 5 | 6 | namespace Pantograph.Tactic 7 | 8 | /-- WARNING: This should be used with a function like `elabTermWithHoles` that properly collects the mvar information from `expr`. -/ 9 | def assign (goal: MVarId) (expr: Expr) (nextGoals: List MVarId): MetaM (List MVarId) := do 10 | goal.checkNotAssigned `Pantograph.Tactic.assign 11 | 12 | -- This run of the unifier is critical in resolving mvars in passing 13 | let exprType ← Meta.inferType expr 14 | let goalType ← goal.getType 15 | unless ← Meta.isDefEq goalType exprType do 16 | throwError s!"{← Meta.ppExpr expr} : {← Meta.ppExpr exprType} ≠ {← Meta.ppExpr goalType}" 17 | goal.assign expr 18 | nextGoals.filterM (not <$> ·.isAssigned) 19 | 20 | def evalAssign : Elab.Tactic.Tactic := fun stx => Elab.Tactic.withMainContext do 21 | let target ← Elab.Tactic.getMainTarget 22 | let goal ← Elab.Tactic.getMainGoal 23 | goal.checkNotAssigned `Pantograph.Tactic.evalAssign 24 | let (expr, nextGoals) ← Elab.Tactic.elabTermWithHoles stx 25 | (expectedType? := .some target) 26 | (tagSuffix := .anonymous ) 27 | (allowNaturalHoles := true) 28 | goal.assign expr 29 | Elab.Tactic.replaceMainGoal nextGoals 30 | 31 | /-- Converts `sorry`s in the source expression to goals. Execute `post` on the 32 | types of the sorrys. -/ 33 | def sorryToHole (src : Expr) (post : Expr → MetaM Expr := pure) 34 | : StateRefT (List MVarId) MetaM Expr := 35 | Meta.transform src λ expr => 36 | if expr.isSorry then do 37 | let type ← instantiateMVars (expr.getArg! 0 |>.bindingBody!) 38 | if type.hasSorry then 39 | throwError s!"Coupling is not allowed in draft tactic: {← Meta.ppExpr type}" 40 | let type ← post type 41 | let mvar ← Meta.mkFreshExprSyntheticOpaqueMVar type 42 | modify (mvar.mvarId! :: .) 43 | return .done mvar 44 | else 45 | return .continue 46 | 47 | -- Given a complete (no holes) expression, extract the sorry's from it and convert them into goals. 48 | def draft (goal : MVarId) (expr : Expr) : MetaM (List MVarId) := do 49 | goal.checkNotAssigned `Pantograph.Tactic.draft 50 | let exprType ← Meta.inferType expr 51 | let goalType ← goal.getType 52 | unless ← Meta.isDefEq goalType exprType do 53 | throwError s!"{← Meta.ppExpr expr} : {← Meta.ppExpr exprType} ≠ {← Meta.ppExpr goalType}" 54 | 55 | let (expr', holes) ← sorryToHole expr |>.run [] 56 | goal.assign expr' 57 | return holes.reverse 58 | 59 | def evalDraft : Elab.Tactic.Tactic := fun stx ↦ Elab.Tactic.withMainContext do 60 | let target ← Elab.Tactic.getMainTarget 61 | let goal ← Elab.Tactic.getMainGoal 62 | let (expr, holeGoals) ← Elab.Tactic.elabTermWithHoles stx 63 | (expectedType? := .some target) 64 | (tagSuffix := .anonymous) 65 | (allowNaturalHoles := true) 66 | let draftGoals ← draft goal expr 67 | Elab.Tactic.replaceMainGoal $ holeGoals ++ draftGoals 68 | 69 | 70 | end Pantograph.Tactic 71 | -------------------------------------------------------------------------------- /doc/rationale.md: -------------------------------------------------------------------------------- 1 | # Design Rationale 2 | 3 | A great problem in machine learning is to use ML agents to automatically prove 4 | mathematical theorems. This sort of proof necessarily involves *search*. 5 | Compatibility for search is the main reason for creating Pantograph. The Lean 4 6 | LSP interface is not conducive to search. Pantograph is designed with this in 7 | mind. It emphasizes the difference between 3 views of a proof: 8 | 9 | - **Presentation View**: The view of a written, polished proof. e.g. Mathlib and 10 | math papers are almost always written in this form. 11 | - **Search View**: The view of a proof exploration trajectory. This is not 12 | explicitly supported by Lean LSP. 13 | - **Kernel View**: The proof viewed as a set of metavariables. 14 | 15 | Pantograph enables proof agents to operate on the search view. 16 | 17 | ## Name 18 | 19 | The name Pantograph is a pun. It means two things 20 | - A pantograph is an instrument for copying down writing. As an agent explores 21 | the vast proof search space, Pantograph records the current state to ensure 22 | the proof is sound. 23 | - A pantograph is also an equipment for an electric train. It supplies power to 24 | a locomotive. In comparison the (relatively) simple Pantograph software powers 25 | theorem proving projects. 26 | 27 | ## Caveats and Limitations 28 | 29 | Pantograph does not exactly mimic Lean LSP's behaviour. That would not grant the 30 | flexibility it offers. To support tree search means Pantograph has to act 31 | differently from Lean in some times, but never at the sacrifice of soundness. 32 | 33 | - When Lean LSP says "don't know how to synthesize placeholder", this indicates 34 | the human operator needs to manually move the cursor to the placeholder and 35 | type in the correct expression. This error therefore should not halt the proof 36 | process, and the placeholder should be turned into a goal. 37 | - When Lean LSP says "unresolved goals", that means a proof cannot finish where 38 | it is supposed to finish at the end of a `by` block. Pantograph will raise the 39 | error in this case, since it indicates the termination of a proof search branch. 40 | 41 | Pantograph cannot perform things that are inherently constrained by Lean. These 42 | include: 43 | 44 | - If a tactic loses track of metavariables, it will not be caught until the end 45 | of the proof search. This is a bug in the tactic itself. 46 | - Although a timeout feature exists in Pantograph, it relies on the coöperative 47 | multitasking from the tactic implementation. There is nothing preventing a 48 | buggy tactic from stalling Lean if it does not check for cancellation often. 49 | - For the same reason as above, there is no graceful way to stop a tactic which 50 | leaks infinite memory. Users who wish to have this behaviour should run 51 | Pantograph in a controlled environment with limited allocations. e.g. 52 | Linux control groups. 53 | - Interceptions of parsing errors generally cannot be turned into goals (e.g. 54 | `def mystery : Nat := :=`) due to Lean's parsing system. This question is also 55 | not well-defined. 56 | 57 | ## References 58 | 59 | * [Pantograph Paper](https://arxiv.org/abs/2410.16429) 60 | 61 | -------------------------------------------------------------------------------- /Test/Frontend/Refactor.lean: -------------------------------------------------------------------------------- 1 | import Pantograph 2 | import Test.Common 3 | 4 | open Lean Pantograph Frontend 5 | 6 | namespace Pantograph.Test.Frontend.Refactor 7 | 8 | abbrev Test := Environment → TestT IO Unit 9 | 10 | deriving instance Repr, DecidableEq for FileMap 11 | 12 | private def test_merge_file_map : Test := λ _env ↦ do 13 | let src1 := " 14 | set_option pp.explicit true 15 | open Nat in 16 | def f : Nat → Nat := id 17 | ".trim 18 | let src2 := " 19 | set_option pp.explicit true 20 | open Nat in 21 | def f : Nat → Nat := 22 | id 23 | ".trim 24 | let filemap1 := s!"{src1}\n{src2}".toFileMap 25 | let filemap2 := Refactor.mergeFileMap src1.toFileMap src2.toFileMap 26 | checkEq "result" filemap1 filemap2 27 | 28 | example : Σ' f : Nat → Nat, ∀ (n : Nat), f n = n := by 29 | constructor 30 | intro n; rfl 31 | 32 | private def test_id : Test := λ env ↦ do 33 | let src := " 34 | set_option pp.explicit true 35 | open Nat in 36 | def f : Nat → Nat := id 37 | " 38 | let expected := " 39 | set_option pp.explicit true 40 | open Nat in 41 | def f : Nat → Nat := 42 | id 43 | ".trim 44 | let result ← runRefactor env src 45 | checkEq "result" result.trim expected 46 | 47 | private def test_simple : Test := λ env ↦ do 48 | let src := " 49 | /-- S1 -/ 50 | def f : Nat → Nat := sorry 51 | theorem mystery (n : Nat) : f n = n := sorry 52 | " 53 | let expected := " 54 | /-- S1 -/ 55 | def f_composite : { f : Nat → Nat // ∀ (n : Nat), f n = n } := 56 | sorry 57 | ".trim 58 | let result ← runRefactor env src 59 | checkEq "result" result.trim expected 60 | 61 | private def test_invalid : Test := λ env ↦ do 62 | let src := " 63 | /-- S1 -/ 64 | def f : Nat → Nat := sorry 65 | theorem mystery (n : Nat) , f n = n := sorry 66 | " 67 | try 68 | let _ ← runRefactor env src 69 | fail "Should fail" 70 | catch ex : IO.Error => 71 | checkEq "error" ex.toString s!"{defaultFileName}:4:25: error: unexpected token ','; expected ':'\n" 72 | 73 | private def test_intercalating : Test := λ env ↦ do 74 | let src := " 75 | def f : Nat → Nat := sorry 76 | def helper (n : Nat) : Nat := n + 1 77 | theorem mystery (n : Nat) : f n = helper n := sorry 78 | " 79 | let expected := " 80 | def helper (n : Nat) : Nat := 81 | n + 1 82 | def f_composite : { f : Nat → Nat // ∀ (n : Nat), f n = helper n } := 83 | sorry 84 | ".trim 85 | let result ← runRefactor env src 86 | checkEq "result" result.trim expected 87 | 88 | private def test_predicate : Test := λ env ↦ do 89 | let src := " 90 | def q : (Nat → Nat) → Prop := sorry 91 | def p : (Nat → Nat) → Prop := sorry 92 | theorem mystery : p Nat.succ := sorry 93 | " 94 | let expected := " 95 | def q : (Nat → Nat) → Prop := 96 | sorry 97 | def p_composite : { p : (Nat → Nat) → Prop // p Nat.succ } := 98 | sorry 99 | ".trim 100 | let result ← runRefactor env src 101 | checkEq "result" result.trim expected 102 | 103 | def suite (env : Environment): List (String × IO LSpec.TestSeq) := 104 | let tests := [ 105 | ("merge file map", test_merge_file_map), 106 | ("id", test_id), 107 | ("simple", test_simple), 108 | ("invalid", test_invalid), 109 | ("intercalating", test_intercalating), 110 | ("predicate", test_predicate), 111 | ] 112 | tests.map λ (name, test) => (name, runTest $ test env) 113 | -------------------------------------------------------------------------------- /Pantograph/Tactic/Prograde.lean: -------------------------------------------------------------------------------- 1 | /- Prograde (forward) reasoning tactics -/ 2 | import Lean.Meta 3 | import Lean.Elab.Tactic 4 | 5 | namespace Pantograph.Tactic 6 | 7 | open Lean 8 | 9 | private def mkUpstreamMVar (goal: MVarId) : MetaM Expr := do 10 | Meta.mkFreshExprSyntheticOpaqueMVar (← goal.getType) (tag := ← goal.getTag) 11 | 12 | 13 | /-- Introduces a fvar to the current mvar -/ 14 | def define (mvarId: MVarId) (binderName: Name) (expr: Expr): MetaM (FVarId × MVarId) := mvarId.withContext do 15 | mvarId.checkNotAssigned `Pantograph.Tactic.define 16 | let type ← Meta.inferType expr 17 | 18 | Meta.withLetDecl binderName type expr λ fvar => do 19 | let mvarUpstream ← mkUpstreamMVar mvarId 20 | mvarId.assign $ ← Meta.mkLetFVars #[fvar] mvarUpstream 21 | pure (fvar.fvarId!, mvarUpstream.mvarId!) 22 | 23 | def evalDefine (binderName: Name) (expr: Syntax): Elab.Tactic.TacticM Unit := do 24 | let goal ← Elab.Tactic.getMainGoal 25 | let expr ← goal.withContext $ Elab.Term.elabTerm (stx := expr) (expectedType? := .none) 26 | let (_, mvarId) ← define goal binderName expr 27 | Elab.Tactic.replaceMainGoal [mvarId] 28 | 29 | structure BranchResult where 30 | fvarId?: Option FVarId := .none 31 | branch: MVarId 32 | main: MVarId 33 | 34 | def «have» (mvarId: MVarId) (binderName: Name) (type: Expr): MetaM BranchResult := mvarId.withContext do 35 | mvarId.checkNotAssigned `Pantograph.Tactic.have 36 | let lctx ← MonadLCtx.getLCtx 37 | -- The branch goal inherits the same context, but with a different type 38 | let mvarBranch ← Meta.mkFreshExprMVarAt lctx (← Meta.getLocalInstances) type 39 | 40 | -- Create the context for the `upstream` goal 41 | let fvarId ← mkFreshFVarId 42 | let lctxUpstream := lctx.mkLocalDecl fvarId binderName type 43 | let mvarUpstream ← 44 | Meta.withLCtx lctxUpstream #[] do 45 | Meta.withNewLocalInstances #[.fvar fvarId] 0 do 46 | let mvarUpstream ← mkUpstreamMVar mvarId 47 | --let expr: Expr := .app (.lam binderName type mvarBranch .default) mvarUpstream 48 | mvarId.assign $ ← Meta.mkLambdaFVars #[.fvar fvarId] mvarUpstream 49 | pure mvarUpstream 50 | 51 | return { 52 | fvarId? := .some fvarId, 53 | branch := mvarBranch.mvarId!, 54 | main := mvarUpstream.mvarId!, 55 | } 56 | 57 | def evalHave (binderName: Name) (type: Syntax): Elab.Tactic.TacticM Unit := do 58 | let goal ← Elab.Tactic.getMainGoal 59 | let nextGoals: List MVarId ← goal.withContext do 60 | let type ← Elab.Term.elabType (stx := type) 61 | let result ← «have» goal binderName type 62 | pure [result.branch, result.main] 63 | Elab.Tactic.replaceMainGoal nextGoals 64 | 65 | def «let» (mvarId: MVarId) (binderName: Name) (type: Expr): MetaM BranchResult := mvarId.withContext do 66 | mvarId.checkNotAssigned `Pantograph.Tactic.let 67 | let lctx ← MonadLCtx.getLCtx 68 | 69 | -- The branch goal inherits the same context, but with a different type 70 | let mvarBranch ← Meta.mkFreshExprMVarAt lctx (← Meta.getLocalInstances) type (userName := binderName) 71 | 72 | assert! ¬ type.hasLooseBVars 73 | let mvarUpstream ← Meta.withLetDecl binderName type mvarBranch $ λ fvar => do 74 | let mvarUpstream ← mkUpstreamMVar mvarId 75 | mvarId.assign $ ← Meta.mkLetFVars #[fvar] mvarUpstream 76 | pure mvarUpstream 77 | 78 | return { 79 | branch := mvarBranch.mvarId!, 80 | main := mvarUpstream.mvarId!, 81 | } 82 | 83 | def evalLet (binderName: Name) (type: Syntax): Elab.Tactic.TacticM Unit := do 84 | let goal ← Elab.Tactic.getMainGoal 85 | let type ← goal.withContext $ Elab.Term.elabType (stx := type) 86 | let result ← «let» goal binderName type 87 | Elab.Tactic.replaceMainGoal [result.branch, result.main] 88 | -------------------------------------------------------------------------------- /flake.lock: -------------------------------------------------------------------------------- 1 | { 2 | "nodes": { 3 | "flake-parts": { 4 | "inputs": { 5 | "nixpkgs-lib": "nixpkgs-lib" 6 | }, 7 | "locked": { 8 | "lastModified": 1754487366, 9 | "narHash": "sha256-pHYj8gUBapuUzKV/kN/tR3Zvqc7o6gdFB9XKXIp1SQ8=", 10 | "owner": "hercules-ci", 11 | "repo": "flake-parts", 12 | "rev": "af66ad14b28a127c5c0f3bbb298218fc63528a18", 13 | "type": "github" 14 | }, 15 | "original": { 16 | "owner": "hercules-ci", 17 | "repo": "flake-parts", 18 | "type": "github" 19 | } 20 | }, 21 | "flake-parts_2": { 22 | "inputs": { 23 | "nixpkgs-lib": "nixpkgs-lib_2" 24 | }, 25 | "locked": { 26 | "lastModified": 1765835352, 27 | "narHash": "sha256-XswHlK/Qtjasvhd1nOa1e8MgZ8GS//jBoTqWtrS1Giw=", 28 | "owner": "hercules-ci", 29 | "repo": "flake-parts", 30 | "rev": "a34fae9c08a15ad73f295041fec82323541400a9", 31 | "type": "github" 32 | }, 33 | "original": { 34 | "owner": "hercules-ci", 35 | "repo": "flake-parts", 36 | "type": "github" 37 | } 38 | }, 39 | "lean4-nix": { 40 | "inputs": { 41 | "flake-parts": "flake-parts_2", 42 | "nixpkgs": "nixpkgs" 43 | }, 44 | "locked": { 45 | "lastModified": 1766008374, 46 | "narHash": "sha256-ABLOpSeybmSa2+eWQVVUKh2C9/JzLIszSK3Zx7szMX0=", 47 | "owner": "lenianiva", 48 | "repo": "lean4-nix", 49 | "rev": "e1e7451050e9a27fd469b544cef174fe9feca0e1", 50 | "type": "github" 51 | }, 52 | "original": { 53 | "owner": "lenianiva", 54 | "repo": "lean4-nix", 55 | "type": "github" 56 | } 57 | }, 58 | "nixpkgs": { 59 | "locked": { 60 | "lastModified": 1765779637, 61 | "narHash": "sha256-KJ2wa/BLSrTqDjbfyNx70ov/HdgNBCBBSQP3BIzKnv4=", 62 | "owner": "nixos", 63 | "repo": "nixpkgs", 64 | "rev": "1306659b587dc277866c7b69eb97e5f07864d8c4", 65 | "type": "github" 66 | }, 67 | "original": { 68 | "owner": "nixos", 69 | "ref": "nixos-unstable", 70 | "repo": "nixpkgs", 71 | "type": "github" 72 | } 73 | }, 74 | "nixpkgs-lib": { 75 | "locked": { 76 | "lastModified": 1753579242, 77 | "narHash": "sha256-zvaMGVn14/Zz8hnp4VWT9xVnhc8vuL3TStRqwk22biA=", 78 | "owner": "nix-community", 79 | "repo": "nixpkgs.lib", 80 | "rev": "0f36c44e01a6129be94e3ade315a5883f0228a6e", 81 | "type": "github" 82 | }, 83 | "original": { 84 | "owner": "nix-community", 85 | "repo": "nixpkgs.lib", 86 | "type": "github" 87 | } 88 | }, 89 | "nixpkgs-lib_2": { 90 | "locked": { 91 | "lastModified": 1765674936, 92 | "narHash": "sha256-k00uTP4JNfmejrCLJOwdObYC9jHRrr/5M/a/8L2EIdo=", 93 | "owner": "nix-community", 94 | "repo": "nixpkgs.lib", 95 | "rev": "2075416fcb47225d9b68ac469a5c4801a9c4dd85", 96 | "type": "github" 97 | }, 98 | "original": { 99 | "owner": "nix-community", 100 | "repo": "nixpkgs.lib", 101 | "type": "github" 102 | } 103 | }, 104 | "nixpkgs_2": { 105 | "locked": { 106 | "lastModified": 1755274400, 107 | "narHash": "sha256-rTInmnp/xYrfcMZyFMH3kc8oko5zYfxsowaLv1LVobY=", 108 | "owner": "nixos", 109 | "repo": "nixpkgs", 110 | "rev": "ad7196ae55c295f53a7d1ec39e4a06d922f3b899", 111 | "type": "github" 112 | }, 113 | "original": { 114 | "owner": "nixos", 115 | "ref": "nixos-25.05", 116 | "repo": "nixpkgs", 117 | "type": "github" 118 | } 119 | }, 120 | "root": { 121 | "inputs": { 122 | "flake-parts": "flake-parts", 123 | "lean4-nix": "lean4-nix", 124 | "nixpkgs": "nixpkgs_2" 125 | } 126 | } 127 | }, 128 | "root": "root", 129 | "version": 7 130 | } 131 | -------------------------------------------------------------------------------- /flake.nix: -------------------------------------------------------------------------------- 1 | { 2 | description = "Pantograph"; 3 | 4 | inputs = { 5 | nixpkgs.url = "github:nixos/nixpkgs/nixos-25.05"; 6 | flake-parts.url = "github:hercules-ci/flake-parts"; 7 | lean4-nix.url = "github:lenianiva/lean4-nix"; 8 | }; 9 | 10 | outputs = inputs @ { 11 | self, 12 | nixpkgs, 13 | flake-parts, 14 | lean4-nix, 15 | ... 16 | }: 17 | flake-parts.lib.mkFlake {inherit inputs;} { 18 | flake = { 19 | }; 20 | systems = [ 21 | "aarch64-linux" 22 | "aarch64-darwin" 23 | "x86_64-linux" 24 | "x86_64-darwin" 25 | ]; 26 | perSystem = { 27 | system, 28 | pkgs, 29 | ... 30 | }: let 31 | manifest = pkgs.lib.importJSON ./lake-manifest.json; 32 | manifest-lspec = builtins.head manifest.packages; 33 | lean = lean4-nix.packages.${system}.lean; 34 | lspecLib = lean.buildLeanPackage { 35 | name = "LSpec"; 36 | roots = ["LSpec"]; 37 | src = builtins.fetchGit {inherit (manifest-lspec) url rev;}; 38 | }; 39 | inherit (pkgs.lib.fileset) unions toSource fileFilter; 40 | src = ./.; 41 | set-project = unions [ 42 | ./Pantograph.lean 43 | (fileFilter (file: file.hasExt "lean") ./Pantograph) 44 | ]; 45 | set-test = unions [ 46 | (fileFilter (file: file.hasExt "lean") ./Test) 47 | ]; 48 | src-project = toSource { 49 | root = src; 50 | fileset = unions [ 51 | set-project 52 | ]; 53 | }; 54 | src-repl = toSource { 55 | root = src; 56 | fileset = unions [ 57 | set-project 58 | ./Main.lean 59 | ./Repl.lean 60 | ]; 61 | }; 62 | src-tomograph = toSource { 63 | root = src; 64 | fileset = unions [ 65 | set-project 66 | ./Tomograph.lean 67 | ]; 68 | }; 69 | src-test = toSource { 70 | root = src; 71 | fileset = unions [ 72 | set-project 73 | ./Repl.lean 74 | set-test 75 | ]; 76 | }; 77 | project = lean.buildLeanPackage { 78 | name = "Pantograph"; 79 | roots = ["Pantograph"]; 80 | src = src-project; 81 | }; 82 | repl = lean.buildLeanPackage { 83 | name = "Repl"; 84 | roots = ["Main" "Repl"]; 85 | deps = [project]; 86 | src = src-repl; 87 | }; 88 | tomograph = lean.buildLeanPackage { 89 | name = "tomograph"; 90 | roots = ["Tomograph"]; 91 | deps = [project]; 92 | src = src-tomograph; 93 | }; 94 | test = lean.buildLeanPackage { 95 | name = "Test"; 96 | # NOTE: The src directory must be ./. since that is where the import 97 | # root begins (e.g. `import Test.Environment` and not `import 98 | # Environment`) and thats where `lakefile.lean` resides. 99 | roots = ["Test.Main"]; 100 | deps = [lspecLib repl]; 101 | src = src-test; 102 | }; 103 | in rec { 104 | packages = { 105 | inherit (project) sharedLib depRoots; 106 | inherit (repl) executable; 107 | tomograph = tomograph.executable; 108 | default = repl.executable; 109 | }; 110 | legacyPackages = { 111 | inherit project; 112 | }; 113 | checks = { 114 | test = 115 | pkgs.runCommand "test" { 116 | buildInputs = [test.executable lean.lean-all]; 117 | } '' 118 | #export LEAN_SRC_PATH="${./.}" 119 | ${test.executable}/bin/test > $out 120 | ''; 121 | }; 122 | formatter = pkgs.alejandra; 123 | devShells.default = pkgs.mkShell { 124 | buildInputs = [ 125 | pkgs.pre-commit 126 | lean.lean-all 127 | ]; 128 | }; 129 | }; 130 | }; 131 | } 132 | -------------------------------------------------------------------------------- /Test/Environment.lean: -------------------------------------------------------------------------------- 1 | import Pantograph.Delate 2 | import Pantograph.Environment 3 | import Repl 4 | import Test.Common 5 | 6 | import LSpec 7 | 8 | open Lean 9 | namespace Pantograph.Test.Environment 10 | 11 | open Pantograph 12 | 13 | deriving instance DecidableEq, Repr for Protocol.InductInfo 14 | deriving instance DecidableEq, Repr for Protocol.ConstructorInfo 15 | deriving instance DecidableEq, Repr for Protocol.RecursorRule 16 | deriving instance DecidableEq, Repr for Protocol.RecursorInfo 17 | deriving instance DecidableEq, Repr for Protocol.EnvInspectResult 18 | 19 | def test_catalog (env : Environment) : IO LSpec.TestSeq := do 20 | let symbolsWithNum := env.constants.fold (init := #[]) (λ acc name info => 21 | match (toFilteredSymbol name info).isSome && (name matches .num _ _) with 22 | | false => acc 23 | | true => acc.push name 24 | ) 25 | return LSpec.check "No num symbols" (symbolsWithNum.size == 0) 26 | 27 | def test_symbol_visibility: IO LSpec.TestSeq := do 28 | let entries: List (Name × Bool) := [ 29 | ("Nat.add_comm".toName, false), 30 | ("foo.bla.Init.Data.List.Basic.2.1.Init.Lean.Expr._hyg.4".toName, true), 31 | ] 32 | let suite := entries.foldl (λ suites (symbol, target) => 33 | let test := LSpec.check symbol.toString ((isNameInternal symbol) == target) 34 | LSpec.TestSeq.append suites test) LSpec.TestSeq.done 35 | return suite 36 | 37 | inductive ConstantCat where 38 | | induct (info: Protocol.InductInfo) 39 | | ctor (info: Protocol.ConstructorInfo) 40 | | recursor (info: Protocol.RecursorInfo) 41 | 42 | def test_inspect (env : Environment) : IO LSpec.TestSeq := do 43 | let testCases: List (Name × ConstantCat) := [ 44 | (`Or, ConstantCat.induct { 45 | numParams := 2, 46 | numIndices := 0, 47 | all := #[`Or], 48 | ctors := #[`Or.inl, `Or.inr], 49 | }), 50 | (`Except.ok, ConstantCat.ctor { 51 | induct := `Except, 52 | cidx := 1, 53 | numParams := 2, 54 | numFields := 1, 55 | }), 56 | (`Eq.rec, ConstantCat.recursor { 57 | all := #[`Eq], 58 | numParams := 2, 59 | numIndices := 1, 60 | numMotives := 1, 61 | numMinors := 1, 62 | rules := #[{ ctor := `Eq.refl, nFields := 0, rhs := { pp? := .some "fun {α} a motive refl => refl" } }] 63 | k := true, 64 | }), 65 | (`ForM.rec, ConstantCat.recursor { 66 | all := #[`ForM], 67 | numParams := 3, 68 | numIndices := 0, 69 | numMotives := 1, 70 | numMinors := 1, 71 | rules := #[{ ctor := `ForM.mk, nFields := 1, rhs := { pp? := .some "fun m γ α motive mk forM => mk forM" } }] 72 | k := false, 73 | }) 74 | ] 75 | let inner: CoreM LSpec.TestSeq := do 76 | let coreContext ← readThe Core.Context 77 | testCases.foldlM (λ acc (name, cat) => do 78 | let args: Protocol.EnvInspect := { name } 79 | let result ← match ← Repl.env_inspect args |>.run { coreContext } |>.run' { env := ← getEnv } with 80 | | .ok result => pure $ result 81 | | .error e => panic! s!"Error: {e.desc}" 82 | let testName := name.toString 83 | let p := match cat with 84 | | .induct info => LSpec.test testName (result.inductInfo? == .some info) 85 | | .ctor info => LSpec.test testName (result.constructorInfo? == .some info) 86 | | .recursor info => LSpec.test testName (result.recursorInfo? == .some info) 87 | return LSpec.TestSeq.append acc p 88 | ) LSpec.TestSeq.done 89 | runCoreMSeq env inner 90 | 91 | def test_symbol_location (env : Environment) : TestT IO Unit := do 92 | addTest $ ← runTestCoreM env do 93 | let .ok result ← Repl.env_inspect { name := `Nat.le_of_succ_le, source? := .some true } 94 | |>.run { coreContext := ← readThe Core.Context } 95 | |>.run' { env := ← getEnv } 96 | | fail "Inspect failed" 97 | checkEq "module" result.module? <| .some `Init.Data.Nat.Basic 98 | 99 | -- Extraction of source doesn't work for symbols in `Init` for some reason 100 | checkTrue "file" result.sourceUri?.isNone 101 | checkEq "sourceStart" (result.sourceStart?.map (·.column)) <| .some 0 102 | checkEq "sourceEnd" (result.sourceEnd?.map (·.column)) <| .some 88 103 | let .ok { imports, constNames, .. } ← Repl.env_module_read ⟨`Init.Data.Nat.Basic⟩ 104 | |>.run { coreContext := ← readThe Core.Context } 105 | |>.run' { env := ← getEnv } 106 | | fail "Module read failed" 107 | checkEq "imports" imports #[`Init.SimpLemmas, `Init.Data.NeZero, `Init.Grind.Tactics] 108 | checkTrue "constNames" $ constNames.contains `Nat.succ_add 109 | 110 | def test_matcher (env : Environment) : TestT IO Unit := do 111 | checkTrue "not matcher" $ ¬ Meta.isMatcherCore env `Nat.strongRecOn 112 | 113 | def suite (env : Environment) : List (String × IO LSpec.TestSeq) := 114 | [ 115 | ("Catalog", test_catalog env), 116 | ("Symbol Visibility", test_symbol_visibility), 117 | ("Inspect", test_inspect env), 118 | ("Symbol Location", runTest $ test_symbol_location env), 119 | ("Matcher", runTest $ test_matcher env), 120 | ] 121 | -------------------------------------------------------------------------------- /Test/Serial.lean: -------------------------------------------------------------------------------- 1 | import Pantograph.Library 2 | import Pantograph.Serial 3 | 4 | import LSpec 5 | import Test.Common 6 | 7 | open Lean 8 | 9 | namespace Pantograph.Test.Serial 10 | 11 | structure MultiState where 12 | coreContext : Core.Context 13 | env: Environment 14 | 15 | abbrev TestM := TestT $ StateRefT MultiState $ IO 16 | 17 | instance : MonadEnv TestM where 18 | getEnv := return (← getThe MultiState).env 19 | modifyEnv f := do modifyThe MultiState fun s => { s with env := f s.env } 20 | 21 | def runCoreM { α } (state : Core.State) (testCoreM : TestT CoreM α) : TestM (α × Core.State) := do 22 | let multiState ← getThe MultiState 23 | let coreM := runTestWithResult testCoreM 24 | match ← (coreM.run multiState.coreContext state).toBaseIO with 25 | | .error e => do 26 | throw $ .userError $ ← e.toMessageData.toString 27 | | .ok ((a, tests), state') => do 28 | set $ (← getThe LSpec.TestSeq) ++ tests 29 | return (a, state') 30 | 31 | def test_pickling_environment : TestM Unit := do 32 | let coreSrc : Core.State := { env := ← getEnv } 33 | let coreDst : Core.State := { env := ← getEnv } 34 | 35 | let name := `mystery 36 | IO.FS.withTempFile λ _ envPicklePath => do 37 | let ((), _) ← runCoreM coreSrc do 38 | let type: Expr := .forallE `p (.sort 0) (.forallE `h (.bvar 0) (.bvar 1) .default) .default 39 | let value: Expr := .lam `p (.sort 0) (.lam `h (.bvar 0) (.bvar 0) .default) .default 40 | let c := Declaration.defnDecl <| mkDefinitionValEx 41 | (name := name) 42 | (levelParams := []) 43 | (type := type) 44 | (value := value) 45 | (hints := mkReducibilityHintsRegularEx 1) 46 | (safety := .safe) 47 | (all := []) 48 | addDecl c 49 | environmentPickle (← getEnv) envPicklePath 50 | 51 | let _ ← runCoreM coreDst do 52 | let (env', _) ← environmentUnpickle envPicklePath 53 | checkTrue s!"Has symbol {name}" (env'.find? name).isSome 54 | let anotherName := `mystery2 55 | checkTrue s!"Doesn't have symbol {anotherName}" (env'.find? anotherName).isNone 56 | 57 | def test_goal_state_simple : TestM Unit := do 58 | let coreSrc : Core.State := { env := ← getEnv } 59 | let coreDst : Core.State := { env := ← getEnv } 60 | IO.FS.withTempFile λ _ statePath => do 61 | let type: Expr := .forallE `p (.sort 0) (.forallE `h (.bvar 0) (.bvar 1) .default) .default 62 | let stateGenerate : MetaM GoalState := runTermElabMInMeta do 63 | GoalState.create type 64 | 65 | let ((), _) ← runCoreM coreSrc do 66 | let state ← stateGenerate.run' 67 | goalStatePickle state statePath 68 | 69 | let ((), _) ← runCoreM coreDst do 70 | let (goalState, _) ← goalStateUnpickle statePath (← getEnv) 71 | let metaM : MetaM (List Expr) := do 72 | goalState.goals.mapM λ goal => goalState.withContext goal goal.getType 73 | let types ← metaM.run' 74 | checkTrue "Goals" $ types[0]!.equal type 75 | 76 | def test_pickling_env_extensions : TestM Unit := do 77 | let coreSrc : Core.State := { env := ← getEnv } 78 | let coreDst : Core.State := { env := ← getEnv } 79 | IO.FS.withTempFile λ _ statePath => do 80 | let ((), _) ← runCoreM coreSrc $ transformTestT runTermElabMInCore do 81 | let .ok e ← elabTerm (← `(term|(2: Nat) ≤ 3 ∧ (3: Nat) ≤ 5)) .none | unreachable! 82 | let state ← GoalState.create e 83 | let .success state _ ← state.tacticOn' 0 (← `(tactic|apply And.intro)) | unreachable! 84 | 85 | let goal := state.goals[0]! 86 | let type ← goal.withContext do 87 | let .ok type ← elabTerm (← `(term|(2: Nat) ≤ 3)) (.some $ .sort 0) | unreachable! 88 | instantiateMVars type 89 | let .success state1 _ ← state.tryTacticM goal (Tactic.assignWithAuxLemma type) | unreachable! 90 | let parentExpr := state1.parentExpr! 91 | checkTrue "src has aux lemma" $ parentExpr.getUsedConstants.any isAuxLemma 92 | goalStatePickle state1 statePath 93 | let ((), _) ← runCoreM coreDst $ transformTestT runTermElabMInCore do 94 | let (state1, _) ← goalStateUnpickle statePath (← getEnv) 95 | let parentExpr := state1.parentExpr! 96 | checkTrue "dst has aux lemma" $ parentExpr.getUsedConstants.any isAuxLemma 97 | 98 | return () 99 | 100 | /-- Synthetic mvars in this case creates closures. These cannot be pickled. -/ 101 | def test_pickling_synthetic_mvars : TestM Unit := do 102 | let coreSrc : Core.State := { env := ← getEnv } 103 | IO.FS.withTempFile λ _ statePath => do 104 | let stateGenerate : MetaM GoalState := runTermElabMInMeta do 105 | let type ← Elab.Term.elabTerm (← `(term|(0 : Nat) < 1)) .none 106 | let state ← GoalState.create type 107 | let .success state _ ← state.tryHave .unfocus `h "0 < 2" | unreachable! 108 | assert! state.savedState.term.elab.syntheticMVars.size > 0 109 | return state 110 | 111 | let ((), _) ← runCoreM coreSrc do 112 | let state ← stateGenerate.run' 113 | goalStatePickle state statePath 114 | 115 | structure Test where 116 | name : String 117 | routine: TestM Unit 118 | 119 | protected def Test.run (test: Test) (env: Environment) : IO LSpec.TestSeq := do 120 | -- Create the state 121 | let state : MultiState := { 122 | coreContext := ← createCoreContext #[], 123 | env, 124 | } 125 | match ← ((runTest $ test.routine).run' state).toBaseIO with 126 | | .ok e => return e 127 | | .error e => 128 | return LSpec.check s!"Emitted exception: {e.toString}" (e.toString == "") 129 | 130 | def suite (env : Environment): List (String × IO LSpec.TestSeq) := 131 | let tests: List Test := [ 132 | { name := "environment", routine := test_pickling_environment, }, 133 | { name := "goal simple", routine := test_goal_state_simple, }, 134 | { name := "goal synthetic mvars", routine := test_pickling_synthetic_mvars, }, 135 | { name := "extensions", routine := test_pickling_env_extensions, }, 136 | ] 137 | tests.map (fun test => (test.name, test.run env)) 138 | 139 | end Pantograph.Test.Serial 140 | -------------------------------------------------------------------------------- /doc/icon.svg: -------------------------------------------------------------------------------- 1 | 2 | 3 | 4 | 16 | 36 | 43 | 50 | 57 | 64 | 71 | 78 | 85 | 92 | 97 | 104 | 105 | 107 | 111 | 115 | 119 | 120 | 124 | 133 | 142 | 147 | 151 | 156 | 157 | 158 | -------------------------------------------------------------------------------- /Pantograph/Serial.lean: -------------------------------------------------------------------------------- 1 | import Pantograph.Goal 2 | import Pantograph.Environment 3 | 4 | import Lean.Environment 5 | import Lean.Replay 6 | import Lean.Util.ShareCommon 7 | import Std.Data.HashMap 8 | 9 | open Lean 10 | 11 | /-! 12 | Input/Output functions borrowed from REPL 13 | 14 | # Pickling and unpickling objects 15 | 16 | By abusing `saveModuleData` and `readModuleData` we can pickle and unpickle objects to disk. 17 | -/ 18 | 19 | namespace Pantograph 20 | 21 | /-- 22 | Save an object to disk. 23 | If you need to write multiple objects from within a single declaration, 24 | you will need to provide a unique `key` for each. 25 | -/ 26 | def pickle {α : Type} (path : System.FilePath) (x : α) (key : Name := by exact decl_name%) : IO Unit := 27 | saveModuleData path key (unsafe unsafeCast x) 28 | 29 | /-- 30 | Load an object from disk. 31 | Note: The returned `CompactedRegion` can be used to free the memory behind the value 32 | of type `α`, using `CompactedRegion.free` (which is only safe once all references to the `α` are 33 | released). Ignoring the `CompactedRegion` results in the data being leaked. 34 | Use `withUnpickle` to call `CompactedRegion.free` automatically. 35 | 36 | This function is unsafe because the data being loaded may not actually have type `α`, and this 37 | may cause crashes or other bad behavior. 38 | -/ 39 | unsafe def unpickle (α : Type) (path : System.FilePath) : IO (α × CompactedRegion) := do 40 | let (x, region) ← readModuleData path 41 | pure (unsafeCast x, region) 42 | 43 | /-- Load an object from disk and run some continuation on it, freeing memory afterwards. -/ 44 | unsafe def withUnpickle [Monad m] [MonadLiftT IO m] {α β : Type} 45 | (path : System.FilePath) (f : α → m β) : m β := do 46 | let (x, region) ← unpickle α path 47 | let r ← f x 48 | region.free 49 | pure r 50 | 51 | /-- 52 | Pickle an `Environment` to disk relative to a background. 53 | 54 | We only store: 55 | * the list of imports 56 | * the new constants from `Environment.constants` 57 | and when unpickling, we build a fresh `Environment` from the imports, 58 | and then add the new constants. 59 | -/ 60 | @[export pantograph_env_pickle_m] 61 | def environmentPickle (env : Environment) (path : System.FilePath) (background? : Option Environment := .none) 62 | : IO Unit := 63 | let distilled := distilEnvironment env background? 64 | pickle path (ShareCommon.shareCommon distilled) 65 | 66 | def resurrectEnvironment 67 | (distilled : DistilledEnvironment) 68 | (background? : Option Environment := .none) 69 | : IO Environment := do 70 | let (imports, constArray) := distilled 71 | let env ← match background? with 72 | | .none => importModules imports {} 0 (loadExts := true) 73 | | .some env => 74 | assert! env.imports == imports 75 | pure env 76 | env.replay (Std.HashMap.ofList constArray.toList) 77 | 78 | /-- 79 | Unpickle an `Environment` from disk. 80 | 81 | We construct a fresh `Environment` with the relevant imports, 82 | and then replace the new constants. 83 | 84 | The `CompactedRegion` must be free'd after using the environment. Otherwise 85 | there could be memory leaks. 86 | -/ 87 | @[export pantograph_env_unpickle_m] 88 | def environmentUnpickle (path : System.FilePath) (background? : Option Environment := .none) 89 | : IO (Environment × CompactedRegion) := unsafe do 90 | let (distilled, region) ← unpickle (Array Import × ConstArray) path 91 | return (← resurrectEnvironment distilled background?, region) 92 | 93 | /-- `CoreM`'s state, with information irrelevant to pickling masked out -/ 94 | structure CompactCoreState where 95 | -- env : Environment 96 | nextMacroScope : MacroScope := firstFrontendMacroScope + 1 97 | ngen : NameGenerator := {} 98 | auxDeclNGen : DeclNameGenerator := {} 99 | -- traceState : TraceState := {} 100 | -- cache : Cache := {} 101 | -- messages : MessageLog := {} 102 | -- infoState : Elab.InfoState := {} 103 | 104 | structure CompactGoalState where 105 | env : DistilledEnvironment 106 | 107 | core : CompactCoreState 108 | «meta» : Meta.State 109 | «elab» : Elab.Term.State 110 | tactic : Elab.Tactic.State 111 | root : MVarId 112 | parentMVars : List MVarId 113 | fragments : FragmentMap 114 | 115 | /-- Pickles a goal state by taking its diff relative to a background 116 | environment. This function eliminates all `MessageData` from synthetic 117 | metavariables, because these `MessageData` objects frequently carry closures, 118 | which cannot be pickled. If there is no synthetic metavariable, this would not 119 | cause a difference. -/ 120 | @[export pantograph_goal_state_pickle_m] 121 | def goalStatePickle (goalState : GoalState) (path : System.FilePath) (background? : Option Environment := .none) 122 | : IO Unit := 123 | let { 124 | savedState := { 125 | term := { 126 | «meta» := { 127 | core := { 128 | env, nextMacroScope, ngen, auxDeclNGen, .. 129 | }, 130 | «meta», 131 | } 132 | «elab» := «elab»@{ 133 | syntheticMVars, .. 134 | }, 135 | }, 136 | tactic 137 | } 138 | root, 139 | parentMVars, 140 | fragments, 141 | } := goalState 142 | -- Delete `MessageData`s 143 | let syntheticMVars : MVarIdMap _ := syntheticMVars.foldl (init := .empty) λ acc key val => 144 | let kind := match val.kind with 145 | | .typeClass _ => .typeClass .none 146 | | .coe header? expectedType e f? _ => .coe header? expectedType e f? .none 147 | | k => k 148 | acc.insert key { val with kind } 149 | let compacted : CompactGoalState := { 150 | env := distilEnvironment env background?, 151 | 152 | core := { nextMacroScope, ngen, auxDeclNGen }, 153 | «meta», 154 | «elab» := { «elab» with syntheticMVars }, 155 | tactic, 156 | 157 | root, 158 | parentMVars, 159 | fragments, 160 | } 161 | pickle path (ShareCommon.shareCommon compacted) 162 | 163 | /-- 164 | Loads a goal state from disk. The background is the environment which the goal 165 | state was cached relative to. 166 | 167 | The `CompactedRegion` must be free'd after using the goal state. Otherwise there 168 | will be memory leaks. 169 | -/ 170 | @[export pantograph_goal_state_unpickle_m] 171 | def goalStateUnpickle (path : System.FilePath) (background? : Option Environment := .none) 172 | : IO (GoalState × CompactedRegion) := unsafe do 173 | let ({ 174 | env, 175 | 176 | core, 177 | «meta», 178 | «elab», 179 | tactic, 180 | 181 | root, 182 | parentMVars, 183 | fragments, 184 | }, region) ← Pantograph.unpickle CompactGoalState path 185 | let env ← resurrectEnvironment env background? 186 | let goalState := { 187 | savedState := { 188 | term := { 189 | «meta» := { 190 | core := { 191 | core with 192 | passedHeartbeats := 0, 193 | env, 194 | }, 195 | «meta», 196 | }, 197 | «elab», 198 | }, 199 | tactic, 200 | }, 201 | root, 202 | parentMVars, 203 | fragments, 204 | } 205 | return (goalState, region) 206 | 207 | end Pantograph 208 | -------------------------------------------------------------------------------- /Pantograph/Frontend/MetaTranslate.lean: -------------------------------------------------------------------------------- 1 | import Lean.Meta 2 | import Std.Data.HashMap 3 | 4 | namespace Pantograph.Frontend 5 | 6 | open Lean 7 | 8 | namespace MetaTranslate 9 | 10 | structure Context where 11 | sourceMCtx : MetavarContext := {} 12 | sourceLCtx : LocalContext := {} 13 | 14 | abbrev FVarMap := Std.HashMap FVarId FVarId 15 | 16 | structure State where 17 | -- Stores mapping from old to new mvar/fvars 18 | mvarMap: Std.HashMap MVarId MVarId := {} 19 | fvarMap: Std.HashMap FVarId FVarId := {} 20 | 21 | /- 22 | Monadic state for translating a frozen meta state. The underlying `MetaM` 23 | operates in the "target" context and state. 24 | -/ 25 | abbrev MetaTranslateM := ReaderT Context StateRefT State MetaM 26 | 27 | def getSourceLCtx : MetaTranslateM LocalContext := do pure (← read).sourceLCtx 28 | def getSourceMCtx : MetaTranslateM MetavarContext := do pure (← read).sourceMCtx 29 | def addTranslatedFVar (src dst: FVarId) : MetaTranslateM Unit := do 30 | modifyGet λ state => ((), { state with fvarMap := state.fvarMap.insert src dst }) 31 | def addTranslatedMVar (src dst: MVarId) : MetaTranslateM Unit := do 32 | modifyGet λ state => ((), { state with mvarMap := state.mvarMap.insert src dst }) 33 | 34 | def saveFVarMap : MetaTranslateM FVarMap := do 35 | return (← get).fvarMap 36 | def restoreFVarMap (map: FVarMap) : MetaTranslateM Unit := do 37 | modifyGet λ state => ((), { state with fvarMap := map }) 38 | def resetFVarMap : MetaTranslateM Unit := do 39 | modifyGet λ state => ((), { state with fvarMap := {} }) 40 | 41 | mutual 42 | private partial def translateLevel (srcLevel: Level) : MetaTranslateM Level := do 43 | let sourceMCtx ← getSourceMCtx 44 | let (_, level) := instantiateLevelMVarsImp sourceMCtx srcLevel 45 | match level with 46 | | .zero => return .zero 47 | | .succ inner => do 48 | let inner' ← translateLevel inner 49 | return .succ inner' 50 | | .max l1 l2 => do 51 | let l1' ← translateLevel l1 52 | let l2' ← translateLevel l2 53 | return .max l1' l2' 54 | | .imax l1 l2 => do 55 | let l1' ← translateLevel l1 56 | let l2' ← translateLevel l2 57 | return .imax l1' l2' 58 | | .param p => return .param p 59 | | .mvar _ => 60 | Meta.mkFreshLevelMVar 61 | private partial def translateExpr (srcExpr: Expr) : MetaTranslateM Expr := do 62 | let sourceMCtx ← getSourceMCtx 63 | -- We want to create as few mvars as possible 64 | let (srcExpr, _) := instantiateMVarsCore (mctx := sourceMCtx) srcExpr 65 | trace[Pantograph.Frontend.MetaTranslate] "Transform src: {srcExpr}" 66 | let result ← Core.transform srcExpr λ e => do 67 | let state ← get 68 | match e with 69 | | .fvar fvarId => 70 | let .some fvarId' := state.fvarMap[fvarId]? | panic! s!"FVar id not registered: {fvarId.name}" 71 | -- Delegating this to `Meta.check` later 72 | --assert! (← getLCtx).contains fvarId' 73 | return .done $ .fvar fvarId' 74 | | .mvar mvarId => do 75 | -- Must not be assigned 76 | assert! !(sourceMCtx.eAssignment.contains mvarId) 77 | match state.mvarMap[mvarId]? with 78 | | .some mvarId' => do 79 | return .done $ .mvar mvarId' 80 | | .none => do 81 | -- Entering another LCtx, must save the current one 82 | let fvarMap ← saveFVarMap 83 | let mvarId' ← translateMVarId mvarId 84 | restoreFVarMap fvarMap 85 | return .done $ .mvar mvarId' 86 | | .sort level => do 87 | let level' ← translateLevel level 88 | return .done $ .sort level' 89 | | _ => return .continue 90 | Meta.check result 91 | return result 92 | 93 | partial def translateLocalInstance (srcInstance: LocalInstance) : MetaTranslateM LocalInstance := do 94 | return { 95 | className := srcInstance.className, 96 | fvar := ← translateExpr srcInstance.fvar 97 | } 98 | partial def translateLocalDecl (srcLocalDecl: LocalDecl) : MetaTranslateM LocalDecl := do 99 | let fvarId ← mkFreshFVarId 100 | addTranslatedFVar srcLocalDecl.fvarId fvarId 101 | match srcLocalDecl with 102 | | .cdecl index _ userName type bi kind => do 103 | trace[Pantograph.Frontend.MetaTranslate] "[CD] {userName} {toString type}" 104 | return .cdecl index fvarId userName (← translateExpr type) bi kind 105 | | .ldecl index _ userName type value nonDep kind => do 106 | trace[Pantograph.Frontend.MetaTranslate] "[LD] {toString type} := {toString value}" 107 | return .ldecl index fvarId userName (← translateExpr type) (← translateExpr value) nonDep kind 108 | 109 | partial def translateLCtx : MetaTranslateM LocalContext := do 110 | resetFVarMap 111 | let lctx ← MonadLCtx.getLCtx 112 | assert! lctx.isEmpty 113 | (← getSourceLCtx).foldlM (λ lctx srcLocalDecl => do 114 | -- Prevents circular proofs and panics 115 | if srcLocalDecl.isAuxDecl && !(lctx.auxDeclToFullName.contains srcLocalDecl.fvarId) then 116 | pure lctx 117 | else 118 | let localDecl ← Meta.withLCtx' lctx do 119 | translateLocalDecl srcLocalDecl 120 | pure $ lctx.addDecl localDecl 121 | ) lctx 122 | 123 | partial def translateMVarId (srcMVarId: MVarId) : MetaTranslateM MVarId := do 124 | if let .some mvarId' := (← get).mvarMap[srcMVarId]? then 125 | trace[Pantograph.Frontend.MetaTranslate] "Existing mvar id {srcMVarId.name} → {mvarId'.name}" 126 | return mvarId' 127 | let mvarId' ← Meta.withLCtx .empty #[] do 128 | let srcDecl := (← getSourceMCtx).findDecl? srcMVarId |>.get! 129 | withTheReader Context (λ ctx => { ctx with sourceLCtx := srcDecl.lctx }) do 130 | let lctx' ← translateLCtx 131 | let localInstances' ← srcDecl.localInstances.mapM translateLocalInstance 132 | Meta.withLCtx lctx' localInstances' do 133 | let target' ← translateExpr srcDecl.type 134 | let mvar' ← Meta.mkFreshExprMVar target' srcDecl.kind srcDecl.userName 135 | let mvarId' := mvar'.mvarId! 136 | if let .some { fvars, mvarIdPending }:= (← getSourceMCtx).getDelayedMVarAssignmentExp srcMVarId then 137 | -- Map the fvars in the pending context. 138 | let mvarIdPending' ← translateMVarId mvarIdPending 139 | let fvars' ← mvarIdPending'.withContext $ fvars.mapM translateExpr 140 | assignDelayedMVar mvarId' fvars' mvarIdPending' 141 | pure mvarId' 142 | trace[Pantograph.Frontend.MetaTranslate] "Translated {srcMVarId.name} → {mvarId'.name}" 143 | addTranslatedMVar srcMVarId mvarId' 144 | return mvarId' 145 | end 146 | 147 | def translateMVarFromTermInfo (termInfo : Elab.TermInfo) (context? : Option Elab.ContextInfo) 148 | : MetaTranslateM MVarId := do 149 | withTheReader Context (λ ctx => { ctx with 150 | sourceMCtx := context?.map (·.mctx) |>.getD {}, 151 | sourceLCtx := termInfo.lctx, 152 | }) do 153 | let type := termInfo.expectedType?.get! 154 | let lctx' ← translateLCtx 155 | let mvar ← Meta.withLCtx lctx' #[] do 156 | Meta.withLocalInstances (lctx'.decls.toList.filterMap id) do 157 | let type' ← translateExpr type 158 | trace[Pantograph.Frontend.MetaTranslate] "Translating from term info {← Meta.ppExpr type'}" 159 | Meta.mkFreshExprSyntheticOpaqueMVar type' 160 | return mvar.mvarId! 161 | 162 | def translateMVarFromTacticInfoBefore (tacticInfo : Elab.TacticInfo) (_context? : Option Elab.ContextInfo) 163 | : MetaTranslateM (List MVarId) := do 164 | withTheReader Context (λ ctx => { ctx with sourceMCtx := tacticInfo.mctxBefore }) do 165 | tacticInfo.goalsBefore.mapM translateMVarId 166 | 167 | end MetaTranslate 168 | 169 | export MetaTranslate (MetaTranslateM) 170 | 171 | initialize 172 | registerTraceClass `Pantograph.Frontend.MetaTranslate 173 | -------------------------------------------------------------------------------- /Pantograph/Frontend/Basic.lean: -------------------------------------------------------------------------------- 1 | import Lean.Parser 2 | import Lean.Elab.Frontend 3 | 4 | open Lean 5 | 6 | namespace Lean.FileMap 7 | 8 | /-- Extract the range of a `Syntax` expressed as lines and columns. -/ 9 | @[export pantograph_frontend_stx_range] 10 | protected def stxRange (fileMap : FileMap) (stx : Syntax) : Position × Position := 11 | let pos := stx.getPos?.getD 0 12 | let endPos := stx.getTailPos?.getD pos 13 | (fileMap.toPosition pos, fileMap.toPosition endPos) 14 | 15 | end Lean.FileMap 16 | namespace Lean.PersistentArray 17 | 18 | /-- 19 | Drop the first `n` elements of a `PersistentArray`, returning the results as a `List`. 20 | 21 | We can't remove the `[Inhabited α]` hypotheses here until 22 | `PersistentArray`'s `GetElem` instance also does. 23 | -/ 24 | protected def drop [Inhabited α] (t : PersistentArray α) (n : Nat) : List α := 25 | List.range (t.size - n) |>.map fun i => t.get! (n + i) 26 | 27 | end Lean.PersistentArray 28 | 29 | namespace Pantograph.Frontend 30 | 31 | @[export pantograph_frontend_stx_byte_range] 32 | def stxByteRange (stx : Syntax) : String.Pos.Raw × String.Pos.Raw := 33 | let pos := stx.getPos?.getD 0 34 | let endPos := stx.getTailPos?.getD 0 35 | (pos, endPos) 36 | 37 | structure Context where 38 | cancelTk? : Option IO.CancelToken := .none 39 | 40 | /-- This `FrontendM` comes with more options. -/ 41 | abbrev FrontendM := ReaderT Context Elab.Frontend.FrontendM 42 | 43 | structure CompilationStep where 44 | scope : Elab.Command.Scope 45 | fileName : String 46 | fileMap : FileMap 47 | src : Substring.Raw 48 | stx : Syntax 49 | before : Environment 50 | after : Environment 51 | msgs : List Message 52 | trees : List Elab.InfoTree 53 | 54 | @[export pantograph_frontend_compilation_step_defined_constants_m] 55 | protected def CompilationStep.newConstants (step : CompilationStep) : IO NameSet := do 56 | step.after.constants.map₂.foldlM (init := .empty) λ acc name _ => do 57 | if step.before.contains name then 58 | return acc 59 | let coreM : CoreM Bool := Option.isSome <$> findDeclarationRanges? name 60 | let hasRange ← coreM.run' 61 | { fileName := step.fileName, fileMap := step.fileMap } 62 | { env := step.after } 63 | |>.toBaseIO 64 | match hasRange with 65 | | .ok true => return acc.insert name 66 | | .ok false => return acc 67 | | .error e => throw $ IO.userError (← e.toMessageData.toString) 68 | 69 | /-- Like `Elab.Frontend.runCommandElabM`, but taking `cancelTk?` into account. -/ 70 | @[inline] def runCommandElabM (x : Elab.Command.CommandElabM α) : FrontendM α := do 71 | let config ← read 72 | let ctx ← readThe Elab.Frontend.Context 73 | let s ← get 74 | let cmdCtx : Elab.Command.Context := { 75 | cmdPos := s.cmdPos 76 | fileName := ctx.inputCtx.fileName 77 | fileMap := ctx.inputCtx.fileMap 78 | snap? := none 79 | cancelTk? := config.cancelTk? 80 | } 81 | match (← liftM <| EIO.toIO' <| (x cmdCtx).run s.commandState) with 82 | | Except.error e => throw <| IO.Error.userError s!"unexpected internal error: {← e.toMessageData.toString}" 83 | | Except.ok (a, sNew) => Elab.Frontend.setCommandState sNew; return a 84 | 85 | def elabCommandAtFrontend (stx : Syntax) : FrontendM Unit := do 86 | runCommandElabM do 87 | let initMsgs ← modifyGet λ st => 88 | (st.messages, { st with messages := {} }) 89 | Elab.Command.elabCommandTopLevel stx 90 | modify λ state => { state with 91 | messages := initMsgs ++ state.messages } 92 | 93 | open Elab.Frontend in 94 | def processCommand : FrontendM Bool := do 95 | updateCmdPos 96 | let cmdState ← getCommandState 97 | let ictx ← getInputContext 98 | let pstate ← getParserState 99 | let scope := cmdState.scopes.head! 100 | let pmctx := { env := cmdState.env, options := scope.opts, currNamespace := scope.currNamespace, openDecls := scope.openDecls } 101 | match profileit "parsing" scope.opts fun _ => Parser.parseCommand ictx pmctx pstate cmdState.messages with 102 | | (cmd, ps, messages) => 103 | modify fun s => { s with commands := s.commands.push cmd } 104 | setParserState ps 105 | setMessages messages 106 | elabCommandAtFrontend cmd 107 | pure (Parser.isTerminalCommand cmd) 108 | 109 | /-- 110 | Process one command, returning a `CompilationStep` and 111 | `done : Bool`, indicating whether this was the last command. 112 | -/ 113 | @[export pantograph_frontend_process_one_command_m] 114 | def processOneCommand: FrontendM (CompilationStep × Bool) := do 115 | let s := (← get).commandState 116 | let before := s.env 117 | let done ← processCommand 118 | let stx := (← get).commands.back! 119 | let src := (← readThe Elab.Frontend.Context).inputCtx.substring 120 | (← get).cmdPos 121 | (← get).parserState.pos 122 | let s' := (← get).commandState 123 | let after := s'.env 124 | let msgs := s'.messages.toList.drop s.messages.toList.length 125 | let trees := s'.infoState.trees.drop s.infoState.trees.size 126 | let { fileName, fileMap, .. } := (← readThe Elab.Frontend.Context).inputCtx 127 | return ({ scope := s.scopes.head!, fileName, fileMap, src, stx, before, after, msgs, trees }, done) 128 | 129 | /-- Executes a `FrontendM`-based monad until completion -/ 130 | partial def executeFrontend { m } [Monad m] [MonadLiftT FrontendM m] 131 | (f : CompilationStep → m Unit) : m Unit := do 132 | let (cmd, done) ← processOneCommand 133 | if done then 134 | if cmd.src.str.isEmpty then 135 | return () 136 | else 137 | f cmd 138 | else 139 | f cmd 140 | executeFrontend f 141 | 142 | def mapCompilationSteps { m α } 143 | [Monad m] [MonadLiftT FrontendM m] [MonadLiftT (ST IO.RealWorld) m] 144 | (f : CompilationStep → m α) : m (List α) := do 145 | let f' (step : CompilationStep) : StateRefT' IO.RealWorld (List α) m Unit := do 146 | let a ← f step 147 | modify (a :: ·) 148 | let (_, li) ← executeFrontend f' |>.run [] 149 | return li.reverse 150 | 151 | @[export pantograph_frontend_find_source_path_m] 152 | def findSourcePath (module : Name) : IO System.FilePath := do 153 | let olean ← findOLean module 154 | return System.FilePath.mk (olean.toString.replace ".lake/build/lib/" "") |>.withExtension "lean" 155 | 156 | def defaultFileName := "" 157 | 158 | /-- 159 | Use with 160 | ```lean 161 | let m: FrontendM α := ... 162 | let (context, state) ← createContextStateFromFile ... 163 | m.run context |>.run' state 164 | ``` 165 | -/ 166 | @[export pantograph_frontend_create_context_state_from_file_m] 167 | def createContextStateFromFile 168 | (file : String) -- Content of the file 169 | (fileName : String := defaultFileName) 170 | (env? : Option Lean.Environment := .none) -- If set to true, assume there's no header. 171 | (opts : Options := {}) 172 | : IO (Elab.Frontend.Context × Elab.Frontend.State) := unsafe do 173 | --let file ← IO.FS.readFile (← findSourcePath module) 174 | let inputCtx := Parser.mkInputContext file fileName 175 | 176 | let (header, parserState, messages) ← Parser.parseHeader inputCtx 177 | let (env, parserState, messages) ← match env? with 178 | | .some env => pure (env, {}, .empty) 179 | | .none => 180 | -- Only process the header if we don't have an environment. 181 | let (env, messages) ← Elab.processHeader header opts messages inputCtx 182 | pure (env, parserState, messages) 183 | let commandState := Elab.Command.mkState env messages opts 184 | let context: Elab.Frontend.Context := { inputCtx } 185 | let state: Elab.Frontend.State := { 186 | commandState := { commandState with infoState.enabled := true }, 187 | parserState, 188 | cmdPos := parserState.pos 189 | } 190 | return (context, state) 191 | 192 | /-- Returns the command state at the end of execution -/ 193 | def collectEndState : FrontendM Elab.Command.State := do 194 | executeFrontend λ _ => pure () 195 | let state ← get 196 | return state.commandState 197 | -------------------------------------------------------------------------------- /Pantograph/Frontend/InfoTree.lean: -------------------------------------------------------------------------------- 1 | /- Adapted from lean-training-data -/ 2 | import Lean.Elab.InfoTree 3 | import Lean.Elab.Util 4 | import Lean.Parser.Term 5 | import Lean.PrettyPrinter 6 | 7 | open Lean 8 | 9 | namespace Lean.Elab 10 | 11 | private def elaboratorToString : Name → String 12 | | .anonymous => "" 13 | | n => s!"⟨{n}⟩ " 14 | 15 | /-- The `Syntax` for a `Lean.Elab.Info`, if there is one. -/ 16 | protected def Info.stx? : Info → Option Syntax 17 | | .ofTacticInfo info => info.stx 18 | | .ofTermInfo info => info.stx 19 | | .ofCommandInfo info => info.stx 20 | | .ofMacroExpansionInfo info => info.stx 21 | | .ofOptionInfo info => info.stx 22 | | .ofFieldInfo info => info.stx 23 | | .ofCompletionInfo info => info.stx 24 | | .ofUserWidgetInfo info => info.stx 25 | | .ofCustomInfo info => info.stx 26 | | .ofFVarAliasInfo _ => none 27 | | .ofDocInfo info => info.stx 28 | | .ofDocElabInfo info => info.stx 29 | | .ofFieldRedeclInfo info => info.stx 30 | | .ofChoiceInfo info => info.stx 31 | | .ofPartialTermInfo info => info.stx 32 | | .ofDelabTermInfo info => info.stx 33 | | .ofErrorNameInfo info => info.stx 34 | /-- Is the `Syntax` for this `Lean.Elab.Info` original, or synthetic? -/ 35 | protected def Info.isOriginal (i : Info) : Bool := 36 | match i.stx? with 37 | | none => true -- Somewhat unclear what to do with `FVarAliasInfo`, so be conservative. 38 | | some stx => match stx.getHeadInfo with 39 | | .original .. => true 40 | | _ => false 41 | 42 | def ContextInfo.ppExpr (ctx : ContextInfo) (lctx : LocalContext) (e : Expr) : IO Format := 43 | ctx.runMetaM lctx (do Meta.ppExpr (← instantiateMVars e)) 44 | 45 | def CommandInfo.toString (info : CommandInfo) (ctx : ContextInfo) : IO String := do 46 | let stx := (← ctx.ppSyntax {} info.stx).pretty 47 | return s!"{elaboratorToString info.elaborator}\n{stx}" 48 | 49 | def TermInfo.toString (info : TermInfo) (ctx : ContextInfo) : IO String := do 50 | let stx := (← ctx.ppSyntax info.lctx info.stx).pretty 51 | let expectedType := (← info.expectedType?.mapM fun ty => do 52 | pure s!": {(← ctx.ppExpr info.lctx ty).pretty}").getD "" 53 | let expr := (← ctx.ppExpr info.lctx info.expr).pretty 54 | return s!"{elaboratorToString info.elaborator}{expr}{expectedType}\n{stx}" 55 | 56 | /-- Find the name for the outermost `Syntax` in this `TacticInfo`. -/ 57 | def TacticInfo.name? (t : TacticInfo) : Option Name := 58 | match t.stx with 59 | | Syntax.node _ n _ => some n 60 | | _ => none 61 | /-- Decide whether a tactic is "substantive", 62 | or is merely a tactic combinator (e.g. `by`, `;`, multiline tactics, parenthesized tactics). -/ 63 | def TacticInfo.isSubstantive (t : TacticInfo) : Bool := 64 | match t.name? with 65 | | none => false 66 | | some `null => false 67 | | some ``cdot => false 68 | | some ``cdotTk => false 69 | | some ``Lean.Parser.Term.byTactic => false 70 | | some ``Lean.Parser.Tactic.tacticSeq => false 71 | | some ``Lean.Parser.Tactic.tacticSeq1Indented => false 72 | | some ``Lean.Parser.Tactic.«tactic_<;>_» => false 73 | | some ``Lean.Parser.Tactic.paren => false 74 | | _ => true 75 | def TacticInfo.pp (info : TacticInfo) (ctx : ContextInfo) : IO Format := 76 | ctx.runMetaM {} try 77 | Lean.PrettyPrinter.ppTactic ⟨info.stx⟩ 78 | catch _ => 79 | pure "" 80 | def TacticInfo.toString (i : TacticInfo) (ctx : ContextInfo) : IO String := do 81 | let name := i.name? 82 | let stx := Format.pretty (← i.pp ctx) 83 | return s!"{name}\n{stx}" 84 | 85 | /-- 86 | Keep `.node` nodes and `.hole` nodes satisfying predicates. 87 | 88 | Returns a `List InfoTree`, although in most situations this will be a singleton. 89 | -/ 90 | partial def InfoTree.filter (p : Info → Bool) (m : MVarId → Bool := fun _ => false) : 91 | InfoTree → List InfoTree 92 | | .context ctx tree => tree.filter p m |>.map (.context ctx) 93 | | .node info children => 94 | if p info then 95 | [.node info (children.toList.map (filter p m)).flatten.toPArray'] 96 | else 97 | (children.toList.map (filter p m)).flatten 98 | | .hole mvar => if m mvar then [.hole mvar] else [] 99 | 100 | /-- Analogue of `Lean.Elab.InfoTree.findInfo?`, but that returns a list of all results. -/ 101 | partial def InfoTree.findAllInfo 102 | (t : InfoTree) 103 | (context?: Option Elab.ContextInfo) 104 | (haltOnMatch : Bool := false) 105 | (pred : Elab.Info → Bool) 106 | : List (Elab.Info × Option Elab.ContextInfo × PersistentArray Elab.InfoTree) := 107 | match t with 108 | | .context inner t => findAllInfo t (inner.mergeIntoOuter? context?) haltOnMatch pred 109 | | .node i children => 110 | let head := if pred i then [(i, context?, children)] else [] 111 | let tail := if haltOnMatch ∧ !head.isEmpty then [] else children.toList.flatMap (fun t => findAllInfo t context? haltOnMatch pred) 112 | head ++ tail 113 | | _ => [] 114 | 115 | /-- Monadic analogue of `findAllInfo`, but predicate controls whether to recurse. -/ 116 | partial def InfoTree.findAllInfoM [Monad m] 117 | (t : InfoTree) 118 | (context?: Option Elab.ContextInfo) 119 | (pred : Elab.Info → Option Elab.ContextInfo → m (Bool × Bool)) 120 | : m (List (Elab.Info × Option Elab.ContextInfo × PersistentArray Elab.InfoTree)) := do 121 | match t with 122 | | .context inner t => t.findAllInfoM (inner.mergeIntoOuter? context?) pred 123 | | .node i children => 124 | let (flagCollect, flagRecurse) ← pred i context? 125 | let head := if flagCollect then [(i, context?, children)] else [] 126 | let tail := if ¬ flagRecurse then pure [] else children.toList.mapM (fun t => t.findAllInfoM context? pred) 127 | return head ++ (← tail).flatten 128 | | _ => return [] 129 | 130 | structure InfoTreePrintOptions where 131 | prettyPrint : Bool := true 132 | 133 | @[export pantograph_infotree_to_string_m] 134 | partial def InfoTree.toString (t : InfoTree) (ctx?: Option Elab.ContextInfo := .none) (options : InfoTreePrintOptions := {}) (indent : Nat := 0) 135 | : IO String := do 136 | let space := String.join $ List.replicate indent " " 137 | match t with 138 | | .context ctx t => 139 | let s ← t.toString (ctx.mergeIntoOuter? ctx?) 140 | pure s!"[context] {s}" 141 | | .node info children => 142 | if let some ctx := ctx? then 143 | let node : String ← match info with 144 | | .ofTermInfo { stx, expr, .. } => 145 | pure s!"[term] {stx.prettyPrint} ➡ {expr}" 146 | | .ofCommandInfo info => 147 | let head := match info.stx.getArg 1 with 148 | | .missing => "" 149 | | other => s!" {other.getKind.toString}" 150 | let s := if options.prettyPrint then 151 | info.stx.prettyPrint 152 | else 153 | s!"{info.stx}" 154 | pure s!"[command/{info.stx.getKind.toString}{head}] {s}" 155 | | .ofTacticInfo info => pure s!"[tactic] {info.stx.prettyPrint}" 156 | | .ofMacroExpansionInfo _ => pure "[macro_exp]" 157 | | .ofOptionInfo info => pure s!"[option] {info.stx.prettyPrint}" 158 | | .ofFieldInfo _ => pure "[field]" 159 | | .ofCompletionInfo info => pure s!"[completion] {info.stx.prettyPrint}" 160 | | .ofUserWidgetInfo _ => pure "[user_widget]" 161 | | .ofCustomInfo _ => pure "[custom]" 162 | | .ofFVarAliasInfo _ => pure "[fvar_alias]" 163 | | .ofFieldRedeclInfo _ => pure "[field_redecl]" 164 | | .ofDocElabInfo _ => pure "[doc_elab]" 165 | | .ofDocInfo _ => pure "[doc]" 166 | | .ofChoiceInfo _ => pure "[choice]" 167 | | .ofPartialTermInfo _ => pure "[partial_term]" 168 | | .ofDelabTermInfo _ => pure "[delab_term]" 169 | | .ofErrorNameInfo _ => pure "[error_name]" 170 | let children ← children.toList.mapM λ t' => do 171 | t'.toString ctx options (indent := indent + 1) 172 | let children := String.join children 173 | return s!"{space}{node}\n{children}" 174 | else throw <| IO.userError "No `ContextInfo` available." 175 | | .hole mvarId => 176 | if let some ctx := ctx? then 177 | let payload := (← ctx.runMetaM {} (do Meta.ppGoal mvarId)).pretty 178 | return s!"{space}[hole] {payload}" 179 | else throw <| IO.userError "No `ContextInfo` available." 180 | 181 | end Lean.Elab 182 | -------------------------------------------------------------------------------- /Test/Frontend/Collect.lean: -------------------------------------------------------------------------------- 1 | import Pantograph.Frontend 2 | import Pantograph.Environment 3 | import Test.Common 4 | 5 | open Lean Pantograph Frontend 6 | 7 | namespace Pantograph.Test.Frontend.Collect 8 | 9 | abbrev Test := Environment → TestT IO Unit 10 | 11 | def runFrontend { α } (env : Environment) (source: String) (f : CompilationStep → FrontendM α) (timeout : UInt32 := 0) 12 | : IO (List α) := do 13 | let (context, state) ← do createContextStateFromFile source (env? := env) 14 | let m := mapCompilationSteps f 15 | let cancelTk? ← match timeout with 16 | | 0 => pure .none 17 | | timeout => .some <$> spawnCancelToken timeout 18 | m.run { cancelTk? } |>.run context |>.run' state 19 | 20 | def test_open : Test := λ env => do 21 | let sketch := " 22 | open Nat 23 | example : ∀ (n : Nat), n + 1 = Nat.succ n := by 24 | intro 25 | apply add_one 26 | " 27 | let errors ← runFrontend env sketch λ step => step.msgs.mapM (·.toString) 28 | checkEq "errors" errors [[], []] 29 | 30 | def collectNewConstants (env : Environment) (source: String) : IO (List (List Name)) := do 31 | let (context, state) ← do Frontend.createContextStateFromFile source (env? := env) 32 | let m := show FrontendM _ from Frontend.mapCompilationSteps λ step => do 33 | step.newConstants 34 | let result ← m.run {} |>.run context |>.run' state 35 | return result.map (·.toList) 36 | 37 | def test_collect_one_constant : Test := λ env => do 38 | let input := " 39 | def mystery : Nat := 123 40 | " 41 | let names ← collectNewConstants env input 42 | checkEq "constants" names [[`mystery]] 43 | 44 | def test_collect_one_theorem : Test := λ env => do 45 | let input := " 46 | theorem mystery [SizeOf α] (as : List α) (i : Fin as.length) : sizeOf (as.get i) < sizeOf as := by 47 | match as, i with 48 | | a::as, ⟨0, _⟩ => simp_arith [get] 49 | | a::as, ⟨i+1, h⟩ => 50 | have ih := sizeOf_get as ⟨i, Nat.le_of_succ_le_succ h⟩ 51 | apply Nat.lt_trans ih 52 | simp_arith 53 | " 54 | let names ← collectNewConstants env input 55 | checkEq "constants" names [[`mystery]] 56 | 57 | def test_collect_stub : Test := λ env => do 58 | let input := " 59 | theorem mystery [SizeOf α] (as : List α) (i : Fin as.length) : sizeOf (as.get i) < sizeOf as := sorry 60 | " 61 | let names ← collectNewConstants env input 62 | checkEq "constants" names [[`mystery]] 63 | 64 | def checkFileConflicts (env : Environment) (src dst : String) : IO (Except String Environment):= do 65 | let srcState ← collectOne src 66 | let dstState ← collectOne dst 67 | ExceptT.run $ checkEnvConflicts env srcState.env dstState.env 68 | where 69 | collectOne (source : String) : IO _ := do 70 | let (context, state) ← do createContextStateFromFile source (env? := env) 71 | let m := collectEndState 72 | m.run { } |>.run context |>.run' state 73 | 74 | def test_conflict_simple : Test := λ env => do 75 | let src := " 76 | def x : Nat := sorry 77 | " 78 | let dst := " 79 | def x : Nat := 123 80 | " 81 | let result? ← checkFileConflicts env src dst 82 | match result? with 83 | | .ok _ => checkTrue "ok" result?.isOk 84 | | .error e => fail s!"Failed with {e}" 85 | def test_conflict_poly : Test := λ env => do 86 | let src := " 87 | def mystery : List α → List α := sorry 88 | " 89 | let dst := " 90 | def helper (li : List β) := li.reverse 91 | def mystery (li : List α) := (helper li) ++ li 92 | " 93 | let result? ← checkFileConflicts env src dst 94 | match result? with 95 | | .ok _ => checkTrue "ok" result?.isOk 96 | | .error e => fail s!"Failed with {e}" 97 | 98 | def test_conflict_auxiliary : Test := λ env => do 99 | let src := " 100 | def f : Nat → Nat := sorry 101 | " 102 | let dst := " 103 | def x : Nat := 123 104 | def f : Nat → Nat := λ y => y + x 105 | " 106 | let result? ← checkFileConflicts env src dst 107 | match result? with 108 | | .ok _ => checkTrue "ok" result?.isOk 109 | | .error e => fail s!"Failed with {e}" 110 | def test_conflict_axiom : Test := λ env => do 111 | let src := " 112 | axiom α : Type 113 | axiom ne : Nonempty α 114 | noncomputable def f : α := sorry 115 | " 116 | let dst := " 117 | axiom α : Type 118 | axiom ne : Nonempty α 119 | noncomputable def f : α := @Classical.choice α ne 120 | " 121 | let result? ← checkFileConflicts env src dst 122 | match result? with 123 | | .ok _ => checkTrue "ok" result?.isOk 124 | | .error e => fail s!"Failed with {e}" 125 | 126 | /-- from `GasStationManager/SafeVerify` -/ 127 | def test_conflict_simple_def : Test := λ env => do 128 | let src := " 129 | def solveAdd (a b:Int):{c:Int//a+c=b} := sorry 130 | " 131 | let dst := " 132 | def solveAdd (a b:Int):{c:Int//a+c=b} := ⟨b-a, by omega⟩ 133 | " 134 | let result? ← checkFileConflicts env src dst 135 | match result? with 136 | | .ok _ => checkTrue "ok" result?.isOk 137 | | .error e => fail s!"Failed with {e}" 138 | /-- from `GasStationManager/SafeVerify` -/ 139 | def test_conflict_fake_implementation : Test := λ env => do 140 | let src := " 141 | noncomputable def definitely_at_least_two : Nat := sorry 142 | theorem definitely_at_least_two_spec : 2 ≤ definitely_at_least_two := sorry 143 | " 144 | let dst := " 145 | @[implemented_by Nat.zero] 146 | noncomputable def definitely_at_least_two : Nat := 147 | Exists.choose (⟨3, by simp⟩ : ∃ x, 2 ≤ x) 148 | 149 | theorem definitely_at_least_two_spec : 2 ≤ definitely_at_least_two := 150 | Exists.choose_spec _ 151 | " 152 | let result? ← checkFileConflicts env src dst 153 | match result? with 154 | | .ok _ => checkTrue "ok" result?.isOk 155 | | .error e => fail s!"Failed with {e}" 156 | 157 | def test_conflict_fail_idempotent : Test := λ env => do 158 | let src := " 159 | def x : Nat := sorry 160 | " 161 | let .error e ← checkFileConflicts env src src 162 | | fail "Must fail" 163 | checkEq "message" e "Definition value has sorry: x" 164 | def test_conflict_fail_delete_definition : Test := λ env => do 165 | let src := " 166 | def x : Nat := sorry 167 | def y : Nat := sorry 168 | " 169 | let dst := " 170 | def x : Nat := 123 171 | " 172 | let .error e ← checkFileConflicts env src dst 173 | | fail "Must fail" 174 | checkEq "message" e "[y] not accounted for" 175 | 176 | def test_conflict_fail_inductive_modification : Test := λ env => do 177 | let src := " 178 | inductive A where 179 | | a 180 | | b 181 | " 182 | let dst := " 183 | inductive A where 184 | | a 185 | | b 186 | | c 187 | " 188 | let .error e ← checkFileConflicts env src dst 189 | | fail "Must fail" 190 | checkEq "message" e "Type clash of A.casesOn" 191 | 192 | /-- from `GasStationManager/SafeVerify` -/ 193 | def test_conflict_fail_noncomputable : Test := λ env => do 194 | let src := " 195 | axiom α : Type 196 | axiom ne : Nonempty α 197 | def f : α := sorry 198 | " 199 | let dst := " 200 | axiom α : Type 201 | axiom ne : Nonempty α 202 | noncomputable def f : α := @Classical.choice α ne 203 | " 204 | let .error e ← checkFileConflicts env src dst 205 | | fail "Must fail" 206 | checkEq "message" e "Must not modify computability on f" 207 | 208 | def test_conflict_fail_add_axiom : Test := λ env => do 209 | let src := " 210 | theorem mystery : False := sorry 211 | " 212 | let dst := " 213 | axiom z : False 214 | theorem mystery : False := z 215 | " 216 | let .error e ← checkFileConflicts env src dst 217 | | fail "Must fail" 218 | checkEq "message" e "Adding axiom is not allowed: z" 219 | 220 | def suite (env : Environment): List (String × IO LSpec.TestSeq) := 221 | let tests := [ 222 | ("open", test_open), 223 | ("collect_one_constant", test_collect_one_constant), 224 | ("collect_one_theorem", test_collect_one_theorem), 225 | ("collect_stub", test_collect_stub), 226 | ("conflict simple", test_conflict_simple), 227 | ("conflict poly", test_conflict_poly), 228 | ("conflict auxiliary", test_conflict_auxiliary), 229 | ("conflict simple def", test_conflict_simple_def), 230 | ("conflict axiom", test_conflict_axiom), 231 | ("conflict fake implementation", test_conflict_fake_implementation), 232 | ("conflict fail idempotent", test_conflict_fail_idempotent), 233 | ("conflict fail delete definition", test_conflict_fail_delete_definition), 234 | ("conflict fail inductive modification", test_conflict_fail_inductive_modification), 235 | ("conflict fail noncomputable", test_conflict_fail_noncomputable), 236 | ("conflict fail add axiom", test_conflict_fail_add_axiom), 237 | ] 238 | tests.map (fun (name, test) => (name, runTest $ test env)) 239 | -------------------------------------------------------------------------------- /Pantograph/Library.lean: -------------------------------------------------------------------------------- 1 | import Pantograph.Environment 2 | import Pantograph.Goal 3 | import Pantograph.Protocol 4 | import Pantograph.Delate 5 | 6 | namespace Pantograph 7 | 8 | open Lean 9 | 10 | /-- This is better than the default version since it handles `.` and doesn't 11 | crash the program when it fails. -/ 12 | def setOptionFromString' (opts : Options) (entry : String) : ExceptT String IO Options := do 13 | let ps := (entry.splitOn "=").map String.trim 14 | let [key, val] ← pure ps | throw "invalid configuration option entry, it must be of the form ' = '" 15 | let key := key.toName 16 | let defValue ← getOptionDefaultValue key 17 | match defValue with 18 | | DataValue.ofString _ => pure $ opts.setString key val 19 | | DataValue.ofBool _ => 20 | match val with 21 | | "true" => pure $ opts.setBool key true 22 | | "false" => pure $ opts.setBool key false 23 | | _ => throw s!"invalid Bool option value '{val}'" 24 | | DataValue.ofName _ => pure $ opts.setName key val.toName 25 | | DataValue.ofNat _ => 26 | match val.toNat? with 27 | | none => throw s!"invalid Nat option value '{val}'" 28 | | some v => pure $ opts.setNat key v 29 | | DataValue.ofInt _ => 30 | match val.toInt? with 31 | | none => throw s!"invalid Int option value '{val}'" 32 | | some v => pure $ opts.setInt key v 33 | | DataValue.ofSyntax _ => throw s!"invalid Syntax option value" 34 | 35 | def runMetaM { α } (metaM: MetaM α): CoreM α := 36 | metaM.run' 37 | 38 | def errorI (type desc: String): Protocol.InteractionError := { error := type, desc := desc } 39 | 40 | /-- Adds the given paths to Lean package search path. This must run with at 41 | least an empty string, otherwise Lean will not be able to find any symbols -/ 42 | @[export pantograph_init_search] 43 | unsafe def initSearch (sp: String := ""): IO Unit := do 44 | Lean.enableInitializersExecution 45 | Lean.initSearchPath (← Lean.findSysroot) (sp := System.SearchPath.parse sp) 46 | 47 | /-- Creates a Core.Context object needed to run all monads -/ 48 | @[export pantograph_create_core_context] 49 | def createCoreContext (options : Array String) : IO Core.Context := do 50 | let options? ← options.foldlM setOptionFromString' Options.empty |>.run 51 | let options ← match options? with 52 | | .ok options => pure options 53 | | .error e => throw $ IO.userError s!"Options cannot be parsed: {e}" 54 | return { 55 | currNamespace := `Cirno, 56 | openDecls := [], -- No 'open' directives needed 57 | fileName := "", 58 | fileMap := { source := "", positions := #[0] }, 59 | options, 60 | maxRecDepth := maxRecDepth.get options, 61 | } 62 | 63 | /-- Creates a Core.State object needed to run all monads -/ 64 | @[export pantograph_create_core_state] 65 | def createCoreState (env : Environment) : IO Core.State := do 66 | return { env } 67 | 68 | @[export pantograph_parse_elab_type_m] 69 | def parseElabType (type : String) : Protocol.FallibleT Elab.TermElabM Expr := do 70 | let env ← MonadEnv.getEnv 71 | let syn ← match parseTerm env type with 72 | | .error str => Protocol.throw $ errorI "parse" str 73 | | .ok syn => pure syn 74 | match ← elabType syn with 75 | | .error str => Protocol.throw $ errorI "elab" str 76 | | .ok expr => return (← instantiateMVars expr) 77 | 78 | /-- This must be a TermElabM since the parsed expr contains extra information -/ 79 | @[export pantograph_parse_elab_expr_m] 80 | def parseElabExpr (expr : String) (expectedType? : Option String := .none) : Protocol.FallibleT Elab.TermElabM Expr := do 81 | let env ← MonadEnv.getEnv 82 | let expectedType? ← expectedType?.mapM parseElabType 83 | let syn ← match parseTerm env expr with 84 | | .error str => Protocol.throw $ errorI "parse" str 85 | | .ok syn => pure syn 86 | match ← elabTerm syn expectedType? with 87 | | .error str => Protocol.throw $ errorI "elab" str 88 | | .ok expr => return (← instantiateMVars expr) 89 | 90 | @[export pantograph_expr_echo_m] 91 | def exprEcho (expr: String) (expectedType?: Option String := .none) (options: @&Protocol.Options := {}): 92 | Protocol.FallibleT Elab.TermElabM Protocol.ExprEchoResult := do 93 | let expr ← parseElabExpr expr expectedType? 94 | try 95 | let type ← unfoldAuxLemmas (← Meta.inferType expr) 96 | return { 97 | type := (← serializeExpression options type), 98 | expr := (← serializeExpression options expr), 99 | } 100 | catch exception => 101 | Protocol.throw $ errorI "typing" (← exception.toMessageData.toString) 102 | 103 | @[export pantograph_goal_start_expr_m] 104 | def goalStartExpr (expr: String) : Protocol.FallibleT Elab.TermElabM GoalState := do 105 | let t ← parseElabType expr 106 | Elab.Term.synthesizeSyntheticMVarsUsingDefault 107 | GoalState.create t 108 | 109 | @[export pantograph_goal_serialize_m] 110 | def goalSerialize (state: GoalState) (options: @&Protocol.Options): CoreM (Array Protocol.Goal) := 111 | runMetaM <| state.serializeGoals options 112 | 113 | @[export pantograph_goal_print_m] 114 | def goalPrint (state: GoalState) (rootExpr: Bool) (parentExprs: Bool) 115 | (goals: Bool) (extraMVars : Array Name) (options: @&Protocol.Options) 116 | : CoreM Protocol.GoalPrintResult := runMetaM do 117 | state.restoreMetaM 118 | 119 | let rootExpr? := state.rootExpr? 120 | let root? ← if rootExpr then 121 | rootExpr?.mapM λ expr => state.withRootContext do 122 | serializeExpression options (← instantiateAll expr) 123 | else 124 | pure .none 125 | let parentExprs? ← if parentExprs then 126 | .some <$> state.parentMVars.mapM λ parent => parent.withContext do 127 | let val? := state.getMVarEAssignment parent 128 | val?.mapM λ val => do 129 | serializeExpression options (← instantiateAll val) 130 | else 131 | pure .none 132 | let goals ← if goals then 133 | goalSerialize state options 134 | else 135 | pure #[] 136 | let extraMVars ← extraMVars.mapM λ name => do 137 | let mvarId: MVarId := ⟨name⟩ 138 | let .some _ ← mvarId.findDecl? | return {} 139 | state.withContext mvarId do 140 | let .some expr ← getExprMVarAssignment? mvarId | return {} 141 | serializeExpression options (← instantiateAll expr) 142 | let env ← getEnv 143 | return { 144 | root?, 145 | parentExprs?, 146 | goals, 147 | extraMVars, 148 | rootHasSorry := rootExpr?.map (·.hasSorry) |>.getD false, 149 | rootHasUnsafe := rootExpr?.map (env.hasUnsafe ·) |>.getD false, 150 | rootHasMVar := rootExpr?.map (·.hasExprMVar) |>.getD false, 151 | } 152 | 153 | @[export pantograph_goal_have_m] 154 | protected def GoalState.tryHave (state: GoalState) (site : Site) (binderName: Name) (type: String): Elab.TermElabM TacticResult := do 155 | let type ← match (← parseTermM type) with 156 | | .ok syn => pure syn 157 | | .error error => return .parseError error 158 | state.restoreElabM 159 | state.tryTacticM site $ Tactic.evalHave binderName type 160 | protected def GoalState.tryLet (state : GoalState) (site : Site) (binderName : Name) (type : String) 161 | : Elab.TermElabM TacticResult := do 162 | state.restoreElabM 163 | let type ← match Parser.runParserCategory 164 | (env := ← MonadEnv.getEnv) 165 | (catName := `term) 166 | (input := type) 167 | (fileName := ← getFileName) with 168 | | .ok syn => pure syn 169 | | .error error => return .parseError error 170 | state.tryTacticM site $ Tactic.evalLet binderName type 171 | @[export pantograph_goal_try_define_m] 172 | protected def GoalState.tryDefine (state: GoalState) (site : Site) (binderName: Name) (expr: String): Elab.TermElabM TacticResult := do 173 | let expr ← match (← parseTermM expr) with 174 | | .ok syn => pure syn 175 | | .error error => return .parseError error 176 | state.restoreElabM 177 | state.tryTacticM site $ Tactic.evalDefine binderName expr 178 | @[export pantograph_goal_try_draft_m] 179 | protected def GoalState.tryDraft (state: GoalState) (site : Site) (expr: String): Elab.TermElabM TacticResult := do 180 | let expr ← match (← parseTermM expr) with 181 | | .ok syn => pure syn 182 | | .error error => return .parseError error 183 | state.restoreElabM 184 | state.tryTacticM site $ Tactic.evalDraft expr 185 | 186 | -- Cancel the token after a timeout. 187 | @[export pantograph_run_cancel_token_with_timeout_m] 188 | def runCancelTokenWithTimeout (cancelToken : IO.CancelToken) (timeout : UInt32) : IO Unit := do 189 | let _ ← IO.asTask do 190 | IO.sleep timeout 191 | cancelToken.set 192 | return () 193 | 194 | def spawnCancelToken (timeout : UInt32) : IO IO.CancelToken := do 195 | let token ← IO.CancelToken.new 196 | runCancelTokenWithTimeout token timeout 197 | return token 198 | -------------------------------------------------------------------------------- /Pantograph/Tactic/Fragment.lean: -------------------------------------------------------------------------------- 1 | /- Fragmented tactics are the tactics which can give incremental feedback and 2 | whose integrity as a block is crucial to its operation. e.g. `calc` or `conv`. 3 | Here, a unified system handles all fragments. 4 | 5 | Inside a tactic fragment, the parser category may be different. An incomplete 6 | fragmented tactic may not be elaboratable.. 7 | 8 | In line with continuation/resumption paradigms, the exit function of a fragment 9 | tactic is responsible for resuming incomplete goals with fragments. For example, 10 | when a conversion tactic finishes, the sentinels should resume the root of the 11 | conversion tactic goal. The user cannot be expected to execute this resumption, 12 | since the root is automatically dormanted at the entry of the conversion tactic 13 | mode. 14 | -/ 15 | import Lean.Meta 16 | import Lean.Elab 17 | 18 | open Lean 19 | 20 | namespace Pantograph 21 | 22 | inductive Fragment where 23 | | calc (prevRhs? : Option Expr) 24 | | conv (rhs : MVarId) 25 | -- This goal is spawned from a `conv` 26 | | convSentinel (parent : MVarId) 27 | deriving BEq, Inhabited 28 | 29 | abbrev FragmentMap := Std.HashMap MVarId Fragment 30 | def FragmentMap.empty : FragmentMap := Std.HashMap.emptyWithCapacity 2 31 | protected def FragmentMap.filter (map : FragmentMap) (pred : MVarId → Fragment → Bool) : FragmentMap := 32 | map.fold (init := FragmentMap.empty) λ acc mvarId fragment => 33 | if pred mvarId fragment then 34 | acc.insert mvarId fragment 35 | else 36 | acc 37 | 38 | protected def Fragment.map (fragment : Fragment) (mapExpr : Expr → CoreM Expr) : CoreM Fragment := 39 | let mapMVar (mvarId : MVarId) : CoreM MVarId := 40 | return (← mapExpr (.mvar mvarId)) |>.mvarId! 41 | match fragment with 42 | | .calc prevRhs? => return .calc (← prevRhs?.mapM mapExpr) 43 | | .conv rhs => do 44 | let rhs' ← mapMVar rhs 45 | return .conv rhs' 46 | | .convSentinel parent => do 47 | return .convSentinel (← mapMVar parent) 48 | 49 | protected def Fragment.enterCalc : Fragment := .calc .none 50 | protected def Fragment.enterConv : Elab.Tactic.TacticM FragmentMap := do 51 | let goal ← Elab.Tactic.getMainGoal 52 | goal.checkNotAssigned `GoalState.conv 53 | let (rhs, newGoal) ← goal.withContext do 54 | let target ← instantiateMVars (← goal.getType) 55 | let (rhs, newGoal) ← Elab.Tactic.Conv.mkConvGoalFor target 56 | pure (rhs.mvarId!, newGoal.mvarId!) 57 | Elab.Tactic.replaceMainGoal [newGoal] 58 | return FragmentMap.empty 59 | |>.insert goal (.conv rhs) 60 | |>.insert newGoal (.convSentinel goal) 61 | 62 | protected partial def Fragment.exit (fragment : Fragment) (goal : MVarId) (fMap: FragmentMap) 63 | : Elab.Tactic.TacticM (FragmentMap × List MVarId) := 64 | match fragment with 65 | | .calc .. => do 66 | Elab.Tactic.setGoals [goal] 67 | return (fMap.erase goal, []) 68 | | .conv rhs => do 69 | let goals := (← Elab.Tactic.getGoals).filter λ descendant => 70 | match fMap[descendant]? with 71 | | .some s => (.convSentinel goal) == s 72 | | _ => false -- Not a conv goal from this 73 | -- Close all existing goals with `refl` 74 | for mvarId in goals do 75 | liftM <| mvarId.refl <|> mvarId.inferInstance <|> pure () 76 | unless (← goals.filterM (·.isAssignedOrDelayedAssigned)).isEmpty do 77 | throwError "convert tactic failed, there are unsolved goals\n{Elab.goalsToMessageData (goals)}" 78 | 79 | -- Ensure the meta tactic runs on `goal` even if its dormant by forcing resumption 80 | Elab.Tactic.setGoals $ goal :: (← Elab.Tactic.getGoals) 81 | let targetNew ← instantiateMVars (.mvar rhs) 82 | let proof ← instantiateMVars (.mvar goal) 83 | 84 | Elab.Tactic.liftMetaTactic1 (·.replaceTargetEq targetNew proof) 85 | 86 | -- Try to solve maiinline by rfl 87 | let mvarId ← Elab.Tactic.getMainGoal 88 | liftM <| mvarId.refl <|> mvarId.inferInstance <|> pure () 89 | Elab.Tactic.pruneSolvedGoals 90 | let fMap := fMap.filter λ mvarId fragment => 91 | !(mvarId == goal || fragment == .convSentinel goal) 92 | return (fMap, [mvarId]) 93 | | .convSentinel parent => 94 | let parentFragment := fMap[parent]! 95 | parentFragment.exit parent (fMap.erase goal) 96 | 97 | protected def Fragment.step (fragment : Fragment) (goal : MVarId) (s : String) (map : FragmentMap) 98 | : Elab.Tactic.TacticM FragmentMap := goal.withContext do 99 | assert! ¬ (← goal.isAssigned) 100 | match fragment with 101 | | .calc prevRhs? => do 102 | let .ok stx := Parser.runParserCategory 103 | (env := ← getEnv) 104 | (catName := `term) 105 | (input := s) 106 | (fileName := ← getFileName) | throwError s!"Failed to parse calc element {s}" 107 | let `(term|$pred) := stx 108 | let decl ← goal.getDecl 109 | let target ← instantiateMVars decl.type 110 | let tag := decl.userName 111 | 112 | let mut step ← Elab.Term.elabType <| ← do 113 | if let some prevRhs := prevRhs? then 114 | Elab.Term.annotateFirstHoleWithType pred (← Meta.inferType prevRhs) 115 | else 116 | pure pred 117 | 118 | let some (_, lhs, rhs) ← Elab.Term.getCalcRelation? step | 119 | throwErrorAt pred "invalid 'calc' step, relation expected{indentExpr step}" 120 | if let some prevRhs := prevRhs? then 121 | unless ← Meta.isDefEqGuarded lhs prevRhs do 122 | throwErrorAt pred "invalid 'calc' step, left-hand-side is{indentD m!"{lhs} : {← Meta.inferType lhs}"}\nprevious right-hand-side is{indentD m!"{prevRhs} : {← Meta.inferType prevRhs}"}" 123 | 124 | -- Creates a mvar to represent the proof that the calc tactic solves the 125 | -- current branch 126 | -- In the Lean `calc` tactic this is gobbled up by 127 | -- `withCollectingNewGoalsFrom` 128 | let mut proof ← Meta.mkFreshExprMVarAt (← getLCtx) (← Meta.getLocalInstances) step 129 | (userName := tag ++ `calc) 130 | let mvarBranch := proof.mvarId! 131 | 132 | let mut proofType ← Meta.inferType proof 133 | let mut remainder? := Option.none 134 | 135 | -- The calc tactic either solves the main goal or leaves another relation. 136 | -- Replace the main goal, and save the new goal if necessary 137 | unless ← Meta.isDefEq proofType target do 138 | let rec throwFailed := 139 | throwError "'calc' tactic failed, has type{indentExpr proofType}\nbut it is expected to have type{indentExpr target}" 140 | let some (_, _, rhs) ← Elab.Term.getCalcRelation? proofType | throwFailed 141 | let some (r, _, rhs') ← Elab.Term.getCalcRelation? target | throwFailed 142 | let lastStep := mkApp2 r rhs rhs' 143 | let lastStepGoal ← Meta.mkFreshExprSyntheticOpaqueMVar lastStep tag 144 | (proof, proofType) ← Elab.Term.mkCalcTrans proof proofType lastStepGoal lastStep 145 | unless ← Meta.isDefEq proofType target do throwFailed 146 | remainder? := .some lastStepGoal.mvarId! 147 | goal.assign proof 148 | let map := map.erase goal 149 | 150 | let goals := [ mvarBranch ] ++ remainder?.toList 151 | Elab.Tactic.setGoals goals 152 | match remainder? with 153 | | .some goal => return map.insert goal $ .calc (.some rhs) 154 | | .none => return map 155 | | .conv .. => do 156 | throwError "Direct operation on conversion tactic parent goal is not allowed" 157 | | fragment@(.convSentinel parent) => do 158 | assert! isLHSGoal? (← goal.getType) |>.isSome 159 | let tactic ← match Parser.runParserCategory 160 | (env := ← MonadEnv.getEnv) 161 | (catName := `conv) 162 | (input := s) 163 | (fileName := ← getFileName) with 164 | | .ok stx => pure $ stx 165 | | .error error => throwError error 166 | let oldGoals ← Elab.Tactic.getGoals 167 | -- Label newly generated goals as conv sentinels 168 | Elab.Tactic.evalTactic tactic 169 | let newConvGoals ← (← Elab.Tactic.getUnsolvedGoals).filterM λ g => do 170 | -- conv tactic might generate non-conv goals 171 | if oldGoals.contains g then 172 | return false 173 | return isLHSGoal? (← g.getType) |>.isSome 174 | -- Conclude the conv by exiting the parent fragment if new goals is empty 175 | if newConvGoals.isEmpty then 176 | let hasSiblingFragment := map.fold (init := false) λ flag _ fragment => 177 | if flag then 178 | true 179 | else match fragment with 180 | | .convSentinel parent' => parent == parent' 181 | | _ => false 182 | if ¬ hasSiblingFragment then 183 | -- This fragment must exist since we have conv goals 184 | let parentFragment := map[parent]! 185 | -- All descendants exhausted. Exit from the parent conv. 186 | return ← Prod.fst <$> parentFragment.exit parent map 187 | return newConvGoals.foldl (init := map) λ acc g => 188 | acc.insert g fragment 189 | 190 | end Pantograph 191 | -------------------------------------------------------------------------------- /Test/Common.lean: -------------------------------------------------------------------------------- 1 | import Pantograph.Goal 2 | import Pantograph.Library 3 | import Pantograph.Protocol 4 | import Pantograph.Frontend 5 | import LSpec 6 | 7 | namespace Pantograph 8 | 9 | open Lean 10 | 11 | deriving instance Repr for Expr 12 | -- Use strict equality check for expressions 13 | instance : BEq Expr := ⟨Expr.equal⟩ 14 | 15 | def uniq (n: Nat): Name := .num (.str .anonymous "_uniq") n 16 | 17 | -- Auxiliary functions 18 | namespace Protocol 19 | def Goal.devolatilizeVars (goal: Goal): Goal := 20 | { 21 | goal with 22 | vars := goal.vars.map removeInternalAux, 23 | 24 | } 25 | where removeInternalAux (v: Variable): Variable := 26 | { 27 | v with 28 | name := .anonymous 29 | } 30 | /-- Set internal names to "" -/ 31 | def Goal.devolatilize (goal: Goal): Goal := 32 | { 33 | goal.devolatilizeVars with 34 | name := .anonymous, 35 | } 36 | 37 | deriving instance DecidableEq, Repr for Name 38 | deriving instance DecidableEq, Repr for Expression 39 | deriving instance DecidableEq, Repr for Variable 40 | deriving instance DecidableEq, Repr for Goal 41 | deriving instance DecidableEq, Repr for ExprEchoResult 42 | deriving instance DecidableEq, Repr for InteractionError 43 | deriving instance DecidableEq, Repr for Option 44 | end Protocol 45 | 46 | namespace Condensed 47 | 48 | deriving instance BEq, Repr for LocalDecl 49 | deriving instance BEq, Repr for Goal 50 | 51 | -- Enable string interpolation 52 | instance : ToString FVarId where 53 | toString id := id.name.toString 54 | instance : ToString MVarId where 55 | toString id := id.name.toString 56 | 57 | protected def LocalDecl.devolatilize (decl: LocalDecl): LocalDecl := 58 | { 59 | decl with fvarId := { name := .anonymous } 60 | } 61 | protected def Goal.devolatilize (goal: Goal): Goal := 62 | { 63 | goal with 64 | mvarId := { name := .anonymous }, 65 | context := goal.context.map LocalDecl.devolatilize 66 | } 67 | 68 | end Condensed 69 | 70 | def GoalState.get! (state: GoalState) (i: Nat): MVarId := state.goals[i]! 71 | def GoalState.tacticOn (state: GoalState) (goalId: Nat) (tactic: String) := state.tryTactic (state.get! goalId) tactic 72 | def GoalState.tacticOn' (state: GoalState) (goalId: Nat) (tactic: TSyntax `tactic) := 73 | state.tryTacticM (state.get! goalId) (Elab.Tactic.evalTactic tactic) true 74 | 75 | def TacticResult.toString : TacticResult → String 76 | | .success state _messages => s!".success ({state.goals.length} goals)" 77 | | .failure messages => 78 | s!".failure ({messages.size} messages)" 79 | | .parseError error => s!".parseError {error}" 80 | | .invalidAction error => s!".invalidAction {error}" 81 | 82 | namespace Test 83 | 84 | def expectationFailure (desc: String) (error: String): LSpec.TestSeq := LSpec.test desc (LSpec.ExpectationFailure "ok _" error) 85 | def assertUnreachable (message: String): LSpec.TestSeq := LSpec.check message false 86 | 87 | def parseFailure (error: String) := expectationFailure "parse" error 88 | def elabFailure (error: String) := expectationFailure "elab" error 89 | 90 | def runCoreMSeq (env: Environment) (coreM: CoreM LSpec.TestSeq) (options: Array String := #[]): IO LSpec.TestSeq := do 91 | let coreContext: Core.Context ← createCoreContext options 92 | match ← (coreM.run' coreContext { env := env }).toBaseIO with 93 | | .error exception => 94 | return LSpec.test "Exception" (s!"internal exception #{← exception.toMessageData.toString}" = "") 95 | | .ok a => return a 96 | def runMetaMSeq (env: Environment) (metaM: MetaM LSpec.TestSeq): IO LSpec.TestSeq := 97 | runCoreMSeq env metaM.run' 98 | def runTermElabMInMeta { α } (termElabM: Elab.TermElabM α): MetaM α := 99 | termElabM.run' (ctx := defaultElabContext) 100 | def runTermElabMInCore { α } (termElabM: Elab.TermElabM α): CoreM α := 101 | (runTermElabMInMeta termElabM).run' 102 | def runTermElabMSeq (env: Environment) (termElabM: Elab.TermElabM LSpec.TestSeq): IO LSpec.TestSeq := 103 | runMetaMSeq env $ termElabM.run' (ctx := defaultElabContext) 104 | 105 | def exprToStr (e: Expr): MetaM String := toString <$> Meta.ppExpr e 106 | 107 | def strToTermSyntax (s: String): CoreM Syntax := do 108 | let .ok stx := Parser.runParserCategory 109 | (env := ← MonadEnv.getEnv) 110 | (catName := `term) 111 | (input := s) 112 | (fileName := ← getFileName) | panic! s!"Failed to parse {s}" 113 | return stx 114 | def parseSentence (s : String) (expectedType? : Option Expr := .none) : Elab.TermElabM Expr := do 115 | let stx ← match Parser.runParserCategory 116 | (env := ← MonadEnv.getEnv) 117 | (catName := `term) 118 | (input := s) 119 | (fileName := ← getFileName) with 120 | | .ok syn => pure syn 121 | | .error error => throwError "Failed to parse: {error}" 122 | Elab.Term.elabTerm (stx := stx) expectedType? 123 | 124 | def runTacticOnMVar (tacticM: Elab.Tactic.TacticM Unit) (goal: MVarId): Elab.TermElabM (List MVarId) := do 125 | let (_, newGoals) ← tacticM { elaborator := .anonymous } |>.run { goals := [goal] } 126 | return newGoals.goals 127 | def mvarUserNameAndType (mvarId: MVarId): MetaM (Name × String) := do 128 | let name := (← mvarId.getDecl).userName 129 | let t ← exprToStr (← mvarId.getType) 130 | return (name, t) 131 | 132 | def extractEnvAfterUnit (env : Environment) (unit : String) 133 | : IO Environment := do 134 | let (context, state) ← Frontend.createContextStateFromFile 135 | (file := unit) (env? := .some env) 136 | let m := show Frontend.FrontendM _ from Frontend.executeFrontend λ _ => pure () 137 | let (_, state) ← m.run {} |>.run context |>.run state 138 | return state.commandState.env 139 | 140 | -- Monadic testing 141 | 142 | abbrev TestT := StateRefT' IO.RealWorld LSpec.TestSeq 143 | 144 | section Monadic 145 | 146 | variable [Monad m] [MonadLiftT (ST IO.RealWorld) m] 147 | 148 | def addTest (test: LSpec.TestSeq) : TestT m Unit := do 149 | set $ (← get) ++ test 150 | 151 | def checkEq [DecidableEq α] [Repr α] (desc : String) (lhs rhs : α) : TestT m Unit := do 152 | addTest $ LSpec.check desc (lhs = rhs) 153 | def checkTrue (desc : String) (flag : Bool) : TestT m Unit := do 154 | addTest $ LSpec.check desc flag 155 | def checkFalse (desc : String) (flag : Bool) : TestT m Unit := do 156 | addTest $ LSpec.check desc !flag 157 | def fail (desc : String) : TestT m Unit := do 158 | addTest $ LSpec.check desc false 159 | 160 | def runTest (t: TestT m Unit): m LSpec.TestSeq := 161 | Prod.snd <$> t.run LSpec.TestSeq.done 162 | def runTestWithResult { α } (t: TestT m α): m (α × LSpec.TestSeq) := 163 | t.run LSpec.TestSeq.done 164 | def runTestCoreM (env: Environment) (coreM: TestT CoreM Unit) (options: Array String := #[]): IO LSpec.TestSeq := do 165 | runCoreMSeq env (runTest coreM) options 166 | 167 | end Monadic 168 | 169 | def runTestTermElabM (env: Environment) (t: TestT Elab.TermElabM Unit): 170 | IO LSpec.TestSeq := 171 | runTermElabMSeq env $ runTest t 172 | def transformTestT { α } { μ μ' : Type → Type } 173 | [Monad μ] [Monad μ'] [MonadLiftT (ST IO.RealWorld) μ] [MonadLiftT (ST IO.RealWorld) μ'] 174 | (tr : {β : Type} → μ β → μ' β) (m : TestT μ α) : TestT μ' α := do 175 | let tests ← get 176 | let (a, tests) ← tr (β := α × LSpec.TestSeq) (m.run tests) 177 | set tests 178 | return a 179 | 180 | def cdeclOf (userName : Name) (type : Expr) : Condensed.LocalDecl := 181 | { userName, type } 182 | 183 | def buildGoal (nameType : List (Name × String)) (target : String) (userName?: Option Name := .none): 184 | Protocol.Goal := 185 | { 186 | userName?, 187 | target := { pp? := .some target}, 188 | vars := (nameType.map fun x => ({ 189 | userName := x.fst, 190 | type? := .some { pp? := .some x.snd }, 191 | })).toArray 192 | } 193 | 194 | namespace Tactic 195 | 196 | /-- Create an aux lemma and assigns it to `mvarId`, which is circuitous, but 197 | exercises the aux lemma generator. -/ 198 | def assignWithAuxLemma (type : Expr) (value? : Option Expr := .none) : Elab.Tactic.TacticM Unit := do 199 | let type ← instantiateMVars type 200 | let value ← match value? with 201 | | .some value => instantiateMVars value 202 | | .none => Meta.mkSorry type (synthetic := false) 203 | if type.hasExprMVar then 204 | throwError "Type has expression mvar" 205 | if value.hasExprMVar then 206 | throwError "value has expression mvar" 207 | let goal ← Elab.Tactic.getMainGoal 208 | goal.withContext do 209 | let name ← Meta.mkAuxLemma [] type value 210 | unless ← Meta.isDefEq type (← goal.getType) do 211 | throwError "Type provided is incorrect" 212 | goal.assign (.const name []) 213 | Elab.Tactic.pruneSolvedGoals 214 | 215 | /-- A function that takes many iterations to finish -/ 216 | partial def collatz (n : Nat) : Elab.Tactic.TacticM Unit := do 217 | if n ≤ 1 then 218 | let g ← Elab.Tactic.getMainGoal 219 | g.assign (.lit (.natVal n)) 220 | Elab.Tactic.pruneSolvedGoals 221 | else if n % 2 == 0 then 222 | Meta.withIncRecDepth do 223 | collatz (n / 2) 224 | else 225 | Meta.withIncRecDepth do 226 | collatz (3 * n + 1) 227 | 228 | end Tactic 229 | 230 | end Test 231 | 232 | end Pantograph 233 | -------------------------------------------------------------------------------- /Test/Delate.lean: -------------------------------------------------------------------------------- 1 | import Pantograph.Delate 2 | import Test.Common 3 | 4 | open Lean Pantograph 5 | 6 | namespace Pantograph.Test.Delate 7 | 8 | deriving instance Repr, DecidableEq for Protocol.BoundExpression 9 | 10 | def test_serializeName: LSpec.TestSeq := 11 | let quote := "\"" 12 | let escape := "\\" 13 | LSpec.test "a.b.1" (serializeName (Name.num (.str (.str .anonymous "a") "b") 1) = "a.b.1") ++ 14 | LSpec.test "seg.«a.b»" (serializeName (Name.str (.str .anonymous "seg") "a.b") = s!"{quote}seg.«a.b»{quote}") ++ 15 | -- Pathological test case 16 | LSpec.test s!"«̈{escape}{quote}»" (serializeName (Name.str .anonymous s!"{escape}{quote}") = s!"{quote}«{escape}{quote}»{quote}") 17 | 18 | def test_expr_to_binder (env: Environment): IO LSpec.TestSeq := do 19 | let entries: List (Name × Protocol.BoundExpression) := [ 20 | ("Nat.add_comm".toName, { binders := #[("n", "Nat"), ("m", "Nat")], target := "n + m = m + n" }), 21 | ("Nat.le_of_succ_le".toName, { binders := #[("n", "Nat"), ("m", "Nat"), ("h", "n.succ ≤ m")], target := "n ≤ m" }) 22 | ] 23 | runCoreMSeq env $ entries.foldlM (λ suites (symbol, target) => do 24 | let env ← MonadEnv.getEnv 25 | let expr := env.find? symbol |>.get! |>.type 26 | let test := LSpec.check symbol.toString ((← typeExprToBound expr) = target) 27 | return LSpec.TestSeq.append suites test) LSpec.TestSeq.done |>.run' 28 | 29 | def test_sexp_of_symbol (env: Environment): IO LSpec.TestSeq := do 30 | let entries: List (String × String) := [ 31 | -- This one contains unhygienic variable names which must be suppressed 32 | ("Nat.add", "(:forall a (:c Nat) (:forall a (:c Nat) (:c Nat)))"), 33 | -- These ones are normal and easy 34 | ("Nat.add_one", "(:forall n (:c Nat) ((:c Eq) (:c Nat) ((:c HAdd.hAdd) (:c Nat) (:c Nat) (:c Nat) ((:c instHAdd) (:c Nat) (:c instAddNat)) 0 ((:c OfNat.ofNat) (:c Nat) (:lit 1) ((:c instOfNatNat) (:lit 1)))) ((:c Nat.succ) 0)))"), 35 | ("Nat.le_of_succ_le", "(:forall n (:c Nat) (:forall m (:c Nat) (:forall h ((:c LE.le) (:c Nat) (:c instLENat) ((:c Nat.succ) 1) 0) ((:c LE.le) (:c Nat) (:c instLENat) 2 1)) :i) :i)"), 36 | -- Handling of higher order types 37 | ("Or", "(:forall a (:sort 0) (:forall b (:sort 0) (:sort 0)))"), 38 | ("List", "(:forall α (:sort (+ u 1)) (:sort (+ u 1)))") 39 | ] 40 | runMetaMSeq env $ entries.foldlM (λ suites (symbol, target) => do 41 | let env ← MonadEnv.getEnv 42 | let expr := env.find? symbol.toName |>.get! |>.type 43 | let test := LSpec.check symbol ((← serializeExpressionSexp expr) = target) 44 | return LSpec.TestSeq.append suites test) LSpec.TestSeq.done 45 | 46 | def test_sexp_of_elab (env: Environment): IO LSpec.TestSeq := do 47 | let entries: List (String × (List Name) × String) := [ 48 | ("λ x: Nat × Bool => x.1", [], "(:lambda x ((:c Prod) (:c Nat) (:c Bool)) ((:c Prod.fst) (:c Nat) (:c Bool) 0))"), 49 | ("λ {α: Sort (u + 1)} => List α", [`u], "(:lambda α (:sort (+ u 1)) ((:c List) 0) :i)"), 50 | ("λ {α} => List α", [], "(:lambda α (:sort (+ (:mv _uniq.4) 1)) ((:c List) 0) :i)"), 51 | ("(2: Nat) <= (5: Nat)", [], "((:c LE.le) (:mv _uniq.16) (:mv _uniq.17) ((:c OfNat.ofNat) (:mv _uniq.4) (:lit 2) (:mv _uniq.5)) ((:c OfNat.ofNat) (:mv _uniq.11) (:lit 5) (:mv _uniq.12)))"), 52 | ] 53 | entries.foldlM (λ suites (source, levels, target) => 54 | let termElabM := do 55 | let env ← MonadEnv.getEnv 56 | let s ← match parseTerm env source with 57 | | .ok s => pure s 58 | | .error e => return parseFailure e 59 | let expr ← match (← elabTerm s) with 60 | | .ok expr => pure expr 61 | | .error e => return elabFailure e 62 | return LSpec.check source ((← serializeExpressionSexp expr) = target) 63 | let metaM := (Elab.Term.withLevelNames levels termElabM).run' (ctx := defaultElabContext) 64 | return LSpec.TestSeq.append suites (← runMetaMSeq env metaM)) 65 | LSpec.TestSeq.done 66 | 67 | def test_sexp_of_expr (env: Environment): IO LSpec.TestSeq := do 68 | let entries: List (Expr × String) := [ 69 | (.lam `p (.sort .zero) 70 | (.lam `q (.sort .zero) 71 | (.lam `k (mkApp2 (.const `And []) (.bvar 1) (.bvar 0)) 72 | (.proj `And 1 (.bvar 0)) 73 | .default) 74 | .implicit) 75 | .implicit, 76 | "(:lambda p (:sort 0) (:lambda q (:sort 0) (:lambda k ((:c And) 1 0) ((:c And.right) _ _ 0)) :i) :i)" 77 | ), 78 | ] 79 | let termElabM: Elab.TermElabM LSpec.TestSeq := entries.foldlM (λ suites (expr, target) => do 80 | let env ← MonadEnv.getEnv 81 | let testCaseName := target.take 10 82 | let test := LSpec.check testCaseName ((← serializeExpressionSexp expr) = target) 83 | return LSpec.TestSeq.append suites test) LSpec.TestSeq.done 84 | runMetaMSeq env $ termElabM.run' (ctx := defaultElabContext) 85 | 86 | -- Instance parsing 87 | def test_instance (env: Environment): IO LSpec.TestSeq := 88 | runMetaMSeq env do 89 | let env ← MonadEnv.getEnv 90 | let source := "λ x y: Nat => HAdd.hAdd Nat Nat Nat (instHAdd Nat instAddNat) x y" 91 | let s := parseTerm env source |>.toOption |>.get! 92 | let _expr := (← runTermElabMInMeta <| elabTerm s) |>.toOption |>.get! 93 | return LSpec.TestSeq.done 94 | 95 | def test_projection_prod (env: Environment) : IO LSpec.TestSeq:= runTest do 96 | let struct := .app (.bvar 1) (.bvar 0) 97 | let expr := .proj `Prod 1 struct 98 | let .field projector numParams := analyzeProjection env expr | 99 | fail "`Prod has fields" 100 | checkEq "projector" projector `Prod.snd 101 | checkEq "numParams" numParams 2 102 | 103 | def test_projection_exists (env: Environment) : IO LSpec.TestSeq:= runTest do 104 | let struct := .app (.bvar 1) (.bvar 0) 105 | let expr := .proj `Exists 1 struct 106 | let .singular recursor numParams numFields := analyzeProjection env expr | 107 | fail "`Exists has no projectors" 108 | checkEq "recursor" recursor `Exists.recOn 109 | checkEq "numParams" numParams 2 110 | checkEq "numFields" numFields 2 111 | 112 | def test_matcher : TestT Elab.TermElabM Unit := do 113 | let t ← parseSentence "Nat → Nat" 114 | let e ← parseSentence "fun (n : Nat) => match n with | 0 => 0 | k => k" (.some t) 115 | let .some _ ← Meta.matchMatcherApp? e.bindingBody! | fail "Must be a matcher app" 116 | let e' ← instantiateAll e 117 | checkTrue "ok" <| ← Meta.isTypeCorrect e' 118 | 119 | def test_intro_delay_intro : TestT Elab.TermElabM Unit := do 120 | let statement ← Elab.Term.elabTerm (← `(term|∀ (i : Nat), { f : Nat → Nat // ∀ (j : Nat), f i = j })) .none 121 | Meta.forallTelescope statement λ _fvars target => do 122 | let goal := (← Meta.mkFreshExprSyntheticOpaqueMVar target).mvarId! 123 | let [cond, f] ← goal.applyConst `Subtype.mk | fail "Incorrect number of goals" 124 | let (_fBinder, fBody) ← f.intro1 125 | cond.withContext do 126 | let opt ← toDelayedMVarInvocation (.mvar f) 127 | checkTrue "condBody/?f" opt.isNone 128 | let sexp ← serializeExpressionSexp (← instantiateAll $ ← cond.getType) 129 | checkTrue "condBody/target" $ sexp.startsWith "(:forall j (:c Nat) ((:c Eq) (:c Nat) ((:lambda a (:c Nat) (:subst" 130 | let (_condBinder, condBody) ← cond.intro1 131 | condBody.withContext do 132 | let opt ← toDelayedMVarInvocation (.mvar f) 133 | checkTrue "condBody/?f" opt.isNone 134 | let opt ← toDelayedMVarInvocation (.mvar fBody) 135 | checkTrue "condBody/?fBody" opt.isNone 136 | let sexp ← serializeExpressionSexp (← instantiateAll $ ← condBody.getType) 137 | checkTrue "condBody/target" $ sexp.startsWith "((:c Eq) (:c Nat) ((:lambda a (:c Nat) (:subst" 138 | 139 | def test_doubly_nested_delayed_assigned : TestT Elab.TermElabM Unit := do 140 | let statement ← Elab.Term.elabTerm (← `(term|∀ (i : Nat), { t : Prop // ∃ f : Nat → Nat → t, ∀ (j : Nat), f j j = f i i })) .none 141 | Meta.forallTelescope statement λ _fvars target => do 142 | let goal := (← Meta.mkFreshExprSyntheticOpaqueMVar target).mvarId! 143 | let [cond1, _t] ← goal.applyConst `Subtype.mk | fail "Incorrect number of goals [1]" 144 | let [cond2, f] ← cond1.applyConst `Exists.intro | fail "Incorrect number of goals [2]" 145 | let (_cond2F, cond2B) ← cond2.intro1 146 | let (_f1F, f1B) ← f.intro1 147 | let (_f12, _f2B) ← f1B.intro1 148 | cond2B.withContext do 149 | let sexp ← serializeExpressionSexp (← instantiateAll $ ← cond2B.getType) 150 | checkTrue "cond2B/target" $ sexp.startsWith "((:c Eq)" 151 | 152 | def suite (env: Environment): List (String × IO LSpec.TestSeq) := 153 | [ 154 | ("serializeName", do pure test_serializeName), 155 | ("Expression binder", test_expr_to_binder env), 156 | ("Sexp from symbol", test_sexp_of_symbol env), 157 | ("Sexp from elaborated expr", test_sexp_of_elab env), 158 | ("Sexp from expr", test_sexp_of_expr env), 159 | ("Instance", test_instance env), 160 | ("Projection Prod", test_projection_prod env), 161 | ("Projection Exists", test_projection_exists env), 162 | ("Matcher", runTestTermElabM env test_matcher), 163 | ("intro delay intro", runTestTermElabM env test_intro_delay_intro), 164 | ("doubly_nested_delayed_assigned", runTestTermElabM env test_doubly_nested_delayed_assigned), 165 | ] 166 | -------------------------------------------------------------------------------- /Pantograph/Environment.lean: -------------------------------------------------------------------------------- 1 | import Pantograph.Elab 2 | import Lean.Environment 3 | import Lean.Replay 4 | import Lean.Util.Path 5 | 6 | open Lean 7 | open Pantograph 8 | 9 | namespace Pantograph 10 | 11 | @[always_inline] 12 | def getAuxLemmaPrefix? (n : Name) : Option String := 13 | match n with 14 | -- `mkAuxLemma` generally allows for arbitrary prefixes but these are the ones produced by core. 15 | | .str _ s => 16 | if "_proof_".isPrefixOf s then 17 | .some "_proof" 18 | else if "_simp_".isPrefixOf s then 19 | some "_simp" 20 | else 21 | none 22 | | _ => .none 23 | @[always_inline] 24 | def isAuxLemma (n : Name) : Bool := 25 | getAuxLemmaPrefix? n |>.isSome 26 | 27 | @[export pantograph_is_name_internal] 28 | def isNameInternal (n: Name): Bool := 29 | -- Returns true if the name is an implementation detail which should not be shown to the user. 30 | isAuxLemma n ∨ n.hasMacroScopes 31 | 32 | /-- Catalog all the non-internal and safe names -/ 33 | @[export pantograph_environment_catalog] 34 | def envCatalog (env : Environment) : Array Name := 35 | env.constants.fold (init := #[]) λ acc name _ => 36 | match isNameInternal name with 37 | | false => acc.push name 38 | | true => acc 39 | 40 | @[export pantograph_environment_module_of_name] 41 | def module_of_name (env: Environment) (name: Name): Option Name := do 42 | let moduleId ← env.getModuleIdxFor? name 43 | if h : moduleId.toNat < env.allImportedModuleNames.size then 44 | return env.allImportedModuleNames[moduleId.toNat] 45 | else 46 | .none 47 | 48 | def toCompactSymbolName (n: Name) (info: ConstantInfo): String := 49 | let pref := match info with 50 | | .axiomInfo _ => "a" 51 | | .defnInfo _ => "d" 52 | | .thmInfo _ => "t" 53 | | .opaqueInfo _ => "o" 54 | | .quotInfo _ => "q" 55 | | .inductInfo _ => "i" 56 | | .ctorInfo _ => "c" 57 | | .recInfo _ => "r" 58 | s!"{pref}{toString n}" 59 | 60 | def toFilteredSymbol (n: Lean.Name) (info: Lean.ConstantInfo): Option String := 61 | if isNameInternal n || info.isUnsafe 62 | then Option.none 63 | else Option.some <| toCompactSymbolName n info 64 | 65 | abbrev ConstArray := Array (Name × ConstantInfo) 66 | abbrev DistilledEnvironment := Array Import × ConstArray 67 | 68 | def envDiff (src dst : Environment) : ConstArray := 69 | dst.constants.map₂.foldl (init := #[]) λ acc name info => 70 | if src.contains name then 71 | acc 72 | else 73 | acc.push (name, info) 74 | 75 | /-- Boil an environment down to minimal components -/ 76 | def distilEnvironment (env : Environment) (background? : Option Environment := .none) 77 | : DistilledEnvironment := 78 | let constants := match background? with 79 | | .some src => envDiff src env 80 | | .none => env.constants.map₂.toArray 81 | (env.header.imports, constants) 82 | 83 | deriving instance BEq for Import 84 | 85 | def checkEnvConflicts (src src' dst : Environment) : ExceptT String IO Environment := do 86 | let (srcImports, srcConstants) := distilEnvironment src' (background? := .some src) 87 | let mut srcConstants := srcConstants.foldl 88 | (init := Std.HashMap.emptyWithCapacity srcConstants.size) 89 | λ acc (k, v) => acc.insert k v 90 | let (dstImports, dstConstants) := distilEnvironment dst (background? := .some src) 91 | -- Replay all `dstConstants` in the environment 92 | if srcImports != dstImports then 93 | throw "Modification of imports is not allowed" 94 | -- Replay all dst constants in src 95 | let env ← try 96 | src.replay $ dstConstants.foldl (init := .emptyWithCapacity dstConstants.size) λ acc (k, v) => acc.insert k v 97 | catch ex : IO.Error => 98 | throw ex.toString 99 | -- check if `constants` can fit into `src'` 100 | for (name, dstInfo) in dstConstants do 101 | if dstInfo.type.hasSorry then 102 | throw s!"Definition type has sorry: {name}" 103 | if dstInfo.value?.map Expr.hasSorry |>.getD false then 104 | throw s!"Definition value has sorry: {name}" 105 | match srcConstants[name]? with 106 | | .some srcInfo => 107 | if srcInfo.type != dstInfo.type then 108 | throw s!"Type clash of {name}" 109 | if srcInfo.levelParams != dstInfo.levelParams then 110 | throw s!"Level param clash of {name}" 111 | if !(← infoCompare srcInfo dstInfo) then 112 | throw s!"Definition clash of {name}" 113 | if !isNoncomputable src' name ∧ isNoncomputable dst name then 114 | throw s!"Must not modify computability on {name}" 115 | | _ => 116 | if dstInfo.isAxiom then 117 | throw s!"Adding axiom is not allowed: {name}" 118 | srcConstants := srcConstants.erase name 119 | let srcConstantsList := srcConstants.keys.filter (not ∘ isInternalSymbol) 120 | if !srcConstantsList.isEmpty then 121 | throw s!"{srcConstantsList} not accounted for" 122 | return env 123 | where 124 | isInternalSymbol (name: Name) : Bool := match name with 125 | | .str _ n => n == "_cstage1" || n == "_cstage2" 126 | | _ => false 127 | /-- Assumes type check has been done. -/ 128 | infoCompare (srcInfo dstInfo : ConstantInfo) : ExceptT String IO Bool := 129 | match srcInfo, dstInfo with 130 | | .axiomInfo _a1, .axiomInfo _a2 => return true 131 | | .defnInfo srcVal, .defnInfo dstVal => do 132 | if srcVal.safety != dstVal.safety then 133 | return false 134 | if srcVal.value.hasSorry then 135 | return true 136 | -- Value modification is prohibited otherwise 137 | return srcVal.value == dstVal.value 138 | | .thmInfo srcVal, .thmInfo dstVal => do 139 | if srcVal.value.hasSorry then 140 | return true 141 | -- Value modification is prohibited otherwise 142 | return srcVal.value == dstVal.value 143 | | .opaqueInfo _, .opaqueInfo _ => do 144 | return true 145 | | .quotInfo _, .quotInfo _ => return true 146 | | .inductInfo _, .inductInfo _ => return true 147 | | .ctorInfo _, .ctorInfo _ => return true 148 | | .recInfo _, .recInfo _ => return true 149 | | _, _ => return false 150 | 151 | /-- Add constants to the environment, renaming the constants if necessary -/ 152 | def replayConstantsRenaming (constants : Std.HashMap Name ConstantInfo) : CoreM (NameMap Name) := do 153 | let env ← getEnv 154 | let nameMap ← constants.foldM (init := .empty) λ acc name _ => do 155 | unless env.contains name do 156 | return acc 157 | let name' ← if let .some pref := getAuxLemmaPrefix? name then 158 | mkAuxDeclName (kind := .str .anonymous pref) 159 | else 160 | mkAuxDeclName (kind := name) 161 | return acc.insert name name' 162 | -- Remap constants 163 | let replaceConst (expr : Expr) : CoreM Expr := Core.transform expr λ 164 | | .const n levels => 165 | let n' := nameMap.getD n n 166 | return .done (.const n' levels) 167 | | e => 168 | return .continue e 169 | let constants ← constants.foldM (init := .emptyWithCapacity constants.size) λ acc name info => do 170 | let info' ← match info with 171 | | .axiomInfo val@{ name, type, .. } => 172 | pure <| .axiomInfo { 173 | val with 174 | name := nameMap.getD name name, 175 | type := ← replaceConst type, 176 | } 177 | | .defnInfo val@{ name, type, value, .. } => 178 | pure <| .defnInfo { 179 | val with 180 | name := nameMap.getD name name, 181 | type := ← replaceConst type, 182 | value := ← replaceConst value, 183 | } 184 | | .thmInfo val@{ name, type, value, .. } => 185 | pure <| .thmInfo { 186 | val with 187 | name := nameMap.getD name name, 188 | type := ← replaceConst type, 189 | value := ← replaceConst value, 190 | } 191 | | .opaqueInfo val@{ name, type, value, .. } => 192 | pure <| .opaqueInfo { 193 | val with 194 | name := nameMap.getD name name, 195 | type := ← replaceConst type, 196 | value := ← replaceConst value, 197 | } 198 | | .quotInfo val@{ name, type, .. } => 199 | pure <| .quotInfo { 200 | val with 201 | name := nameMap.getD name name, 202 | type := ← replaceConst type, 203 | } 204 | | .inductInfo val@{ name, type, all, ctors, .. } => 205 | pure <| .inductInfo { 206 | val with 207 | name := nameMap.getD name name, 208 | type := ← replaceConst type, 209 | all := all.map λ n => nameMap.getD n n, 210 | ctors := ctors.map λ n => nameMap.getD n n, 211 | } 212 | | .ctorInfo val@{ name, type, induct, .. } => 213 | pure <| .ctorInfo { 214 | val with 215 | name := nameMap.getD name name, 216 | type := ← replaceConst type, 217 | induct := nameMap.getD induct induct 218 | } 219 | | .recInfo val@{ name, type, all, .. } => 220 | pure <| .recInfo { 221 | val with 222 | name := nameMap.getD name name, 223 | type := ← replaceConst type, 224 | all := all.map λ n => nameMap.getD n n, 225 | } 226 | return acc.insert name info' 227 | let env' ← (← getEnv).replay constants 228 | setEnv env' 229 | return nameMap 230 | -------------------------------------------------------------------------------- /doc/repl.md: -------------------------------------------------------------------------------- 1 | # REPL 2 | 3 | This documentation is about interacting with the REPL. 4 | 5 | ## Examples 6 | 7 | After building the `repl`, it will be available in `.lake/build/bin/repl`. 8 | Execute it by either directly referring to its name, or `lake exe repl`. 9 | 10 | ``` sh 11 | repl MODULES|LEAN_OPTIONS 12 | ``` 13 | 14 | The `repl` executable must be given with a list of modules to import. By default 15 | it will import nothing, not even `Init`. It can also accept lean options of the 16 | form `--key=value` e.g. `--pp.raw=true`. 17 | 18 | Running repl with `--version` shows the version and then exits. 19 | 20 | After it emits the `ready.` signal, `repl` accepts commands as single-line JSON 21 | inputs and outputs either an `Error:` (indicating malformed command) or a JSON 22 | return value indicating the result of a command execution. The command must be 23 | given in one of two formats 24 | 25 | ``` 26 | command { ... } 27 | { "cmd": command, "payload": ... } 28 | ``` 29 | 30 | The list of available commands can be found below. An empty command aborts the 31 | REPL. 32 | 33 | Example: (~5k symbols) 34 | ``` 35 | $ repl Init 36 | env.catalog {} 37 | env.inspect {"name": "Nat.le_add_left"} 38 | ``` 39 | 40 | Example with `mathlib4` (~90k symbols, may stack overflow, see troubleshooting) 41 | 42 | ``` 43 | $ repl Mathlib.Analysis.Seminorm 44 | env.catalog {} 45 | ``` 46 | 47 | Example proving a theorem: (alternatively use `goal.start {"copyFrom": "Nat.add_comm"}`) 48 | to prime the proof 49 | 50 | ``` 51 | $ repl Init 52 | goal.start {"expr": "∀ (n m : Nat), n + m = m + n"} 53 | goal.tactic {"stateId": 0, "tactic": "intro n m"} 54 | goal.tactic {"stateId": 1, "tactic": "assumption"} 55 | goal.delete {"stateIds": [0]} 56 | stat {} 57 | goal.tactic {"stateId": 1, "tactic": "rw [Nat.add_comm]"} 58 | stat 59 | ``` 60 | where the application of `assumption` should lead to a failure. 61 | 62 | ### Project Environment 63 | 64 | To use Pantograph in a project environment, setup the `LEAN_PATH` environment 65 | variable so it contains the library path of lean libraries. The libraries must 66 | be built in advance. For example, if `mathlib4` is stored at `../lib/mathlib4`, 67 | the environment might be setup like this: 68 | 69 | ``` sh 70 | LIB="../lib" 71 | LIB_MATHLIB="$LIB/mathlib4/.lake" 72 | export LEAN_PATH="$LIB_MATHLIB:$LIB_MATHLIB/aesop/build/lib:$LIB_MATHLIB/Qq/build/lib:$LIB_MATHLIB/std/build/lib" 73 | 74 | LEAN_PATH=$LEAN_PATH repl $@ 75 | ``` 76 | The `$LEAN_PATH` executable of any project can be extracted by 77 | ``` sh 78 | lake env printenv LEAN_PATH 79 | ``` 80 | 81 | Additional modules cannot be imported after the perennial process starts, either 82 | via `env.load` or the frontend functions. The technical reason for this is when 83 | Lean cannot determine whether an imported module's initializer has run. 84 | 85 | ## Commands 86 | 87 | See `Pantograph/Protocol.lean` for a description of the parameters and return values in JSON. 88 | * `reset`: Delete all cached expressions and proof trees 89 | * `stat`: Display resource usage 90 | * `options.set { key: value, ... }`: Set one or more options. These are not Lean 91 | `CoreM` options; those have to be set via command line arguments.), for 92 | options see below. 93 | * `options.print`: Display the current set of options 94 | * `expr.echo {"expr": , "type": , ["levels": []]}`: Determine the 95 | type of an expression and format it. 96 | * `env.catalog`: Display a list of all safe Lean symbols in the current environment 97 | * `env.inspect {"name": , "value": }`: Show the type and package of a 98 | given symbol; If value flag is set, the value is printed or hidden. By default 99 | only the values of definitions are printed. 100 | * `env.save { "path": }`, `env.load { "path": }`: Save/Load the 101 | current environment to/from a file 102 | * `env.module_read { "module": }`: Reads a list of symbols from a module 103 | * `env.describe {}`: Describes the imports and modules in the current environment 104 | * `env.parse { "input": , "category": }`: Parse a bit 105 | of syntax and returns the parser's terminal position. 106 | * `goal.start {["name": ], ["expr": ], ["levels": []], ["copyFrom": ]}`: 107 | Start a new proof from a given expression or symbol 108 | * `goal.tactic {"stateId": , ["goalId": ], ["autoResume": ], ...}`: 109 | Execute a tactic string on a given goal site. The tactic is supplied as additional 110 | key-value pairs in one of the following formats: 111 | - `{ "tactic": }`: Executes a tactic or a sequence of tactics in the 112 | current mode. 113 | - `{ "mode": }`: Enter a different tactic mode. The permitted values 114 | are `tactic` (default), `conv`, `calc`. In case of `calc`, each step must 115 | be of the form `lhs op rhs`. An `lhs` of `_` indicates that it should be set 116 | to the previous `rhs`. 117 | - `{ "expr": }`: Assign the given proof term to the current goal 118 | - `{ "have": , "binderName": }`: Execute `have` and creates a branch goal 119 | - `{ "let": , "binderName": }`: Execute `let` and creates a branch goal 120 | - `{ "draft": }`: Draft an expression with `sorry`s, turning them into 121 | goals. Coupling is not allowed. 122 | If the `goals` field does not exist, the tactic execution has failed. Read 123 | `messages` to find the reason. 124 | * `goal.continue {"stateId": , ["branch": ], ["goals": ]}`: 125 | Execute continuation/resumption 126 | - `{ "branch": }`: Continue on branch state. The current state must have no goals. 127 | - `{ "goals": }`: Resume the given goals 128 | * `goal.subsume {"stateId": , "goal": , "candidates": 129 | , ["srcStateId": ]}`: determine if any goal in `candidates` (coming 130 | from either the provided state id or `srcStateId`) subsumes `goal`. It returns 131 | the *subsumptor* (goal providing the solution) and a new state id if the 132 | subsumption is not a cycle, in which case the *subsumend* `goal` is erased. 133 | * `goal.remove {"stateIds": []}"`: Drop the goal states specified in the list 134 | * `goal.print {"stateId": }"`: Print a goal state 135 | * `goal.save { "id": , "path": }`, `goal.load { "path": }`: 136 | Save/Load a goal state to/from a file. The environment is not carried with the 137 | state. The user is responsible to ensure the sender/receiver instances share 138 | the same environment. 139 | * `frontend.process { ["fileName": ,] ["file": ], readHeader: 140 | , inheritEnv: , invocations: , newConstants: }`: 141 | Executes the Lean frontend on a file, collecting the tactic invocations 142 | (`"invocations": output-path`), or new constants (`newConstants`) 143 | * `frontend.distil { "file": , ["binderName": ], "ignoreValues": bool 144 | }`: Extract condensed search targets from a file, where coupled search targets 145 | will be condensed into one. Set `binderName` to override the binder name to 146 | e.g. `f`. Set `ignoreValues` to false to incorporate existing solutions. 147 | 148 | Note that `example`s are not search targets! 149 | * `frontend.track { "src": , "dst": }`: Check if one file conforms to 150 | another. The declarations in `src` could have `sorry`s and the declarations in 151 | `dst` would fill them. 152 | * [Experimental] `frontend.refactor { "file": , "coreOptions": 153 | [["="]] }`: Group dependent `sorry`s into one single `sorry`. 154 | Currently only flat dependencies are supported (i.e. an object with a list of 155 | properties). 156 | 157 | ## Options 158 | 159 | The full list of options can be found in `Pantograph/Protocol.lean`. Particularly: 160 | - `automaticMode` (default on): Goals will not become dormant when this is 161 | turned on. By default it is turned on, with all goals automatically resuming. 162 | This makes Pantograph act like a gym, with no resumption necessary to manage 163 | your goals. 164 | - `timeout` (default 0): Set `timeout` to a non-zero number to specify timeout 165 | (milliseconds) for all `CoreM` and frontend operations. 166 | 167 | ## Errors 168 | 169 | When an error pertaining to the execution of a command happens, the returning JSON structure is 170 | 171 | ``` json 172 | { "error": "type", "desc": "description" } 173 | ``` 174 | Common error forms: 175 | * `command`: Indicates malformed command structure which results from either 176 | invalid command or a malformed JSON structure that cannot be fed to an 177 | individual command. 178 | * `index`: Indicates an invariant maintained by the output of one command and 179 | input of another is broken. For example, attempting to query a symbol not 180 | existing in the library or indexing into a non-existent proof state. 181 | * `parse`: Indicates parsing errors 182 | * `elab`: Indicates elaboration errors 183 | * `frontend`: Indicates whole-file parsing and elaboration errors 184 | * `io`: Generic IO error 185 | * `command`: The command's argument is malformed 186 | 187 | ## Troubleshooting 188 | 189 | If lean encounters stack overflow problems when printing catalog, execute this before running lean: 190 | ```sh 191 | ulimit -s unlimited 192 | ``` 193 | -------------------------------------------------------------------------------- /LICENSE: -------------------------------------------------------------------------------- 1 | Apache License 2 | Version 2.0, January 2004 3 | http://www.apache.org/licenses/ 4 | 5 | TERMS AND CONDITIONS FOR USE, REPRODUCTION, AND DISTRIBUTION 6 | 7 | 1. Definitions. 8 | 9 | "License" shall mean the terms and conditions for use, reproduction, 10 | and distribution as defined by Sections 1 through 9 of this document. 11 | 12 | "Licensor" shall mean the copyright owner or entity authorized by 13 | the copyright owner that is granting the License. 14 | 15 | "Legal Entity" shall mean the union of the acting entity and all 16 | other entities that control, are controlled by, or are under common 17 | control with that entity. For the purposes of this definition, 18 | "control" means (i) the power, direct or indirect, to cause the 19 | direction or management of such entity, whether by contract or 20 | otherwise, or (ii) ownership of fifty percent (50%) or more of the 21 | outstanding shares, or (iii) beneficial ownership of such entity. 22 | 23 | "You" (or "Your") shall mean an individual or Legal Entity 24 | exercising permissions granted by this License. 25 | 26 | "Source" form shall mean the preferred form for making modifications, 27 | including but not limited to software source code, documentation 28 | source, and configuration files. 29 | 30 | "Object" form shall mean any form resulting from mechanical 31 | transformation or translation of a Source form, including but 32 | not limited to compiled object code, generated documentation, 33 | and conversions to other media types. 34 | 35 | "Work" shall mean the work of authorship, whether in Source or 36 | Object form, made available under the License, as indicated by a 37 | copyright notice that is included in or attached to the work 38 | (an example is provided in the Appendix below). 39 | 40 | "Derivative Works" shall mean any work, whether in Source or Object 41 | form, that is based on (or derived from) the Work and for which the 42 | editorial revisions, annotations, elaborations, or other modifications 43 | represent, as a whole, an original work of authorship. For the purposes 44 | of this License, Derivative Works shall not include works that remain 45 | separable from, or merely link (or bind by name) to the interfaces of, 46 | the Work and Derivative Works thereof. 47 | 48 | "Contribution" shall mean any work of authorship, including 49 | the original version of the Work and any modifications or additions 50 | to that Work or Derivative Works thereof, that is intentionally 51 | submitted to Licensor for inclusion in the Work by the copyright owner 52 | or by an individual or Legal Entity authorized to submit on behalf of 53 | the copyright owner. For the purposes of this definition, "submitted" 54 | means any form of electronic, verbal, or written communication sent 55 | to the Licensor or its representatives, including but not limited to 56 | communication on electronic mailing lists, source code control systems, 57 | and issue tracking systems that are managed by, or on behalf of, the 58 | Licensor for the purpose of discussing and improving the Work, but 59 | excluding communication that is conspicuously marked or otherwise 60 | designated in writing by the copyright owner as "Not a Contribution." 61 | 62 | "Contributor" shall mean Licensor and any individual or Legal Entity 63 | on behalf of whom a Contribution has been received by Licensor and 64 | subsequently incorporated within the Work. 65 | 66 | 2. Grant of Copyright License. Subject to the terms and conditions of 67 | this License, each Contributor hereby grants to You a perpetual, 68 | worldwide, non-exclusive, no-charge, royalty-free, irrevocable 69 | copyright license to reproduce, prepare Derivative Works of, 70 | publicly display, publicly perform, sublicense, and distribute the 71 | Work and such Derivative Works in Source or Object form. 72 | 73 | 3. Grant of Patent License. Subject to the terms and conditions of 74 | this License, each Contributor hereby grants to You a perpetual, 75 | worldwide, non-exclusive, no-charge, royalty-free, irrevocable 76 | (except as stated in this section) patent license to make, have made, 77 | use, offer to sell, sell, import, and otherwise transfer the Work, 78 | where such license applies only to those patent claims licensable 79 | by such Contributor that are necessarily infringed by their 80 | Contribution(s) alone or by combination of their Contribution(s) 81 | with the Work to which such Contribution(s) was submitted. If You 82 | institute patent litigation against any entity (including a 83 | cross-claim or counterclaim in a lawsuit) alleging that the Work 84 | or a Contribution incorporated within the Work constitutes direct 85 | or contributory patent infringement, then any patent licenses 86 | granted to You under this License for that Work shall terminate 87 | as of the date such litigation is filed. 88 | 89 | 4. Redistribution. You may reproduce and distribute copies of the 90 | Work or Derivative Works thereof in any medium, with or without 91 | modifications, and in Source or Object form, provided that You 92 | meet the following conditions: 93 | 94 | (a) You must give any other recipients of the Work or 95 | Derivative Works a copy of this License; and 96 | 97 | (b) You must cause any modified files to carry prominent notices 98 | stating that You changed the files; and 99 | 100 | (c) You must retain, in the Source form of any Derivative Works 101 | that You distribute, all copyright, patent, trademark, and 102 | attribution notices from the Source form of the Work, 103 | excluding those notices that do not pertain to any part of 104 | the Derivative Works; and 105 | 106 | (d) If the Work includes a "NOTICE" text file as part of its 107 | distribution, then any Derivative Works that You distribute must 108 | include a readable copy of the attribution notices contained 109 | within such NOTICE file, excluding those notices that do not 110 | pertain to any part of the Derivative Works, in at least one 111 | of the following places: within a NOTICE text file distributed 112 | as part of the Derivative Works; within the Source form or 113 | documentation, if provided along with the Derivative Works; or, 114 | within a display generated by the Derivative Works, if and 115 | wherever such third-party notices normally appear. The contents 116 | of the NOTICE file are for informational purposes only and 117 | do not modify the License. You may add Your own attribution 118 | notices within Derivative Works that You distribute, alongside 119 | or as an addendum to the NOTICE text from the Work, provided 120 | that such additional attribution notices cannot be construed 121 | as modifying the License. 122 | 123 | You may add Your own copyright statement to Your modifications and 124 | may provide additional or different license terms and conditions 125 | for use, reproduction, or distribution of Your modifications, or 126 | for any such Derivative Works as a whole, provided Your use, 127 | reproduction, and distribution of the Work otherwise complies with 128 | the conditions stated in this License. 129 | 130 | 5. Submission of Contributions. Unless You explicitly state otherwise, 131 | any Contribution intentionally submitted for inclusion in the Work 132 | by You to the Licensor shall be under the terms and conditions of 133 | this License, without any additional terms or conditions. 134 | Notwithstanding the above, nothing herein shall supersede or modify 135 | the terms of any separate license agreement you may have executed 136 | with Licensor regarding such Contributions. 137 | 138 | 6. Trademarks. This License does not grant permission to use the trade 139 | names, trademarks, service marks, or product names of the Licensor, 140 | except as required for reasonable and customary use in describing the 141 | origin of the Work and reproducing the content of the NOTICE file. 142 | 143 | 7. Disclaimer of Warranty. Unless required by applicable law or 144 | agreed to in writing, Licensor provides the Work (and each 145 | Contributor provides its Contributions) on an "AS IS" BASIS, 146 | WITHOUT WARRANTIES OR CONDITIONS OF ANY KIND, either express or 147 | implied, including, without limitation, any warranties or conditions 148 | of TITLE, NON-INFRINGEMENT, MERCHANTABILITY, or FITNESS FOR A 149 | PARTICULAR PURPOSE. You are solely responsible for determining the 150 | appropriateness of using or redistributing the Work and assume any 151 | risks associated with Your exercise of permissions under this License. 152 | 153 | 8. Limitation of Liability. In no event and under no legal theory, 154 | whether in tort (including negligence), contract, or otherwise, 155 | unless required by applicable law (such as deliberate and grossly 156 | negligent acts) or agreed to in writing, shall any Contributor be 157 | liable to You for damages, including any direct, indirect, special, 158 | incidental, or consequential damages of any character arising as a 159 | result of this License or out of the use or inability to use the 160 | Work (including but not limited to damages for loss of goodwill, 161 | work stoppage, computer failure or malfunction, or any and all 162 | other commercial damages or losses), even if such Contributor 163 | has been advised of the possibility of such damages. 164 | 165 | 9. Accepting Warranty or Additional Liability. While redistributing 166 | the Work or Derivative Works thereof, You may choose to offer, 167 | and charge a fee for, acceptance of support, warranty, indemnity, 168 | or other liability obligations and/or rights consistent with this 169 | License. However, in accepting such obligations, You may act only 170 | on Your own behalf and on Your sole responsibility, not on behalf 171 | of any other Contributor, and only if You agree to indemnify, 172 | defend, and hold each Contributor harmless for any liability 173 | incurred by, or claims asserted against, such Contributor by reason 174 | of your accepting any such warranty or additional liability. 175 | 176 | END OF TERMS AND CONDITIONS 177 | 178 | Copyright 2024 Leni Aniva 179 | 180 | Licensed under the Apache License, Version 2.0 (the "License"); 181 | you may not use this file except in compliance with the License. 182 | You may obtain a copy of the License at 183 | 184 | http://www.apache.org/licenses/LICENSE-2.0 185 | 186 | Unless required by applicable law or agreed to in writing, software 187 | distributed under the License is distributed on an "AS IS" BASIS, 188 | WITHOUT WARRANTIES OR CONDITIONS OF ANY KIND, either express or implied. 189 | See the License for the specific language governing permissions and 190 | limitations under the License. 191 | -------------------------------------------------------------------------------- /Test/Tactic/Prograde.lean: -------------------------------------------------------------------------------- 1 | import Pantograph.Tactic.Prograde 2 | import Test.Common 3 | 4 | open Lean 5 | 6 | namespace Pantograph.Test.Tactic.Prograde 7 | 8 | def test_define : TestT Elab.TermElabM Unit := do 9 | let expr := "forall (p q : Prop) (h: p), And (Or p q) (Or p q)" 10 | let expr ← parseSentence expr 11 | Meta.forallTelescope expr $ λ _ body => do 12 | let e ← match Parser.runParserCategory 13 | (env := ← MonadEnv.getEnv) 14 | (catName := `term) 15 | (input := "Or.inl h") 16 | (fileName := ← getFileName) with 17 | | .ok syn => pure syn 18 | | .error error => throwError "Failed to parse: {error}" 19 | -- Apply the tactic 20 | let goal ← Meta.mkFreshExprSyntheticOpaqueMVar body 21 | let target: Expr := mkAnd 22 | (mkOr (.fvar ⟨uniq 8⟩) (.fvar ⟨uniq 9⟩)) 23 | (mkOr (.fvar ⟨uniq 8⟩) (.fvar ⟨uniq 9⟩)) 24 | let h := .fvar ⟨uniq 8⟩ 25 | addTest $ LSpec.test "goals before" ((← toCondensedGoal goal.mvarId!).devolatilize == { 26 | context := #[ 27 | cdeclOf `p (.sort 0), 28 | cdeclOf `q (.sort 0), 29 | cdeclOf `h h 30 | ], 31 | target, 32 | }) 33 | let tactic := Tactic.evalDefine `h2 e 34 | let m := .mvar ⟨uniq 13⟩ 35 | let [newGoal] ← runTacticOnMVar tactic goal.mvarId! | panic! "Incorrect goal number" 36 | addTest $ LSpec.test "goals after" ((← toCondensedGoal newGoal).devolatilize == { 37 | context := #[ 38 | cdeclOf `p (.sort 0), 39 | cdeclOf `q (.sort 0), 40 | cdeclOf `h h, 41 | { 42 | userName := `h2, 43 | type := mkOr h m, 44 | value? := .some $ mkApp3 (mkConst `Or.inl) h m (.fvar ⟨uniq 10⟩) 45 | } 46 | ], 47 | target, 48 | }) 49 | let .some e ← getExprMVarAssignment? goal.mvarId! | panic! "Tactic must assign" 50 | addTest $ LSpec.test "assign" e.isLet 51 | 52 | def test_define_proof : TestT Elab.TermElabM Unit := do 53 | let rootExpr ← parseSentence "∀ (p q: Prop), p → ((p ∨ q) ∨ (p ∨ q))" 54 | let state0 ← GoalState.create rootExpr 55 | let tactic := "intro p q h" 56 | let state1 ← match ← state0.tacticOn (goalId := 0) (tactic := tactic) with 57 | | .success state _ => pure state 58 | | other => do 59 | addTest $ assertUnreachable $ other.toString 60 | return () 61 | checkEq tactic ((← state1.serializeGoals).map (·.devolatilize)) 62 | #[buildGoal [(`p, "Prop"), (`q, "Prop"), (`h, "p")] "(p ∨ q) ∨ p ∨ q"] 63 | 64 | let expr := "Or.inl (Or.inl h)" 65 | let state2 ← match ← state1.tryAssign (state1.get! 0) (expr := expr) with 66 | | .success state _ => pure state 67 | | other => do 68 | addTest $ assertUnreachable $ other.toString 69 | return () 70 | addTest $ LSpec.check s!":= {expr}" ((← state2.serializeGoals).map (·.devolatilize) = 71 | #[]) 72 | 73 | let evalBind := `y 74 | let evalExpr := "Or.inl h" 75 | let state2 ← match ← state1.tryDefine (state1.get! 0) (binderName := evalBind) (expr := evalExpr) with 76 | | .success state _ => pure state 77 | | other => do 78 | addTest $ assertUnreachable $ other.toString 79 | return () 80 | checkEq s!"eval {evalBind} := {evalExpr}" ((← state2.serializeGoals).map (·.devolatilize)) 81 | #[{ 82 | target := { pp? := .some "(p ∨ q) ∨ p ∨ q"}, 83 | vars := #[ 84 | { userName := `p, type? := .some { pp? := .some "Prop" } }, 85 | { userName := `q, type? := .some { pp? := .some "Prop" } }, 86 | { userName := `h, type? := .some { pp? := .some "p" } }, 87 | { userName := `y, 88 | type? := .some { pp? := .some "p ∨ ?m.9" }, 89 | value? := .some { pp? := .some "Or.inl h" }, 90 | } 91 | ] 92 | }] 93 | 94 | let expr := "Or.inl y" 95 | let state3 ← match ← state2.tryAssign (state2.get! 0) (expr := expr) with 96 | | .success state _ => pure state 97 | | other => do 98 | addTest $ assertUnreachable $ other.toString 99 | return () 100 | addTest $ LSpec.check s!":= {expr}" ((← state3.serializeGoals).map (·.devolatilize) = 101 | #[]) 102 | 103 | addTest $ LSpec.check "(3 root)" state3.rootExpr?.isSome 104 | 105 | def fun_define_root_expr: ∀ (p: Prop), PProd (Nat → p) Unit → p := by 106 | intro p x 107 | apply x.fst 108 | exact 5 109 | 110 | def test_define_root_expr : TestT Elab.TermElabM Unit := do 111 | --let rootExpr ← parseSentence "Nat" 112 | --let state0 ← GoalState.create rootExpr 113 | --let .success state1 _ ← state0.tacticOn (goalId := 0) "exact 5" | addTest $ assertUnreachable "exact 5" 114 | --let .some rootExpr := state1.rootExpr? | addTest $ assertUnreachable "Root expr" 115 | --addTest $ LSpec.check "root" ((toString $ ← Meta.ppExpr rootExpr) = "5") 116 | let rootExpr ← parseSentence "∀ (p: Prop), PProd (Nat → p) Unit → p" 117 | let state0 ← GoalState.create rootExpr 118 | let tactic := "intro p x" 119 | let .success state1 _ ← state0.tacticOn (goalId := 0) tactic | addTest $ assertUnreachable tactic 120 | let binderName := `binder 121 | let value := "x.fst" 122 | let expr ← state1.goals[0]!.withContext $ strToTermSyntax value 123 | let tacticM := Tactic.evalDefine binderName expr 124 | let .success state2 _ ← state1.tryTacticM (state1.get! 0) tacticM | addTest $ assertUnreachable s!"define {binderName} := {value}" 125 | let tactic := s!"apply {binderName}" 126 | let .success state3 _ ← state2.tacticOn (goalId := 0) tactic | addTest $ assertUnreachable tactic 127 | let tactic := s!"exact 5" 128 | let .success state4 _ ← state3.tacticOn (goalId := 0) tactic | addTest $ assertUnreachable tactic 129 | let .some rootExpr := state4.rootExpr? | addTest $ assertUnreachable "Root expr" 130 | addTest $ LSpec.check "root" ((toString $ ← Meta.ppExpr rootExpr) = "fun p x =>\n let binder := x.fst;\n binder 5") 131 | 132 | def test_have_proof : TestT Elab.TermElabM Unit := do 133 | let rootExpr ← parseSentence "∀ (p q: Prop), p → ((p ∨ q) ∨ (p ∨ q))" 134 | let state0 ← GoalState.create rootExpr 135 | let tactic := "intro p q h" 136 | let state1 ← match ← state0.tacticOn (goalId := 0) (tactic := tactic) with 137 | | .success state _ => pure state 138 | | other => do 139 | addTest $ assertUnreachable $ other.toString 140 | return () 141 | checkEq tactic ((← state1.serializeGoals).map (·.devolatilize)) 142 | #[buildGoal [(`p, "Prop"), (`q, "Prop"), (`h, "p")] "(p ∨ q) ∨ p ∨ q"] 143 | 144 | let expr := "Or.inl (Or.inl h)" 145 | let state2 ← match ← state1.tryAssign (state1.get! 0) (expr := expr) with 146 | | .success state _ => pure state 147 | | other => do 148 | addTest $ assertUnreachable $ other.toString 149 | return () 150 | addTest $ LSpec.check s!":= {expr}" ((← state2.serializeGoals).map (·.devolatilize) = 151 | #[]) 152 | 153 | let haveBind := `y 154 | let haveType := "p ∨ q" 155 | let state2 ← match ← state1.tryHave (state1.get! 0) (binderName := haveBind) (type := haveType) with 156 | | .success state _ => pure state 157 | | other => do 158 | addTest $ assertUnreachable $ other.toString 159 | return () 160 | checkEq s!"have {haveBind}: {haveType}" ((← state2.serializeGoals).map (·.devolatilize)) 161 | #[ 162 | buildGoal [(`p, "Prop"), (`q, "Prop"), (`h, "p")] "p ∨ q", 163 | buildGoal [(`p, "Prop"), (`q, "Prop"), (`h, "p"), (`y, "p ∨ q")] "(p ∨ q) ∨ p ∨ q" 164 | ] 165 | 166 | let expr := "Or.inl h" 167 | let state3 ← match ← state2.tryAssign (state2.get! 0) (expr := expr) with 168 | | .success state _ => pure state 169 | | other => do 170 | addTest $ assertUnreachable $ other.toString 171 | return () 172 | checkEq s!":= {expr}" ((← state3.serializeGoals).map (·.devolatilize)) #[] 173 | 174 | let state2b ← match state3.continue state2 with 175 | | .ok state => pure state 176 | | .error e => do 177 | addTest $ assertUnreachable e 178 | return () 179 | let expr := "Or.inl y" 180 | let state4 ← match ← state2b.tryAssign (state2b.get! 0) (expr := expr) with 181 | | .success state _ => pure state 182 | | other => do 183 | addTest $ assertUnreachable $ other.toString 184 | return () 185 | checkEq s!":= {expr}" ((← state4.serializeGoals).map (·.devolatilize)) #[] 186 | 187 | let .some rootExpr := state4.rootExpr? | fail "Root expr" 188 | checkEq "root" (toString $ ← Meta.ppExpr rootExpr) "fun p q h y => Or.inl y" 189 | 190 | def test_let (specialized: Bool): TestT Elab.TermElabM Unit := do 191 | let rootExpr ← parseSentence "∀ (p q: Prop), p → ((p ∨ q) ∨ (p ∨ q))" 192 | let state0 ← GoalState.create rootExpr 193 | let tactic := "intro a p h" 194 | let state1 ← match ← state0.tacticOn (goalId := 0) (tactic := tactic) with 195 | | .success state _ => pure state 196 | | other => do 197 | addTest $ assertUnreachable $ other.toString 198 | return () 199 | addTest $ LSpec.check tactic ((← state1.serializeGoals).map (·.devolatilize) = 200 | #[{ 201 | target := { pp? := .some mainTarget }, 202 | vars := interiorVars, 203 | }]) 204 | 205 | let letType := "Nat" 206 | let expr := s!"let b: {letType} := _; _" 207 | let result2 ← match specialized with 208 | | true => state1.tryLet (state1.get! 0) (binderName := `b) (type := letType) 209 | | false => state1.tryAssign (state1.get! 0) (expr := expr) 210 | let state2 ← match result2 with 211 | | .success state _ => pure state 212 | | other => do 213 | addTest $ assertUnreachable $ other.toString 214 | return () 215 | let serializedState2 ← state2.serializeGoals 216 | let letBindName := if specialized then `b else `_1 217 | checkEq expr (serializedState2.map (·.devolatilize)) 218 | #[{ 219 | target := { pp? := .some letType }, 220 | vars := interiorVars, 221 | userName? := .some letBindName 222 | }, 223 | { 224 | target := { pp? := .some mainTarget }, 225 | vars := interiorVars ++ #[{ 226 | userName := `b, 227 | type? := .some { pp? := .some letType }, 228 | value? := .some { pp? := .some s!"?{letBindName}" }, 229 | }], 230 | userName? := if specialized then .none else .some `_2, 231 | }] 232 | 233 | let tactic := "exact 1" 234 | let state3 ← match ← state2.tacticOn (goalId := 0) (tactic := tactic) with 235 | | .success state _ => pure state 236 | | other => do 237 | addTest $ assertUnreachable $ other.toString 238 | return () 239 | addTest $ LSpec.check tactic ((← state3.serializeGoals).map (·.devolatilize) = #[]) 240 | 241 | let state3r ← match state3.continue state2 with 242 | | .error msg => do 243 | addTest $ assertUnreachable $ msg 244 | return () 245 | | .ok state => pure state 246 | checkEq "(continue)" ((← state3r.serializeGoals).map (·.devolatilize)) 247 | #[{ 248 | target := { pp? := .some mainTarget }, 249 | vars := interiorVars ++ #[{ 250 | userName := `b, 251 | type? := .some { pp? := .some "Nat" }, 252 | value? := .some { pp? := .some "1" }, 253 | }], 254 | userName? := if specialized then .none else `_2, 255 | }] 256 | 257 | let tactic := "exact h" 258 | match ← state3r.tacticOn (goalId := 0) (tactic := tactic) with 259 | | .failure #[message] => 260 | checkEq tactic 261 | (← message.toString) 262 | s!"{← getFileName}:0:0: error: Type mismatch\n h\nhas type\n a\nbut is expected to have type\n {mainTarget}\n" 263 | | other => do 264 | fail s!"Should be a failure: {other.toString}" 265 | 266 | let tactic := "exact Or.inl (Or.inl h)" 267 | let state4 ← match ← state3r.tacticOn (goalId := 0) (tactic := tactic) with 268 | | .success state _ => pure state 269 | | other => do 270 | addTest $ assertUnreachable $ other.toString 271 | return () 272 | addTest $ LSpec.test "(4 root)" state4.rootExpr?.isSome 273 | where 274 | mainTarget := "(a ∨ p) ∨ a ∨ p" 275 | interiorVars: Array Protocol.Variable := #[ 276 | { userName := `a, type? := .some { pp? := .some "Prop" }, }, 277 | { userName := `p, type? := .some { pp? := .some "Prop" }, }, 278 | { userName := `h, type? := .some { pp? := .some "a" }, } 279 | ] 280 | 281 | def suite (env: Environment): List (String × IO LSpec.TestSeq) := 282 | [ 283 | ("define", test_define), 284 | ("define proof", test_define_proof), 285 | ("define root expr", test_define_root_expr), 286 | ("have proof", test_have_proof), 287 | ("let via assign", test_let false), 288 | ("let via tryLet", test_let true), 289 | ] |>.map (λ (name, t) => (name, runTestTermElabM env t)) 290 | -------------------------------------------------------------------------------- /Pantograph/Frontend/Distil.lean: -------------------------------------------------------------------------------- 1 | import Pantograph.Frontend.InfoTree 2 | import Pantograph.Frontend.MetaTranslate 3 | import Pantograph.Frontend.Refactor 4 | import Pantograph.Goal 5 | import Pantograph.Protocol 6 | 7 | open Lean 8 | 9 | namespace Pantograph.Frontend 10 | 11 | -- Info tree filtering functions 12 | 13 | /- Adapted from lean-training-data -/ 14 | structure TacticInvocation where 15 | info : Elab.TacticInfo 16 | ctx : Elab.ContextInfo 17 | children : PersistentArray Elab.InfoTree 18 | namespace TacticInvocation 19 | 20 | /-- Return the range of the tactic, as a pair of file positions. -/ 21 | @[export pantograph_frontend_tactic_invocation_range] 22 | protected def range (t : TacticInvocation) : Position × Position := t.ctx.fileMap.stxRange t.info.stx 23 | 24 | /-- Pretty print a tactic. -/ 25 | protected def pp (t : TacticInvocation) : IO Format := 26 | t.ctx.runMetaM {} try 27 | Lean.PrettyPrinter.ppTactic ⟨t.info.stx⟩ 28 | catch _ => 29 | pure "" 30 | 31 | /-- Run a tactic on the goals stored in a `TacticInvocation`. -/ 32 | protected def runMetaMGoalsBefore (t : TacticInvocation) (x : List MVarId → MetaM α) : IO α := do 33 | t.ctx.runMetaM {} <| Meta.withMCtx t.info.mctxBefore <| x t.info.goalsBefore 34 | 35 | /-- Run a tactic on the after goals stored in a `TacticInvocation`. -/ 36 | protected def runMetaMGoalsAfter (t : TacticInvocation) (x : List MVarId → MetaM α) : IO α := do 37 | t.ctx.runMetaM {} <| Meta.withMCtx t.info.mctxAfter <| x t.info.goalsAfter 38 | 39 | /-- Run a tactic on the main goal stored in a `TacticInvocation`. -/ 40 | protected def runMetaM (t : TacticInvocation) (x : MVarId → MetaM α) : IO α := do 41 | match t.info.goalsBefore.head? with 42 | | none => throw <| IO.userError s!"No goals at {← t.pp}" 43 | | some g => t.runMetaMGoalsBefore fun _ => do g.withContext <| x g 44 | 45 | protected def goalState (t : TacticInvocation) : IO (List Format) := do 46 | t.runMetaMGoalsBefore (fun gs => gs.mapM fun g => do Meta.ppGoal g) 47 | 48 | protected def goalStateAfter (t : TacticInvocation) : IO (List Format) := do 49 | t.runMetaMGoalsAfter (fun gs => gs.mapM fun g => do Meta.ppGoal g) 50 | 51 | protected def usedConstants (t: TacticInvocation) : NameSet := 52 | let info := t.info 53 | info.goalsBefore 54 | |>.filterMap info.mctxAfter.getExprAssignmentCore? 55 | |>.map Expr.getUsedConstantsAsSet 56 | |>.foldl .append .empty 57 | 58 | end TacticInvocation 59 | 60 | /-- Return all `TacticInfo` nodes in an `InfoTree` corresponding to tactics, 61 | each equipped with its relevant `ContextInfo`, and any children info trees. -/ 62 | private def collectTacticNodes (t : Elab.InfoTree) : List TacticInvocation := 63 | let infos := t.findAllInfo none false fun i => match i with 64 | | .ofTacticInfo _ => true 65 | | _ => false 66 | infos.filterMap fun p => match p with 67 | | (.ofTacticInfo i, some ctx, children) => .some ⟨i, ctx, children⟩ 68 | | _ => none 69 | 70 | def collectTactics (t : Elab.InfoTree) : List TacticInvocation := 71 | collectTacticNodes t |>.filter fun i => i.info.isSubstantive 72 | 73 | @[export pantograph_frontend_collect_tactics_from_compilation_step_m] 74 | def collectTacticsFromCompilationStep (step : CompilationStep) : IO (List Protocol.InvokedTactic) := do 75 | let tacticInfoTrees := step.trees.flatMap λ tree => tree.filter λ 76 | | info@(.ofTacticInfo _) => info.isOriginal 77 | | _ => false 78 | let tactics := tacticInfoTrees.flatMap collectTactics 79 | tactics.mapM λ invocation => do 80 | let goalBefore := (Format.joinSep (← invocation.goalState) "\n").pretty 81 | let goalAfter := (Format.joinSep (← invocation.goalStateAfter) "\n").pretty 82 | let tactic ← invocation.ctx.runMetaM {} <| Meta.withMCtx invocation.info.mctxBefore do 83 | return (← invocation.ctx.ppSyntax {} invocation.info.stx).pretty 84 | -- FIXME: Why does this not work? There are problems with `term.pseudo.antiquot` 85 | --PrettyPrinter.ppTactic ⟨invocation.info.stx⟩ 86 | --return t.pretty 87 | let usedConstants := invocation.usedConstants.toArray 88 | return { 89 | goalBefore, 90 | goalAfter, 91 | tactic, 92 | usedConstants, 93 | } 94 | 95 | structure InfoWithContext where 96 | info: Elab.Info 97 | context?: Option Elab.ContextInfo := .none 98 | 99 | structure GoalCollectionOptions where 100 | collectTypeErrors : Bool := false 101 | 102 | private def collectSorrysInTree (t : Elab.InfoTree) (options : GoalCollectionOptions := {}) 103 | : IO (List InfoWithContext) := do 104 | let infos ← t.findAllInfoM none fun i ctx? => match i with 105 | | .ofTermInfo { expectedType?, expr, stx, lctx, isBinder := false, .. } => do 106 | let .some ctx := ctx? | return (false, true) 107 | if expr.isSorry ∧ stx.isOfKind `Lean.Parser.Term.sorry then 108 | if expectedType?.isNone then 109 | throw $ .userError "Sorry of indeterminant type is not allowed" 110 | return (true, false) 111 | unless options.collectTypeErrors do 112 | return (false, true) 113 | let .some expectedType := expectedType? | return (false, true) 114 | let typeMatch ← ctx.runMetaM lctx do 115 | let type ← Meta.inferType expr 116 | Meta.isExprDefEqGuarded type expectedType 117 | return match typeMatch, expr.hasSorry with 118 | | false, true => (true, false) -- Types mismatch but has sorry -> collect, halt 119 | | false, false => (true, false) -- Types mistmatch but no sorry -> collect, halt 120 | | true, true => (false, true) -- Types match but has sorry -> continue 121 | | true, false => (false, false) -- Types match but no sorries -> halt 122 | | .ofTacticInfo { stx, goalsBefore, .. } => 123 | -- The `sorry` term is distinct from the `sorry` tactic 124 | let isSorry := stx.isOfKind `Lean.Parser.Tactic.tacticSorry 125 | return (isSorry ∧ !goalsBefore.isEmpty, ¬ isSorry) 126 | | _ => return (false, true) 127 | return infos.map fun (info, context?, _) => { info, context? } 128 | 129 | -- NOTE: Plural deliberately not spelled "sorries" 130 | @[export pantograph_frontend_collect_sorrys_m] 131 | def collectSorrys (step: CompilationStep) (options : GoalCollectionOptions := {}) 132 | : IO (List InfoWithContext) := do 133 | return (← step.trees.mapM $ λ tree => collectSorrysInTree tree options).flatten 134 | 135 | structure AnnotatedGoalState where 136 | state : GoalState 137 | srcBoundaries : List (String.Pos.Raw × String.Pos.Raw) 138 | 139 | /-- 140 | Since we cannot directly merge `MetavarContext`s, we have to get creative. This 141 | function duplicates frozen mvars in term and tactic info nodes, and add them to 142 | the current `MetavarContext`. 143 | 144 | DEPRECATED: Behaviour is unstable when there are multiple `sorry`s. Consider using 145 | the draft tactic instead. 146 | -/ 147 | @[export pantograph_frontend_sorrys_to_goal_state_m] 148 | def sorrysToGoalState (sorrys : List InfoWithContext) : MetaM AnnotatedGoalState := do 149 | assert! !sorrys.isEmpty 150 | let goalsM := sorrys.mapM λ i => do 151 | match i.info with 152 | | .ofTermInfo termInfo => do 153 | let mvarId ← MetaTranslate.translateMVarFromTermInfo termInfo i.context? 154 | if (← mvarId.getType).hasSorry then 155 | throwError s!"Coupling is not allowed in drafting" 156 | return [(mvarId, stxByteRange termInfo.stx)] 157 | | .ofTacticInfo tacticInfo => do 158 | let mvarIds ← MetaTranslate.translateMVarFromTacticInfoBefore tacticInfo i.context? 159 | for mvarId in mvarIds do 160 | if (← mvarId.getType).hasSorry then 161 | throwError s!"Coupling is not allowed in drafting" 162 | let range := stxByteRange tacticInfo.stx 163 | return mvarIds.map (·, range) 164 | | _ => panic! "Invalid info" 165 | let annotatedGoals := List.flatten (← goalsM.run {} |>.run' {}) 166 | let goals := annotatedGoals.map Prod.fst 167 | let srcBoundaries := annotatedGoals.map Prod.snd 168 | let root := match goals with 169 | | [] => panic! "No MVars generated" 170 | | [g] => g 171 | | _ => { name := .anonymous } 172 | let state ← GoalState.createFromMVars goals root 173 | return { state, srcBoundaries } 174 | 175 | structure DistilConfig where 176 | binderName? : Option Name := .none 177 | ignoreValues : Bool := true 178 | structure DistilledSearchTarget where 179 | goalState : GoalState 180 | 181 | def distilGoalStateFrom (head : Refactor.Command) (tail : List Refactor.Command) (config : DistilConfig) 182 | : RefactorM DistilledSearchTarget := do 183 | if head.constants.isEmpty then 184 | throw $ .userError "No constants in head declaration" 185 | let headName := head.constants.toList.head! 186 | let binderName := match config.binderName?, headName with 187 | | .some n, _ => n 188 | | _, .str _ binderName => Name.mkSimple binderName 189 | | _, _ => `x 190 | Refactor.distilSearchTarget head tail λ (witness, witnessValue) companions => do 191 | if companions.isEmpty then 192 | Meta.check witness 193 | -- Without companions, we can directly construct a goal state 194 | let goalState ← GoalState.create witness 195 | let goalState ← if !config.ignoreValues then 196 | goalState.step .unfocus do 197 | let goal ← Elab.Tactic.getMainGoal 198 | let (value, witnessGoals) ← Tactic.sorryToHole witnessValue |>.run [] 199 | Meta.check value 200 | goal.assign value 201 | Elab.Tactic.setGoals witnessGoals 202 | else 203 | pure goalState 204 | return { goalState } 205 | else 206 | let companion ← Meta.withLocalDeclD binderName witness λ binder => do 207 | let companion ← Refactor.mkProdElem ``And <| companions.map (·.fst.instantiate1 binder) 208 | Meta.mkLambdaFVars #[binder] companion 209 | let target ← Meta.mkAppOptM ``Subtype #[witness, companion] 210 | let goalState ← GoalState.create target 211 | let goalState ← if !config.ignoreValues && (!witnessValue.isSorry || companions.any (!·.snd.isSorry)) then 212 | goalState.step .unfocus do 213 | let goal ← Elab.Tactic.getMainGoal 214 | -- Construct the solution expression 215 | let (witnessValue', witnessGoals) ← Tactic.sorryToHole witnessValue |>.run [] 216 | assert! !witnessValue'.hasSorry 217 | let companionValue ← Meta.withLocalDeclD binderName witness λ binder => do 218 | let v ← Refactor.mkProdElem ``And.intro 219 | <| companions.map λ (_, value) => value.instantiate1 binder 220 | Meta.mkLambdaFVars #[binder] v 221 | let (companionValue', companionGoals) ← Tactic.sorryToHole 222 | (companionValue.beta #[witnessValue']) |>.run [] 223 | if h : witnessGoals.length = 1 then 224 | let witnessGoal := witnessGoals[0] 225 | witnessGoal.setTag binderName 226 | let value ← Meta.mkAppOptM ``Subtype.mk #[ 227 | witness, companion, 228 | witnessValue', 229 | companionValue', 230 | ] 231 | Meta.check value 232 | goal.assign value 233 | Elab.Tactic.setGoals (witnessGoals ++ companionGoals.reverse) 234 | else 235 | pure goalState 236 | return { goalState } 237 | 238 | open Refactor in 239 | @[export pantograph_distil_search_targets_m] 240 | def distilSearchTargets (env : Environment) (source : String) (config : DistilConfig := {}) (fileName : String := defaultFileName) 241 | : IO (List DistilledSearchTarget) := do 242 | let (fContext, fState) ← createContextStateFromFile source fileName env {} 243 | let commands ← Refactor.preprocess.run {} |>.run fContext |>.run' fState 244 | let errors := commands.filter (·.hasError) 245 | if let .some error := errors.head? then 246 | let message ← error.messages.mapM (·.toString) 247 | throw $ IO.userError $ "\n".intercalate message 248 | let m : RefactorM (List DistilledSearchTarget) := do 249 | let mut targets := [] 250 | while !(← get).commands.isEmpty do 251 | let { commands, .. } ← get 252 | let decl :: commands := commands 253 | | Refactor.fail "No commands left" 254 | modify ({ · with commands }) -- Prevents infinite loop 255 | let isSearchTarget ← decl.runCoreM do 256 | if decl.constants.isEmpty then 257 | return false 258 | let name := decl.constants.toList.head! 259 | let info := (← getEnv).find? name |>.get! 260 | let .some value := info.value? | return false 261 | return value.hasSorry 262 | if !isSearchTarget then 263 | pushNewCommand' (⟨decl.stx⟩ : Syntax.Command) 264 | continue 265 | let depstr ← extractDependencyStructure decl commands 266 | modify ({ · with commands := depstr.tail }) 267 | for command in depstr.intercalating do 268 | pushNewCommand' (⟨command.stx⟩ : Syntax.Command) 269 | let searchTarget ← distilGoalStateFrom decl depstr.component config 270 | targets := targets ++ [searchTarget] 271 | pure targets 272 | let outContext := { 273 | fContext.inputCtx with 274 | inputString := "", 275 | fileMap := "".toFileMap, 276 | endPos := "".rawEndPos, 277 | endPos_valid := by simp, 278 | } 279 | let parserState := {} 280 | let outState := { 281 | commandState := Elab.Command.mkState env {} {}, 282 | parserState, 283 | cmdPos := parserState.pos, 284 | } 285 | m.run { inContext := fContext.inputCtx } 286 | |>.run' { outContext, outState, commands } 287 | -------------------------------------------------------------------------------- /Test/Tactic/Fragment.lean: -------------------------------------------------------------------------------- 1 | import Pantograph.Goal 2 | import Test.Common 3 | 4 | open Lean 5 | 6 | namespace Pantograph.Test.Tactic.Fragment 7 | 8 | private def buildGoal (nameType: List (Name × String)) (target: String): 9 | Protocol.Goal := 10 | { 11 | target := { pp? := .some target}, 12 | vars := (nameType.map fun x => ({ 13 | userName := x.fst, 14 | type? := .some { pp? := .some x.snd }, 15 | })).toArray 16 | } 17 | 18 | abbrev TestM := TestT $ Elab.TermElabM 19 | 20 | example : ∀ (a b c1 c2: Nat), (b + a) + c1 = (b + a) + c2 → (a + b) + c1 = (b + a) + c2 := by 21 | intro a b c1 c2 h 22 | conv => 23 | lhs 24 | congr 25 | . rw [Nat.add_comm] 26 | . rfl 27 | exact h 28 | 29 | def test_conv_simple: TestM Unit := do 30 | let rootTarget ← parseSentence "∀ (a b c1 c2: Nat), (b + a) + c1 = (b + a) + c2 → (a + b) + c1 = (b + a) + c2" 31 | let state0 ← GoalState.create rootTarget 32 | 33 | let tactic := "intro a b c1 c2 h" 34 | let state1 ← match ← state0.tacticOn (goalId := 0) (tactic := tactic) with 35 | | .success state _ => pure state 36 | | other => do 37 | addTest $ assertUnreachable $ other.toString 38 | return () 39 | addTest $ LSpec.check tactic ((← state1.serializeGoals).map (·.devolatilize) = 40 | #[interiorGoal [] "a + b + c1 = b + a + c2"]) 41 | 42 | let goalConv := state1.goals[0]! 43 | let state2 ← match ← state1.convEnter (state1.get! 0) with 44 | | .success state _ => pure state 45 | | other => do 46 | addTest $ assertUnreachable $ other.toString 47 | return () 48 | addTest $ LSpec.check "conv => ..." ((← state2.serializeGoals).map (·.devolatilize) = 49 | #[{ interiorGoal [] "a + b + c1 = b + a + c2" with fragment := .conv }]) 50 | 51 | let convTactic := "rhs" 52 | let state3R ← match ← state2.tacticOn (goalId := 0) convTactic with 53 | | .success state _ => pure state 54 | | other => do 55 | addTest $ assertUnreachable $ other.toString 56 | return () 57 | addTest $ LSpec.check s!" {convTactic} (discard)" ((← state3R.serializeGoals).map (·.devolatilize) = 58 | #[{ interiorGoal [] "b + a + c2" with fragment := .conv }]) 59 | 60 | let convTactic := "lhs" 61 | let state3L ← match ← state2.tacticOn (goalId := 0) convTactic with 62 | | .success state _ => pure state 63 | | other => do 64 | addTest $ assertUnreachable $ other.toString 65 | return () 66 | addTest $ LSpec.check s!" {convTactic}" ((← state3L.serializeGoals).map (·.devolatilize) = 67 | #[{ interiorGoal [] "a + b + c1" with fragment := .conv }]) 68 | 69 | let convTactic := "congr" 70 | let state4 ← match ← state3L.tacticOn (goalId := 0) convTactic with 71 | | .success state _ => pure state 72 | | other => do 73 | addTest $ assertUnreachable $ other.toString 74 | return () 75 | addTest $ LSpec.check s!" {convTactic}" ((← state4.serializeGoals).map (·.devolatilize) = 76 | #[ 77 | { interiorGoal [] "a + b" with fragment := .conv, userName? := `a }, 78 | { interiorGoal [] "c1" with fragment := .conv, userName? := `a } 79 | ]) 80 | 81 | let convTactic := "rw [Nat.add_comm]" 82 | let state5_1 ← match ← state4.tacticOn (goalId := 0) convTactic with 83 | | .success state _ => pure state 84 | | other => do 85 | addTest $ assertUnreachable $ other.toString 86 | return () 87 | addTest $ LSpec.check s!" · {convTactic}" ((← state5_1.serializeGoals).map (·.devolatilize) = 88 | #[{ interiorGoal [] "b + a" with fragment := .conv, userName? := `a }]) 89 | 90 | let convTactic := "rfl" 91 | let state6_1 ← match ← state5_1.tacticOn (goalId := 0) convTactic with 92 | | .success state _ => pure state 93 | | other => do 94 | addTest $ assertUnreachable $ other.toString 95 | return () 96 | addTest $ LSpec.check s!" {convTactic}" ((← state6_1.serializeGoals).map (·.devolatilize) = 97 | #[]) 98 | 99 | let state4_1 ← match state6_1.continue state4 with 100 | | .ok state => pure state 101 | | .error e => do 102 | addTest $ expectationFailure "continue" e 103 | return () 104 | 105 | let convTactic := "rfl" 106 | let state1_1 ← match ← state4_1.tacticOn (goalId := 0) convTactic with 107 | | .success state _ => pure state 108 | | other => do 109 | addTest $ assertUnreachable $ other.toString 110 | return () 111 | addTest $ LSpec.check s!" · {convTactic}" ((← state1_1.serializeGoals).map (·.devolatilize) = 112 | #[interiorGoal [] "b + a + c1 = b + a + c2"]) 113 | checkEq "(fragments)" state1_1.fragments.size 0 114 | 115 | /- 116 | let state1_1 ← match ← state6.fragmentExit goalConv with 117 | | .success state _ => pure state 118 | | other => do 119 | addTest $ assertUnreachable $ other.toString 120 | return () 121 | -/ 122 | 123 | let tactic := "exact h" 124 | let stateF ← match ← state1_1.tacticOn (goalId := 0) (tactic := tactic) with 125 | | .success state _ => pure state 126 | | other => do 127 | addTest $ assertUnreachable $ other.toString 128 | return () 129 | checkEq tactic ((← stateF.serializeGoals).map (·.devolatilize)) #[] 130 | checkEq "fragments" stateF.fragments.size 0 131 | 132 | where 133 | h := "b + a + c1 = b + a + c2" 134 | interiorGoal (free: List (Name × String)) (target: String) := 135 | let free := [(`a, "Nat"), (`b, "Nat"), (`c1, "Nat"), (`c2, "Nat"), (`h, h)] ++ free 136 | buildGoal free target 137 | 138 | example (p : Prop) (x y z : Nat) : p → (p → x = y) → x + z = y + z ∧ p := by 139 | intro hp hi 140 | apply And.intro 141 | conv => 142 | rhs 143 | arg 1 144 | rw [←hi] 145 | rfl 146 | tactic => exact hp 147 | exact hp 148 | 149 | def test_conv_unshielded : TestM Unit := do 150 | let rootTarget ← parseSentence "∀ (p : Prop) (x y z : Nat), p → (p → x = y) → x + z = y + z ∧ p" 151 | let state ← GoalState.create rootTarget 152 | let tactic := "intro p x y z hp hi" 153 | let .success state _ ← state.tacticOn 0 tactic | fail "intro failed" 154 | let tactic := "apply And.intro" 155 | let .success state _ ← state.tacticOn 0 tactic | fail "apply failed" 156 | let .success state _ ← state.convEnter (.prefer state.goals[0]!) | fail "Cannot enter conversion tactic mode" 157 | let .success state _ ← state.tryTactic .unfocus "rhs" | fail "rhs failed" 158 | let tactic := "arg 1" 159 | let .success state _ ← state.tryTactic .unfocus tactic | fail s!"{tactic} failed" 160 | checkEq s!" {tactic}" ((← state.serializeGoals).map (·.devolatilize)) 161 | #[ 162 | { interiorGoal [] "y" with fragment := .conv }, 163 | { interiorGoal [] "p" with userName? := `right, }, 164 | ] 165 | let tactic := "rw [←hi]" 166 | let .success state _ ← state.tryTactic .unfocus tactic | fail s!"{tactic} failed" 167 | checkEq s!" {tactic}" state.goals.length 3 168 | let tactic := "rfl" 169 | let .success state _ ← state.tryTactic .unfocus tactic | fail s!"{tactic} failed" 170 | checkEq s!" {tactic}" ((← state.serializeGoals).map (·.devolatilize)) 171 | #[ 172 | interiorGoal [] "p", 173 | { interiorGoal [] "p" with userName? := `right, }, 174 | ] 175 | checkEq "(n goals)" state.goals.length 2 176 | checkEq "(fragments)" state.fragments.size 0 177 | let tactic := "exact hp" 178 | let .success state _ ← state.tryTactic .unfocus tactic | fail s!"{tactic} failed" 179 | let tactic := "exact hp" 180 | let .success state _ ← state.tryTactic .unfocus tactic | fail s!"{tactic} failed" 181 | let root? := state.rootExpr? 182 | checkTrue "root" root?.isSome 183 | checkEq "fragments" state.fragments.size 0 184 | where 185 | interiorGoal (free: List (Name × String)) (target: String) := 186 | let free := [(`p, "Prop"), (`x, "Nat"), (`y, "Nat"), (`z, "Nat"), (`hp, "p"), (`hi, "p → x = y")] ++ free 187 | buildGoal free target 188 | 189 | example : ∀ (x y z w : Nat), y = z → x + z = w → x + y = w := by 190 | intro x y z w hyz hxzw 191 | conv => 192 | lhs 193 | arg 2 194 | rw [hyz] 195 | rfl 196 | exact hxzw 197 | 198 | def test_conv_unfinished : TestM Unit := do 199 | let rootTarget ← parseSentence "∀ (x y z w : Nat), y = z → x + z = w → x + y = w" 200 | let state ← GoalState.create rootTarget 201 | let tactic := "intro x y z w hyz hxzw" 202 | let .success state _ ← state.tacticOn 0 tactic | fail "intro failed" 203 | let convParent := state.goals[0]! 204 | let .success state _ ← state.convEnter (.prefer convParent) | fail "Cannot enter conversion tactic mode" 205 | let .success state _ ← state.tryTactic .unfocus "lhs" | fail "rhs failed" 206 | let tactic := "arg 2" 207 | let .success state _ ← state.tryTactic .unfocus tactic | fail s!"{tactic} failed" 208 | checkEq s!" {tactic}" ((← state.serializeGoals).map (·.devolatilize)) 209 | #[ 210 | { interiorGoal [] "y" with fragment := .conv }, 211 | ] 212 | let tactic := "rw [hyz]" 213 | let .success state _ ← state.tryTactic .unfocus tactic | fail s!"{tactic} failed" 214 | checkEq s!" {tactic}" ((← state.serializeGoals).map (·.devolatilize)) 215 | #[ 216 | { interiorGoal [] "z" with fragment := .conv }, 217 | ] 218 | checkTrue " (fragment)" $ state.fragments.contains state.mainGoal?.get! 219 | checkTrue " (fragment parent)" $ state.fragments.contains convParent 220 | checkTrue " (main goal)" state.mainGoal?.isSome 221 | let tactic := "rfl" 222 | let .success state _ ← state.tryTactic .unfocus tactic | fail s!"{tactic} failed" 223 | checkEq s!" {tactic}" ((← state.serializeGoals).map (·.devolatilize)) 224 | #[ 225 | interiorGoal [] "x + z = w", 226 | ] 227 | checkEq "(fragments)" state.fragments.size 0 228 | checkEq s!" {tactic}" state.goals.length 1 229 | let tactic := "exact hxzw" 230 | let .success state _ ← state.tryTactic .unfocus tactic | fail s!"{tactic} failed" 231 | let root? := state.rootExpr? 232 | checkTrue "root" root?.isSome 233 | where 234 | interiorGoal (free: List (Name × String)) (target: String) := 235 | let free := [(`x, "Nat"), (`y, "Nat"), (`z, "Nat"), (`w, "Nat"), (`hyz, "y = z"), (`hxzw, "x + z = w")] ++ free 236 | buildGoal free target 237 | 238 | def test_conv_exit : TestM Unit := do 239 | let rootTarget ← parseSentence "∀ (a b : Nat), b = 2 → 1 + a + 1 = a + b" 240 | let state ← GoalState.create rootTarget 241 | let tactic := "intro a b h" 242 | let .success state _ ← state.tacticOn 0 tactic | fail "intro failed" 243 | let .success state _ ← state.calcEnter (.prefer state.goals[0]!) | fail "Cannot enter conversion tactic mode" 244 | let .success state _ ← state.tryTactic .unfocus "1 + a + 1 = a + 1 + 1" | fail "rhs failed" 245 | let .success state _ ← state.convEnter (.prefer state.goals[0]!) | fail "Cannot enter conversion tactic mode" 246 | let .success state _ ← state.tryTactic .unfocus "rhs" | fail "rhs failed" 247 | let .success state _ ← state.tryTactic .unfocus "rw [Nat.add_comm]" | fail "rhs failed" 248 | let .success state _ ← state.fragmentExit .unfocus | fail "exit failed" 249 | checkEq "(parents)" state.parentExprs.length 1 250 | where 251 | interiorGoal (free: List (Name × String)) (target: String) := 252 | let free := [(`a, "Nat"), (`b, "Nat"), (`h, "b = 2")] ++ free 253 | buildGoal free target 254 | 255 | example : ∀ (a b c d: Nat), a + b = b + c → b + c = c + d → a + b = c + d := by 256 | intro a b c d h1 h2 257 | calc a + b = b + c := by apply h1 258 | _ = c + d := by apply h2 259 | 260 | def test_calc: TestM Unit := do 261 | let rootTarget ← parseSentence "∀ (a b c d: Nat), a + b = b + c → b + c = c + d → a + b = c + d" 262 | let state0 ← GoalState.create rootTarget 263 | let tactic := "intro a b c d h1 h2" 264 | let state1 ← match ← state0.tacticOn (goalId := 0) (tactic := tactic) with 265 | | .success state _ => pure state 266 | | other => do 267 | addTest $ assertUnreachable $ other.toString 268 | return () 269 | checkEq tactic ((← state1.serializeGoals).map (·.devolatilize)) 270 | #[interiorGoal [] "a + b = c + d"] 271 | let pred := "a + b = b + c" 272 | let .success state1 _ ← state1.calcEnter state1.mainGoal?.get! | fail "Could not enter calc" 273 | let state2 ← match ← state1.tacticOn 0 pred with 274 | | .success state _ => pure state 275 | | other => do 276 | addTest $ assertUnreachable $ other.toString 277 | return () 278 | checkEq s!"calc {pred} := _" ((← state2.serializeGoals).map (·.devolatilize)) 279 | #[ 280 | { interiorGoal [] "a + b = b + c" with userName? := `calc }, 281 | { interiorGoal [] "b + c = c + d" with fragment := .calc }, 282 | ] 283 | checkFalse "(2.0 prev rhs)" <| state2.fragments.contains (state2.get! 0) 284 | checkTrue "(2.1 prev rhs)" <| state2.fragments[(state2.get! 1)]? matches .some (.calc ..) 285 | 286 | let tactic := "apply h1" 287 | let state2m ← match ← state2.tacticOn (goalId := 0) (tactic := tactic) with 288 | | .success state _ => pure state 289 | | other => do 290 | addTest $ assertUnreachable $ other.toString 291 | return () 292 | let state3 ← match state2m.continue state2 with 293 | | .ok state => pure state 294 | | .error e => do 295 | addTest $ expectationFailure "continue" e 296 | return () 297 | let pred := "_ = c + d" 298 | let state4 ← match ← state3.tacticOn 0 pred with 299 | | .success state _ => pure state 300 | | other => do 301 | addTest $ assertUnreachable $ other.toString 302 | return () 303 | checkEq s!"calc {pred} := _" ((← state4.serializeGoals).map (·.devolatilize)) 304 | #[ 305 | { interiorGoal [] "b + c = c + d" with userName? := `calc }, 306 | ] 307 | checkFalse "(4.0 prev rhs)" <| state4.fragments.contains (state4.get! 0) 308 | let tactic := "apply h2" 309 | let state4m ← match ← state4.tacticOn (goalId := 0) (tactic := tactic) with 310 | | .success state _ => pure state 311 | | other => do 312 | addTest $ assertUnreachable $ other.toString 313 | return () 314 | checkEq "(fragments)" state4m.fragments.size 0 315 | addTest $ LSpec.test "(4m root)" state4m.rootExpr?.isSome 316 | where 317 | interiorGoal (free: List (Name × String)) (target: String) := 318 | let free := [(`a, "Nat"), (`b, "Nat"), (`c, "Nat"), (`d, "Nat"), 319 | (`h1, "a + b = b + c"), (`h2, "b + c = c + d")] ++ free 320 | buildGoal free target 321 | 322 | def suite (env: Environment): List (String × IO LSpec.TestSeq) := 323 | [ 324 | ("conv simple", test_conv_simple), 325 | ("conv unshielded", test_conv_unshielded), 326 | ("conv unfinished", test_conv_unfinished), 327 | ("conv exit", test_conv_exit), 328 | ("calc", test_calc), 329 | ] |>.map (λ (name, t) => (name, runTestTermElabM env t)) 330 | 331 | end Pantograph.Test.Tactic.Fragment 332 | -------------------------------------------------------------------------------- /Pantograph/Frontend/Refactor.lean: -------------------------------------------------------------------------------- 1 | /- A scrolling refactor algorithm: The algorithm ingests Lean compilation units 2 | on one end and outputs compilation units on the other. -/ 3 | import Pantograph.Frontend.Basic 4 | import Pantograph.Frontend.InfoTree 5 | import Pantograph.Delate 6 | 7 | open Lean 8 | 9 | namespace Pantograph.Frontend 10 | 11 | namespace Refactor 12 | 13 | /-- A command in the input file, frozen in context -/ 14 | structure Command where 15 | dependencies : NameSet := .empty 16 | stx : Syntax 17 | trees : List Elab.InfoTree 18 | hasError : Bool := false 19 | constants : NameSet := .empty 20 | state : Elab.Command.State 21 | messages : List Message 22 | 23 | inductive CommandCategory where 24 | -- Definition of data 25 | | data 26 | -- Definition fo theorems 27 | | declaration 28 | -- section, variable, universe 29 | | flow 30 | -- Things which can be discarded. 31 | | auxiliary 32 | -- No refactor units may cross an `.unknown` boundary. 33 | | unknown 34 | 35 | protected def Command.category (command : Command) : CommandCategory := 36 | match command.stx.getKind with 37 | | `Lean.Parser.Command.declaration => 38 | match command.stx.getArg 2 |>.getKind with 39 | | `Lean.Parser.Command.structure 40 | | `Lean.Parser.Command.inductive 41 | => .data 42 | | `Lean.Parser.Command.theorem 43 | | `Lean.Parser.Command.definition 44 | => .data 45 | | _ => .unknown 46 | | `Lean.Parser.Command.section 47 | | `Lean.Parser.Command.namespace 48 | | `Lean.Parser.Command.variable 49 | | `Lean.Parser.Command.end 50 | => .flow 51 | | `Lean.Parser.Command.set_option 52 | => .auxiliary 53 | | _ => .unknown 54 | protected def Command.comments (command : Command) : Syntax := 55 | let modifiers := command.stx.getArg 0 56 | let comments := modifiers.getArg 0 57 | comments[0] 58 | 59 | structure Config where 60 | coreOptions : Options := {} 61 | deriving Inhabited 62 | 63 | structure Context where 64 | inContext : Parser.InputContext 65 | config : Config := {} 66 | 67 | structure State where 68 | outContext : Parser.InputContext 69 | outState : Elab.Frontend.State 70 | 71 | -- Collected top-level units, scrolling 72 | commands : List Command := [] 73 | 74 | /-- Two monads rolled into one -/ 75 | abbrev RefactorM := ReaderT Context $ StateRefT State IO 76 | 77 | def readConfig : RefactorM Config := do 78 | return (← read).config 79 | def readCoreOptions : RefactorM Options := do 80 | return (← readConfig).coreOptions 81 | 82 | def fail { α } (s : String) : IO α := 83 | throw <| .userError s 84 | 85 | def mergeFileMap (fm1 fm2 : FileMap) : FileMap := 86 | let bias := fm1.source.rawEndPos.byteIdx + 1 87 | let mappedPos := fm2.positions.map λ pos => { byteIdx := pos.byteIdx + bias } 88 | { 89 | source := s!"{fm1.source}\n{fm2.source}", 90 | positions := fm1.positions.take (fm1.positions.size - 1) ++ mappedPos 91 | } 92 | 93 | /-- Add one command to the refactored file -/ 94 | def pushNewCommand (f : Format) : RefactorM Unit := do 95 | modify λ state@{ outContext := outContext@{ fileMap, .. }, .. } => 96 | let payload := f.pretty 97 | let merged := mergeFileMap fileMap payload.toFileMap 98 | { 99 | state with outContext := { 100 | outContext with 101 | inputString := merged.source, 102 | fileMap := merged, 103 | endPos := merged.source.rawEndPos, 104 | endPos_valid := by simp, 105 | } 106 | } 107 | -- After modification, run the parser ahead by one position 108 | let { outContext := inputCtx, outState, .. } ← get 109 | let (_endFlag, outState@ { commandState := { messages, .. }, .. }) ← 110 | Elab.Frontend.processCommand.run { inputCtx } |>.run outState 111 | -- Ensure no error has occurred 112 | if messages.hasErrors then 113 | let messages ← messages.toList.mapM (·.toString) 114 | throw (.userError s!"Error messages: {messages}") 115 | modify ({ · with outState }) 116 | 117 | /-- Run `FrontendM` at the tail of the out file -/ 118 | def liftFrontend { α } (x : FrontendM α) : RefactorM α := do 119 | let { outContext := inputCtx, outState, .. } ← get 120 | x.run {} |>.run { inputCtx } |>.run' outState 121 | /-- Run `CoreM` at the tail of the out file -/ 122 | def runCoreM { α } (x : CoreM α) : RefactorM α := do 123 | liftFrontend $ runCommandElabM $ Elab.Command.liftCoreM x 124 | 125 | def pushNewCommand' (command : Syntax.Command) : RefactorM Unit := do 126 | let f ← runCoreM do 127 | PrettyPrinter.ppCommand command 128 | pushNewCommand f 129 | 130 | /-- runs a "frozen" `CommandElabM` that can't modify anything. -/ 131 | @[inline] protected 132 | def Command.runCommandElabM (command : Command) (x : Elab.Command.CommandElabM α) : RefactorM α := do 133 | let inputCtx := (← read).inContext 134 | let cmdCtx : Elab.Command.Context := { 135 | fileName := inputCtx.fileName 136 | fileMap := inputCtx.fileMap 137 | snap? := none, 138 | cancelTk? := .none, 139 | } 140 | match (← liftM <| EIO.toIO' <| (x cmdCtx).run command.state) with 141 | | Except.error e => throw <| IO.Error.userError s!"unexpected internal error: {← e.toMessageData.toString}" 142 | | Except.ok (a, _sNew) => return a 143 | 144 | @[inline] protected 145 | def Command.runCoreM { α } (command : Command) (c : CoreM α) : RefactorM α := 146 | command.runCommandElabM $ Elab.Command.liftCoreM c 147 | 148 | def constantDependencies (env : Environment) (name : Name) : NameSet := 149 | let const := env.find? name |>.get! 150 | let s := const.type.getUsedConstantsAsSet 151 | let s := match const.value? with 152 | | .some v => s.append v.getUsedConstantsAsSet 153 | | .none => s 154 | s 155 | 156 | def hasSorry (step : CompilationStep) : Bool := 157 | step.trees.any λ tree => 158 | let nodes := tree.filter λ 159 | | .ofTermInfo { expr, .. } => expr.isSorry 160 | | .ofTacticInfo { stx, .. } => stx.isOfKind `Lean.Parser.Tactic.tacticSorry 161 | | _ => false 162 | !nodes.isEmpty 163 | 164 | /-- Scroll to the end of the file, reading all compilation units in the process -/ 165 | def preprocess : FrontendM (List Command) := mapCompilationSteps λ step => do 166 | let constants ← step.newConstants 167 | let dependencies := constants.foldl (init := NameSet.empty) λ acc c => 168 | acc.append $ constantDependencies step.after c 169 | let commandState := (← getThe Elab.Frontend.State).commandState 170 | let unit := if step.msgs.any (·.severity == .error) then 171 | { 172 | hasError := true, 173 | stx := step.stx, 174 | trees := step.trees, 175 | state := commandState, 176 | messages := step.msgs, 177 | } 178 | else 179 | { 180 | stx := step.stx, 181 | trees := step.trees, 182 | dependencies, 183 | constants, 184 | state := commandState, 185 | messages := step.msgs, 186 | } 187 | return unit 188 | 189 | /-- Creates `combine (combine a[0] a[1]) a[2] ...` -/ 190 | def mkProdElem (combine : Name := ``And.intro) : List Expr → MetaM Expr 191 | | .nil => return .const `Unit [] 192 | | [a] => return a 193 | | x :: xs => do 194 | let r ← mkProdElem combine xs 195 | Meta.mkAppM combine #[x, r] 196 | 197 | private def mkDocComment (s : String) := 198 | mkNode ``Parser.Command.docComment #[mkAtom "/--", mkAtom s!"{s} -/"] 199 | 200 | protected def _root_.Array.last! { α } [Inhabited α] (a : Array α) : α := 201 | match h : a.size with 202 | | 0 => default 203 | | n+1 => a[n] 204 | 205 | def distilSearchTarget { α } (head : Command) (tail : List Command) (f : (Expr × Expr) → List (Expr × Expr) → Elab.Term.TermElabM α) 206 | : RefactorM α := do 207 | let (headName, witness, witnessValue) ← head.runCommandElabM do 208 | Elab.Command.liftCoreM do 209 | let env := head.state.env 210 | let name := head.constants.toList.head! 211 | let info := env.find? name |>.get! 212 | let .some value := info.value? 213 | | throwError s!"Constant has no value: {name}" 214 | return (name, ← normalize info.type, ← normalize value) 215 | let companions ← tail.mapM λ command => command.runCommandElabM do 216 | Elab.Command.liftTermElabM do 217 | let env := command.state.env 218 | let name := command.constants.toList.head! 219 | let info := env.find? name |>.get! 220 | let type ← normalize info.type 221 | let c ← mkConstWithLevelParams headName 222 | let type ← Meta.kabstract type c 223 | let .some value := info.value? 224 | | throwError s!"Constant has no value: {name}" 225 | -- Normalization strips away matchers. 226 | let value ← normalize value 227 | let value ← Meta.kabstract value c 228 | return (type, value) 229 | liftFrontend $ runCommandElabM $ Elab.Command.liftTermElabM do 230 | f (witness, witnessValue) companions 231 | where 232 | normalize (e : Expr) : CoreM Expr := do 233 | unfoldAuxLemmas $ ← unfoldMatchers e 234 | 235 | /-- Fold `sorry`s into one definition -/ 236 | def foldTheoremsFlat (head : Command) (tail : List Command) : RefactorM Syntax.Command := do 237 | -- Concatenate all doc comments 238 | let allDocs := "\n".intercalate $ (head :: tail).filterMap λ command => 239 | let `(docComment|$comment) := command.comments 240 | let s := comment.getDocString 241 | if s.isEmpty then .none else s 242 | let headName := head.constants.toList.head! 243 | let binderName := match headName with 244 | | .str _ binderName => Name.mkSimple binderName 245 | | _ => `x 246 | let coreOptions ← readCoreOptions 247 | distilSearchTarget head tail λ (witness, _) companions => do 248 | -- Construct the companion 249 | let companion ← Meta.withLocalDeclD binderName witness λ binder => do 250 | let companion ← mkProdElem ``And <| companions.map (·.fst.instantiate1 binder) 251 | Meta.mkLambdaFVars #[binder] companion 252 | let target ← Meta.mkAppOptM ``Subtype #[witness, companion] 253 | Meta.check target 254 | -- Delaborate this back into syntax 255 | let target ← withOptions (λ _ => pp.analyze.set coreOptions true) do 256 | PrettyPrinter.delab target 257 | let theoremIdent := mkIdent $ Name.mkSimple s!"{binderName}_composite" 258 | let comment? := if allDocs.isEmpty then .none else .some $ mkDocComment allDocs 259 | `(command|$[$comment?:docComment]? def $theoremIdent : $target := sorry) 260 | 261 | structure DependencyTracker where 262 | -- Constants generated during the next batch of commands to be processed 263 | innerConstants : NameSet := {} 264 | 265 | isNonFlat : Bool := false 266 | structure DependencyStructure where 267 | -- Commands which depend on each other 268 | component : List Command := [] 269 | -- Intercalating commands 270 | intercalating : List Command := [] 271 | -- Remainder 272 | tail : List Command 273 | -- Structure 274 | tracker : DependencyTracker 275 | 276 | def extractDependencyStructure (head : Command) (commands : List Command) 277 | : RefactorM DependencyStructure := do 278 | let ((series, other), tracker) := (λ (z : StateM DependencyTracker _) => z.run {}) $ 279 | commands.zipIdx.partitionM λ (command, _) => do 280 | let tracker ← get 281 | if command.dependencies.any tracker.innerConstants.contains then 282 | let innerConstants := command.constants.foldl 283 | (init := tracker.innerConstants) 284 | λ acc n => acc.insert n 285 | modify ({· with innerConstants, isNonFlat := true}) 286 | return true 287 | if command.dependencies.any head.constants.contains then 288 | let innerConstants := command.constants.foldl 289 | (init := tracker.innerConstants) 290 | λ acc n => acc.insert n 291 | modify ({· with innerConstants }) 292 | return true 293 | else 294 | return false 295 | if series.isEmpty then 296 | return { 297 | tail := commands, 298 | tracker, 299 | } 300 | -- Find all intercalating declarations 301 | let maxIdx := series.map Prod.snd |>.max?.get! 302 | let (intercalating, tail) := other.partition λ (_, idx) => idx < maxIdx 303 | return { 304 | component := series.map Prod.fst, 305 | intercalating := intercalating.map Prod.fst, 306 | tail := tail.map Prod.fst, 307 | tracker 308 | } 309 | 310 | /-- Scroll one unit down from the top -/ 311 | def collectNextCommand : RefactorM Unit := do 312 | let { commands, .. } ← get 313 | let decl :: commands := commands | Refactor.fail "No commands left" 314 | modify ({ · with commands }) -- Prevents infinite loop 315 | 316 | let isSearchTarget ← decl.runCoreM do 317 | if decl.constants.isEmpty then 318 | return false 319 | let name := decl.constants.toList.head! 320 | let info := (← getEnv).find? name |>.get! 321 | let .some value := info.value? | return false 322 | return value.hasSorry 323 | if !isSearchTarget then 324 | pushNewCommand' (⟨decl.stx⟩ : Syntax.Command) 325 | return 326 | 327 | let depstr ← extractDependencyStructure decl commands 328 | if depstr.component.isEmpty then 329 | pushNewCommand' (⟨decl.stx⟩ : Syntax.Command) 330 | return 331 | 332 | if depstr.tracker.isNonFlat then 333 | Refactor.fail "Cannot refactor non-flat dependency structure" 334 | modify ({ · with commands := depstr.tail }) 335 | 336 | -- Push all intercalating commands 337 | for command in depstr.intercalating do 338 | pushNewCommand' (⟨command.stx⟩ : Syntax.Command) 339 | 340 | let f ← foldTheoremsFlat decl depstr.component 341 | pushNewCommand' f 342 | 343 | end Refactor 344 | 345 | open Refactor in 346 | def runRefactor (env : Environment) (source : String) 347 | (config : Refactor.Config := {}) (fileName := defaultFileName) : IO String := do 348 | let (fContext, fState) ← createContextStateFromFile source fileName env {} 349 | let commands ← preprocess.run {} |>.run fContext |>.run' fState 350 | let errors := commands.filter (·.hasError) 351 | if let .some error := errors.head? then 352 | let message ← error.messages.mapM (·.toString) 353 | throw $ IO.userError $ "\n".intercalate message 354 | let m : RefactorM Unit := do 355 | while !(← get).commands.isEmpty do 356 | collectNextCommand 357 | let outContext := { 358 | fContext.inputCtx with 359 | inputString := "", 360 | fileMap := "".toFileMap, 361 | endPos := "".rawEndPos, 362 | endPos_valid := by simp, 363 | } 364 | let parserState := {} 365 | let outState := { 366 | commandState := Elab.Command.mkState env {} {}, 367 | parserState, 368 | cmdPos := parserState.pos, 369 | } 370 | let (_, state) ← m.run { config, inContext := fContext.inputCtx } 371 | |>.run { outContext, outState, commands } 372 | return state.outContext.inputString 373 | 374 | export Refactor (RefactorM) 375 | --------------------------------------------------------------------------------