├── .github
└── workflows
│ └── ci.yml
├── .gitignore
├── .vscode
├── copyright.code-snippets
├── extensions.json
├── module-docstring.code-snippets
└── settings.json
├── README.md
├── REPL.lean
├── REPL
├── Frontend.lean
├── JSON.lean
├── Lean
│ ├── ContextInfo.lean
│ ├── Environment.lean
│ ├── InfoTree.lean
│ └── InfoTree
│ │ └── ToJson.lean
├── Main.lean
├── Snapshots.lean
└── Util
│ ├── Path.lean
│ └── Pickle.lean
├── lake-manifest.json
├── lakefile.toml
├── lean-toolchain
├── test.sh
└── test
├── Mathlib
├── .gitignore
├── H20231110.sh
├── ReplMathlibTests.lean
├── lake-manifest.json
├── lakefile.toml
├── lean-toolchain
├── test.sh
└── test
│ ├── 20240209.expected.out
│ ├── 20240209.in
│ ├── 20240209.lean
│ ├── H20231020.expected.out
│ ├── H20231020.in
│ ├── H20231020.lean
│ ├── H20231110.expected.out
│ ├── H20231110.in
│ ├── H20231115.expected.out
│ ├── H20231115.in
│ ├── H20231115_2.expected.out
│ ├── H20231115_2.in
│ ├── H20231115_3.expected.out
│ ├── H20231115_3.in
│ ├── H20231214.in
│ ├── H20231214.lean
│ ├── H20231215.expected.out
│ ├── H20231215.in
│ ├── H20231215_2.expected.out
│ ├── H20231215_2.in
│ ├── exact.expected.out
│ ├── exact.in
│ ├── import_Mathlib.lean
│ ├── induction.expected.out
│ ├── induction.in
│ ├── induction.lean
│ ├── on_goal.expected.out
│ ├── on_goal.in
│ ├── pickle.expected.out
│ ├── pickle.in
│ ├── pickle_2.expected.out
│ ├── pickle_2.in
│ ├── placeholder_synthesis.expected.out
│ └── placeholder_synthesis.in
├── all_tactics.expected.out
├── all_tactics.in
├── app_type_mismatch.expected.out
├── app_type_mismatch.in
├── app_type_mismatch2.expected.out
├── app_type_mismatch2.in
├── assumption_proof.expected.out
├── assumption_proof.in
├── by_cases.expected.out
├── by_cases.in
├── by_cases.lean
├── calc.expected.out
├── calc.in
├── def_eval.expected.out
├── def_eval.in
├── dup_msg.expected.out
├── dup_msg.in
├── dup_sorries.expected.out
├── dup_sorries.in
├── file.expected.out
├── file.in
├── file.lean
├── have_by_sorry.expected.out
├── have_by_sorry.in
├── import_lean.expected.out
├── import_lean.in
├── incomplete.expected.out
├── incomplete.in
├── incomplete.lean
├── infotree.expected.out
├── infotree.in
├── invalid_tactic.expected.out
├── invalid_tactic.in
├── name_generator.expected.out
├── name_generator.in
├── no_goal_sorry.expected.out
├── no_goal_sorry.in
├── no_goal_sorry_2.expected.out
├── no_goal_sorry_2.in
├── options.expected.out
├── options.in
├── pickle_environment.expected.out
├── pickle_environment.in
├── pickle_environment_with_imports.expected.out
├── pickle_environment_with_imports.in
├── pickle_open.expected.out
├── pickle_open.in
├── pickle_open_2.expected.out
├── pickle_open_2.in
├── pickle_open_scoped.expected.out
├── pickle_open_scoped.in
├── pickle_open_scoped_2.expected.out
├── pickle_open_scoped_2.in
├── pickle_proof_state_1.expected.out
├── pickle_proof_state_1.in
├── pickle_proof_state_2.expected.out
├── pickle_proof_state_2.in
├── pickle_proof_state_env.expected.out
├── pickle_proof_state_env.in
├── pickle_scoped_notation.in
├── pickle_scoped_notation_2.in
├── proof_branching.expected.out
├── proof_branching.in
├── proof_branching2.expected.out
├── proof_branching2.in
├── proof_step.expected.out
├── proof_step.in
├── proof_transitivity.expected.out
├── proof_transitivity.in
├── readme.expected.out
├── readme.in
├── root_goals.expected.out
├── root_goals.in
├── self_proof_apply_check.expected.out
├── self_proof_apply_check.in
├── self_proof_check.expected.out
├── self_proof_check.in
├── self_proof_exact_check.expected.out
├── self_proof_exact_check.in
├── self_proof_rw.expected.out
├── self_proof_rw.in
├── sorry_hypotheses.expected.out
├── sorry_hypotheses.in
├── synthesize_placeholder.expected.out
├── synthesize_placeholder.in
├── tactic_mode_sorry.expected.out
├── tactic_mode_sorry.in
├── tactic_sorry.expected.out
├── tactic_sorry.in
├── term_sorry.expected.out
├── term_sorry.in
├── trace_simp.expected.out
├── trace_simp.in
├── unfinished_tactic_block.expected.out
├── unfinished_tactic_block.in
├── unknown_environment.expected.out
├── unknown_environment.in
├── unknown_proof_state.expected.out
├── unknown_proof_state.in
├── unknown_tactic.expected.out
├── unknown_tactic.in
├── variables.expected.out
└── variables.in
/.github/workflows/ci.yml:
--------------------------------------------------------------------------------
1 | name: Run Tests
2 |
3 | on: [push, pull_request]
4 |
5 | jobs:
6 | test:
7 | runs-on: ubuntu-latest
8 |
9 | steps:
10 | - name: Checkout code
11 | uses: actions/checkout@v2
12 |
13 | - name: install elan
14 | run: |
15 | set -o pipefail
16 | curl -sSfL https://github.com/leanprover/elan/releases/download/v3.0.0/elan-x86_64-unknown-linux-gnu.tar.gz | tar xz
17 | ./elan-init -y --default-toolchain none
18 | echo "$HOME/.elan/bin" >> $GITHUB_PATH
19 |
20 | - name: build
21 | run: lake build
22 |
23 | - name: Run tests
24 | run: ./test.sh
25 |
--------------------------------------------------------------------------------
/.gitignore:
--------------------------------------------------------------------------------
1 | /build
2 | /lake-packages/*
3 | /lakefile.olean
4 | /.lake
5 | /test/Mathlib/.lake
6 | /test/*.olean
7 | /test/*.olean.tmp
8 | /test/*.produced.out
9 |
--------------------------------------------------------------------------------
/.vscode/copyright.code-snippets:
--------------------------------------------------------------------------------
1 | {
2 | "Copyright header for mathlib": {
3 | "scope": "lean4",
4 | "prefix": "copyright",
5 | "body": [
6 | "/-",
7 | "Copyright (c) ${CURRENT_YEAR} $1. All rights reserved.",
8 | "Released under Apache 2.0 license as described in the file LICENSE.",
9 | "Authors: $1",
10 | "-/"
11 | ]
12 | }
13 | }
14 |
--------------------------------------------------------------------------------
/.vscode/extensions.json:
--------------------------------------------------------------------------------
1 | {
2 | // See https://go.microsoft.com/fwlink/?LinkId=827846 to learn about workspace recommendations.
3 | // Extension identifier format: ${publisher}.${name}. Example: vscode.csharp
4 |
5 | // List of extensions which should be recommended for users of this workspace.
6 | "recommendations": [
7 | "leanprover.lean4"
8 | ],
9 | // List of extensions recommended by VS Code that should not be recommended for users of this workspace.
10 | "unwantedRecommendations": [
11 | "ms-vscode-remote.remote-containers"
12 | ]
13 | }
14 |
--------------------------------------------------------------------------------
/.vscode/module-docstring.code-snippets:
--------------------------------------------------------------------------------
1 | {
2 | "Module docstring for mathlib": {
3 | "scope": "lean4",
4 | "prefix": "module docstring",
5 | "body": [
6 | "/-!",
7 | "# ${TM_FILENAME_BASE/([^_]*)(_?)/${1:/capitalize}${2:+ }/g}",
8 | "",
9 | "## Main definitions",
10 | "",
11 | "* `FooBar`",
12 | "",
13 | "## Main statements",
14 | "",
15 | "* `fooBar_unique`",
16 | "",
17 | "## Notation",
18 | "",
19 | "",
20 | "",
21 | "## Implementation details",
22 | "",
23 | "",
24 | "",
25 | "## References",
26 | "",
27 | "* [F. Bar, *Quuxes*][bibkey]",
28 | "",
29 | "## Tags",
30 | "",
31 | "Foobars, barfoos",
32 | "-/",
33 | "",
34 | ]},
35 | }
36 |
--------------------------------------------------------------------------------
/.vscode/settings.json:
--------------------------------------------------------------------------------
1 | {
2 | "editor.insertSpaces": true,
3 | "editor.tabSize": 2,
4 | "editor.rulers" : [100],
5 | "files.encoding": "utf8",
6 | "files.eol": "\n",
7 | "files.insertFinalNewline": true,
8 | // We don't use this: it messes up our test files!
9 | // "files.trimFinalNewlines": true,
10 | "files.trimTrailingWhitespace": true,
11 | }
12 |
--------------------------------------------------------------------------------
/README.md:
--------------------------------------------------------------------------------
1 | # A read-eval-print-loop for Lean 4
2 |
3 | Run using `lake exe repl`.
4 | Communicates via JSON on stdin and stdout.
5 | Commands should be separated by blank lines.
6 |
7 | The REPL works both in "command" mode and "tactic" mode.
8 |
9 | ## Command mode
10 |
11 | In command mode, you send complete commands (e.g. declarations) to the REPL.
12 |
13 | Commands may be of the form
14 |
15 | ```json
16 | { "cmd" : "def f := 2" }
17 | ```
18 |
19 | ```json
20 | { "cmd" : "example : f = 2 := rfl", "env" : 1 }
21 | ```
22 |
23 | The `env` field, if present,
24 | must contain a number received in the `env` field of a previous response,
25 | and causes the command to be run in the existing environment.
26 |
27 | If there is no `env` field, a new environment is created.
28 |
29 | You can only use `import` commands when you do not specify the `env` field.
30 |
31 | You can backtrack simply by using earlier values for `env`.
32 |
33 | The response includes:
34 | * A numeric label for the `Environment` after your command,
35 | which you can use as the starting point for subsequent commands.
36 | * Any messages generated while processing your command.
37 | * A list of the `sorry`s in your command, including
38 | * their expected type, and
39 | * a numeric label for the proof state at the `sorry`, which you can then use in tactic mode.
40 |
41 | Example output:
42 |
43 | ```json
44 | {"sorries":
45 | [{"pos": {"line": 1, "column": 18},
46 | "endPos": {"line": 1, "column": 23},
47 | "goal": "⊢ Nat",
48 | "proofState": 0}],
49 | "messages":
50 | [{"severity": "error",
51 | "pos": {"line": 1, "column": 23},
52 | "endPos": {"line": 1, "column": 26},
53 | "data":
54 | "type mismatch\n rfl\nhas type\n f = f : Prop\nbut is expected to have type\n f = 2 : Prop"}],
55 | "env": 6}
56 | ```
57 |
58 | showing any messages generated, and sorries with their goal states.
59 |
60 | ## File mode
61 |
62 | There is a simple wrapper around command mode that allows reading in an entire file.
63 |
64 | If `test/file.lean` contains
65 | ```lean
66 | def f : Nat := 37
67 |
68 | def g := 2
69 |
70 | theorem h : f + g = 39 := by exact rfl
71 | ```
72 |
73 | then
74 | ```
75 | echo '{"path": "test/file.lean", "allTactics": true}' | lake exe repl
76 | ```
77 | results in output
78 | ```json
79 | {"tactics":
80 | [{"tactic": "exact rfl",
81 | "proofState": 0,
82 | "pos": {"line": 5, "column": 29},
83 | "goals": "⊢ f + g = 39",
84 | "endPos": {"line": 5, "column": 38}}],
85 | "env": 0}
86 | ```
87 |
88 | ## Tactic mode (experimental)
89 |
90 | To enter tactic mode issue a command containing a `sorry`,
91 | and then use the `proofState` index returned for each `sorry`.
92 |
93 | Example usage:
94 | ```json
95 | {"cmd" : "def f (x : Unit) : Nat := by sorry"}
96 |
97 | {"sorries":
98 | [{"proofState": 0,
99 | "pos": {"line": 1, "column": 29},
100 | "goal": "x : Unit\n⊢ Nat",
101 | "endPos": {"line": 1, "column": 34}}],
102 | "messages":
103 | [{"severity": "warning",
104 | "pos": {"line": 1, "column": 4},
105 | "endPos": {"line": 1, "column": 5},
106 | "data": "declaration uses 'sorry'"}],
107 | "env": 0}
108 |
109 | {"tactic": "apply Int.natAbs", "proofState": 0}
110 |
111 | {"proofState": 1, "goals": ["x : Unit\n⊢ Int"]}
112 |
113 | {"tactic": "exact -37", "proofState": 1}
114 |
115 | {"proofState": 2, "goals": []}
116 | ```
117 |
118 | You can use `sorry` in tactic mode.
119 | The result will contain additional `proofState` identifiers for the goal at each sorry.
120 |
121 | At present there is nothing you can do with a completed proof state:
122 | we would like to extend this so that you can replace the original `sorry` with your tactic script,
123 | and obtain the resulting `Environment`
124 |
125 | ## Pickling
126 |
127 | The REPL supports pickling environments and proof states to disk as `.olean` files.
128 | As long as the same imports are available, it should be possible to move such an `.olean` file
129 | to another machine and unpickle into a new REPL session.
130 |
131 | The commands are
132 |
133 | ```json
134 | {"pickleTo": "path/to/file.olean", "env": 7}
135 |
136 | {"pickleTo": "path/to/file.olean", "proofState": 17}
137 |
138 | {"unpickleEnvFrom": "path/to/file.olean"}
139 |
140 | {"unpickleProofStateFrom": "path/to/file.olean"}
141 | ```
142 |
143 | The unpickling commands will report the new "env" or "proofState" identifier that
144 | you can use in subsequent commands.
145 |
146 | Pickling is quite efficient:
147 | * we don't record full `Environment`s, only the changes relative to imports
148 | * unpickling uses memory mapping
149 | * file sizes are generally small, but see https://github.com/digama0/leangz if compression is
150 | desirable
151 |
152 | ## Using the REPL from another project
153 |
154 | Set up your project as usual using `lake new` or `lake init`
155 | (or the interactive setup GUI available via the VSCode extension under the `∀` menu).
156 |
157 | In that project, add `require` statements in the `lakefile.lean` for any dependencies you need
158 | (e.g. Mathlib). (You probably should verify that `lake build` works as expected in that project.)
159 |
160 | Now you can run the REPL as:
161 | ```shell
162 | lake env ../path/to/repl/.lake/build/bin/repl < commands.in
163 | ```
164 | (Here `../path/to/repl/` represents the path to your checkout of this repository,
165 | in which you've already run `lake build`.)
166 |
167 | The `lake env` prefix sets up the environment associated to your local project, so that the REPL
168 | can find needed imports.
169 |
170 | ## Future work
171 |
172 | * Replay tactic scripts from tactic mode back into the original `sorry`.
173 | * Currently if you create scoped environment extensions (e.g. scoped notations) in a session
174 | these are not correctly pickled and unpickled in later sessions.
175 |
--------------------------------------------------------------------------------
/REPL.lean:
--------------------------------------------------------------------------------
1 | import REPL.Frontend
2 | import REPL.Lean.InfoTree
3 | import REPL.JSON
4 | import REPL.Main
5 |
--------------------------------------------------------------------------------
/REPL/Frontend.lean:
--------------------------------------------------------------------------------
1 | /-
2 | Copyright (c) 2023 Scott Morrison. All rights reserved.
3 | Released under Apache 2.0 license as described in the file LICENSE.
4 | Authors: Scott Morrison
5 | -/
6 | import Lean.Elab.Frontend
7 |
8 | open Lean Elab
9 |
10 | namespace Lean.Elab.IO
11 |
12 | /--
13 | Wrapper for `IO.processCommands` that enables info states, and returns
14 | * the new command state
15 | * messages
16 | * info trees
17 | -/
18 | def processCommandsWithInfoTrees
19 | (inputCtx : Parser.InputContext) (parserState : Parser.ModuleParserState)
20 | (commandState : Command.State) : IO (Command.State × List Message × List InfoTree) := do
21 | let commandState := { commandState with infoState.enabled := true }
22 | let s ← IO.processCommands inputCtx parserState commandState <&> Frontend.State.commandState
23 | pure (s, s.messages.toList, s.infoState.trees.toList)
24 |
25 | /--
26 | Process some text input, with or without an existing command state.
27 | If there is no existing environment, we parse the input for headers (e.g. import statements),
28 | and create a new environment.
29 | Otherwise, we add to the existing environment.
30 |
31 | Returns:
32 | 1. The header-only command state (only useful when cmdState? is none)
33 | 2. The resulting command state after processing the entire input
34 | 3. List of messages
35 | 4. List of info trees
36 | -/
37 | def processInput (input : String) (cmdState? : Option Command.State)
38 | (opts : Options := {}) (fileName : Option String := none) :
39 | IO (Command.State × Command.State × List Message × List InfoTree) := unsafe do
40 | Lean.initSearchPath (← Lean.findSysroot)
41 | enableInitializersExecution
42 | let fileName := fileName.getD ""
43 | let inputCtx := Parser.mkInputContext input fileName
44 |
45 | match cmdState? with
46 | | none => do
47 | -- Split the processing into two phases to prevent self-reference in proofs in tactic mode
48 | let (header, parserState, messages) ← Parser.parseHeader inputCtx
49 | let (env, messages) ← processHeader header opts messages inputCtx
50 | let headerOnlyState := Command.mkState env messages opts
51 | let (cmdState, messages, trees) ← processCommandsWithInfoTrees inputCtx parserState headerOnlyState
52 | return (headerOnlyState, cmdState, messages, trees)
53 |
54 | | some cmdStateBefore => do
55 | let parserState : Parser.ModuleParserState := {}
56 | let (cmdStateAfter, messages, trees) ← processCommandsWithInfoTrees inputCtx parserState cmdStateBefore
57 | return (cmdStateBefore, cmdStateAfter, messages, trees)
58 |
--------------------------------------------------------------------------------
/REPL/JSON.lean:
--------------------------------------------------------------------------------
1 | /-
2 | Copyright (c) 2023 Scott Morrison. All rights reserved.
3 | Released under Apache 2.0 license as described in the file LICENSE.
4 | Authors: Scott Morrison
5 | -/
6 | import Lean.Data.Json
7 | import Lean.Message
8 | import Lean.Elab.InfoTree.Main
9 |
10 | open Lean Elab InfoTree
11 |
12 | namespace REPL
13 |
14 | structure CommandOptions where
15 | allTactics : Option Bool := none
16 | rootGoals : Option Bool := none
17 | /--
18 | Should be "full", "tactics", "original", or "substantive".
19 | Anything else is ignored.
20 | -/
21 | infotree : Option String
22 |
23 | /-- Run Lean commands.
24 | If `env = none`, starts a new session (in which you can use `import`).
25 | If `env = some n`, builds on the existing environment `n`.
26 | -/
27 | structure Command extends CommandOptions where
28 | env : Option Nat
29 | cmd : String
30 | deriving ToJson, FromJson
31 |
32 | /-- Process a Lean file in a fresh environment. -/
33 | structure File extends CommandOptions where
34 | path : System.FilePath
35 | deriving FromJson
36 |
37 | /--
38 | Run a tactic in a proof state.
39 | -/
40 | structure ProofStep where
41 | proofState : Nat
42 | tactic : String
43 | deriving ToJson, FromJson
44 |
45 | /-- Line and column information for error messages and sorries. -/
46 | structure Pos where
47 | line : Nat
48 | column : Nat
49 | deriving ToJson, FromJson
50 |
51 | /-- Severity of a message. -/
52 | inductive Severity
53 | | trace | info | warning | error
54 | deriving ToJson, FromJson
55 |
56 | /-- A Lean message. -/
57 | structure Message where
58 | pos : Pos
59 | endPos : Option Pos
60 | severity : Severity
61 | data : String
62 | deriving ToJson, FromJson
63 |
64 | /-- Construct the JSON representation of a Lean message. -/
65 | def Message.of (m : Lean.Message) : IO Message := do pure <|
66 | { pos := ⟨m.pos.line, m.pos.column⟩,
67 | endPos := m.endPos.map fun p => ⟨p.line, p.column⟩,
68 | severity := match m.severity with
69 | | .information => .info
70 | | .warning => .warning
71 | | .error => .error,
72 | data := (← m.data.toString).trim }
73 |
74 | /-- A Lean `sorry`. -/
75 | structure Sorry where
76 | pos : Pos
77 | endPos : Pos
78 | goal : String
79 | /--
80 | The index of the proof state at the sorry.
81 | You can use the `ProofStep` instruction to run a tactic at this state.
82 | -/
83 | proofState : Option Nat
84 | deriving FromJson
85 |
86 | instance : ToJson Sorry where
87 | toJson r := Json.mkObj <| .flatten [
88 | [("goal", r.goal)],
89 | [("proofState", toJson r.proofState)],
90 | if r.pos.line ≠ 0 then [("pos", toJson r.pos)] else [],
91 | if r.endPos.line ≠ 0 then [("endPos", toJson r.endPos)] else [],
92 | ]
93 |
94 | /-- Construct the JSON representation of a Lean sorry. -/
95 | def Sorry.of (goal : String) (pos endPos : Lean.Position) (proofState : Option Nat) : Sorry :=
96 | { pos := ⟨pos.line, pos.column⟩,
97 | endPos := ⟨endPos.line, endPos.column⟩,
98 | goal,
99 | proofState }
100 |
101 | structure Tactic where
102 | pos : Pos
103 | endPos : Pos
104 | goals : String
105 | tactic : String
106 | proofState : Option Nat
107 | usedConstants : Array Name
108 | deriving ToJson, FromJson
109 |
110 | /-- Construct the JSON representation of a Lean tactic. -/
111 | def Tactic.of (goals tactic : String) (pos endPos : Lean.Position) (proofState : Option Nat) (usedConstants : Array Name) : Tactic :=
112 | { pos := ⟨pos.line, pos.column⟩,
113 | endPos := ⟨endPos.line, endPos.column⟩,
114 | goals,
115 | tactic,
116 | proofState,
117 | usedConstants }
118 |
119 | /--
120 | A response to a Lean command.
121 | `env` can be used in later calls, to build on the stored environment.
122 | -/
123 | structure CommandResponse where
124 | env : Nat
125 | messages : List Message := []
126 | sorries : List Sorry := []
127 | tactics : List Tactic := []
128 | infotree : Option Json := none
129 | deriving FromJson
130 |
131 | def Json.nonemptyList [ToJson α] (k : String) : List α → List (String × Json)
132 | | [] => []
133 | | l => [⟨k, toJson l⟩]
134 |
135 | instance : ToJson CommandResponse where
136 | toJson r := Json.mkObj <| .flatten [
137 | [("env", r.env)],
138 | Json.nonemptyList "messages" r.messages,
139 | Json.nonemptyList "sorries" r.sorries,
140 | Json.nonemptyList "tactics" r.tactics,
141 | match r.infotree with | some j => [("infotree", j)] | none => []
142 | ]
143 |
144 | /--
145 | A response to a Lean tactic.
146 | `proofState` can be used in later calls, to run further tactics.
147 | -/
148 | structure ProofStepResponse where
149 | proofState : Nat
150 | goals : List String
151 | messages : List Message := []
152 | sorries : List Sorry := []
153 | traces : List String
154 | proofStatus : String
155 | deriving ToJson, FromJson
156 |
157 | instance : ToJson ProofStepResponse where
158 | toJson r := Json.mkObj <| .flatten [
159 | [("proofState", r.proofState)],
160 | [("goals", toJson r.goals)],
161 | Json.nonemptyList "messages" r.messages,
162 | Json.nonemptyList "sorries" r.sorries,
163 | Json.nonemptyList "traces" r.traces,
164 | [("proofStatus", r.proofStatus)]
165 | ]
166 |
167 | /-- Json wrapper for an error. -/
168 | structure Error where
169 | message : String
170 | deriving ToJson, FromJson
171 |
172 | structure PickleEnvironment where
173 | env : Nat
174 | pickleTo : System.FilePath
175 | deriving ToJson, FromJson
176 |
177 | structure UnpickleEnvironment where
178 | unpickleEnvFrom : System.FilePath
179 | deriving ToJson, FromJson
180 |
181 | structure PickleProofState where
182 | proofState : Nat
183 | pickleTo : System.FilePath
184 | deriving ToJson, FromJson
185 |
186 | structure UnpickleProofState where
187 | unpickleProofStateFrom : System.FilePath
188 | env : Option Nat
189 | deriving ToJson, FromJson
190 |
191 | end REPL
192 |
--------------------------------------------------------------------------------
/REPL/Lean/ContextInfo.lean:
--------------------------------------------------------------------------------
1 | import Lean
2 |
3 | namespace Lean.Elab.ContextInfo
4 |
5 | /-- Pretty print an expression in the given `ContextInfo` with the given `LocalContext`. -/
6 | def ppExpr (ctx : ContextInfo) (lctx : LocalContext) (e : Expr) : IO Format :=
7 | ctx.runMetaM lctx (do Meta.ppExpr (← instantiateMVars e))
8 |
9 | end Lean.Elab.ContextInfo
10 |
--------------------------------------------------------------------------------
/REPL/Lean/Environment.lean:
--------------------------------------------------------------------------------
1 | import REPL.Util.Pickle
2 | import Lean.Replay
3 |
4 | open System (FilePath)
5 |
6 | namespace Lean.Environment
7 |
8 | /--
9 | Pickle an `Environment` to disk.
10 |
11 | We only store:
12 | * the list of imports
13 | * the new constants from `Environment.constants`
14 | and when unpickling, we build a fresh `Environment` from the imports,
15 | and then add the new constants.
16 | -/
17 | def pickle (env : Environment) (path : FilePath) : IO Unit :=
18 | _root_.pickle path (env.header.imports, env.constants.map₂)
19 |
20 | /--
21 | Unpickle an `Environment` from disk.
22 |
23 | We construct a fresh `Environment` with the relevant imports,
24 | and then replace the new constants.
25 | -/
26 | def unpickle (path : FilePath) : IO (Environment × CompactedRegion) := unsafe do
27 | let ((imports, map₂), region) ← _root_.unpickle (Array Import × PHashMap Name ConstantInfo) path
28 | let env ← importModules imports {} 0 (loadExts := true)
29 | return (← env.replay (Std.HashMap.ofList map₂.toList), region)
30 |
31 | end Lean.Environment
32 |
--------------------------------------------------------------------------------
/REPL/Lean/InfoTree.lean:
--------------------------------------------------------------------------------
1 | /-
2 | Copyright (c) 2023 Scott Morrison. All rights reserved.
3 | Released under Apache 2.0 license as described in the file LICENSE.
4 | Authors: Scott Morrison
5 | -/
6 | import Lean
7 |
8 | /-!
9 | Additional functions to deal with `InfoTree`.
10 | -/
11 |
12 | open Lean Elab Meta
13 |
14 | namespace Lean.FileMap
15 |
16 | /-- Extract the range of a `Syntax` expressed as lines and columns. -/
17 | -- Extracted from the private declaration `Lean.Elab.formatStxRange`,
18 | -- in `Lean.Elab.InfoTree.Main`.
19 | def stxRange (fileMap : FileMap) (stx : Syntax) : Position × Position :=
20 | let pos := stx.getPos?.getD 0
21 | let endPos := stx.getTailPos?.getD pos
22 | (fileMap.toPosition pos, fileMap.toPosition endPos)
23 |
24 | end Lean.FileMap
25 |
26 | namespace Lean.Syntax
27 |
28 | /-- Check if a `Syntax` is an explicit invocation of the `sorry` tactic. -/
29 | def isSorryTactic (stx : Syntax) : Bool :=
30 | s!"{stx}" = "(Tactic.tacticSorry \"sorry\")"
31 |
32 | /-- Check if a `Syntax` is an explicit `sorry` term. -/
33 | def isSorryTerm (stx : Syntax) : Bool :=
34 | s!"{stx}" = "(Term.sorry \"sorry\")"
35 |
36 | end Lean.Syntax
37 |
38 | namespace Lean.Elab
39 |
40 | /-- Extract the range of a `Syntax` expressed as lines and columns. -/
41 | -- Extracted from the private declaration `Lean.Elab.formatStxRange`,
42 | -- in `Lean.Elab.InfoTree.Main`.
43 | def stxRange (fileMap : FileMap) (stx : Syntax) : Position × Position :=
44 | let pos := stx.getPos?.getD 0
45 | let endPos := stx.getTailPos?.getD pos
46 | (fileMap.toPosition pos, fileMap.toPosition endPos)
47 |
48 | end Lean.Elab
49 |
50 | namespace Lean.Elab.Info
51 |
52 | /-- The type of a `Lean.Elab.Info`, as a string. -/
53 | def kind : Info → String
54 | | .ofTacticInfo _ => "TacticInfo"
55 | | .ofTermInfo _ => "TermInfo"
56 | | ofPartialTermInfo _ => "PartialTermInfo"
57 | | .ofCommandInfo _ => "CommandInfo"
58 | | .ofMacroExpansionInfo _ => "MacroExpansionInfo"
59 | | .ofOptionInfo _ => "OptionInfo"
60 | | .ofFieldInfo _ => "FieldInfo"
61 | | .ofCompletionInfo _ => "CompletionInfo"
62 | | .ofUserWidgetInfo _ => "UserWidgetInfo"
63 | | .ofCustomInfo _ => "CustomInfo"
64 | | .ofFVarAliasInfo _ => "FVarAliasInfo"
65 | | .ofFieldRedeclInfo _ => "FieldRedeclInfo"
66 | | .ofChoiceInfo _ => "ChoiceInfo"
67 | | .ofDelabTermInfo _ => "DelabTermInfo"
68 |
69 | /-- The `Syntax` for a `Lean.Elab.Info`, if there is one. -/
70 | def stx? : Info → Option Syntax
71 | | .ofTacticInfo info => info.stx
72 | | .ofTermInfo info => info.stx
73 | | ofPartialTermInfo info => info.stx
74 | | .ofCommandInfo info => info.stx
75 | | .ofMacroExpansionInfo info => info.stx
76 | | .ofOptionInfo info => info.stx
77 | | .ofFieldInfo info => info.stx
78 | | .ofCompletionInfo info => info.stx
79 | | .ofUserWidgetInfo info => info.stx
80 | | .ofCustomInfo info => info.stx
81 | | .ofFVarAliasInfo _ => none
82 | | .ofFieldRedeclInfo info => info.stx
83 | | .ofChoiceInfo info => info.stx
84 | | .ofDelabTermInfo info => info.stx
85 |
86 | /-- Is the `Syntax` for this `Lean.Elab.Info` original, or synthetic? -/
87 | def isOriginal (i : Info) : Bool :=
88 | match i.stx? with
89 | | none => true -- Somewhat unclear what to do with `FVarAliasInfo`, so be conservative.
90 | | some stx => match stx.getHeadInfo with
91 | | .original .. => true
92 | | _ => false
93 |
94 | end Lean.Elab.Info
95 | namespace Lean.Elab.TacticInfo
96 |
97 | /-- Find the name for the outermost `Syntax` in this `TacticInfo`. -/
98 | def name? (t : TacticInfo) : Option Name :=
99 | match t.stx with
100 | | Syntax.node _ n _ => some n
101 | | _ => none
102 |
103 | /-- Decide whether a tactic is "substantive",
104 | or is merely a tactic combinator (e.g. `by`, `;`, multiline tactics, parenthesized tactics). -/
105 | def isSubstantive (t : TacticInfo) : Bool :=
106 | match t.name? with
107 | | none => false
108 | | some `null => false
109 | | some ``cdot => false
110 | | some ``cdotTk => false
111 | | some ``Lean.Parser.Term.byTactic => false
112 | | some ``Lean.Parser.Tactic.tacticSeq => false
113 | | some ``Lean.Parser.Tactic.tacticSeq1Indented => false
114 | | some ``Lean.Parser.Tactic.«tactic_<;>_» => false
115 | | some ``Lean.Parser.Tactic.paren => false
116 | | _ => true
117 |
118 | def getUsedConstantsAsSet (t : TacticInfo) : NameSet :=
119 | t.goalsBefore
120 | |>.filterMap t.mctxAfter.getExprAssignmentCore?
121 | |>.map Expr.getUsedConstantsAsSet
122 | |>.foldl .union .empty
123 |
124 | end Lean.Elab.TacticInfo
125 |
126 | namespace Lean.Elab.InfoTree
127 |
128 | /--
129 | Keep `.node` nodes and `.hole` nodes satisfying predicates.
130 |
131 | Returns a `List InfoTree`, although in most situations this will be a singleton.
132 | -/
133 | partial def filter (p : Info → Bool) (m : MVarId → Bool := fun _ => false) :
134 | InfoTree → List InfoTree
135 | | .context ctx tree => tree.filter p m |>.map (.context ctx)
136 | | .node info children =>
137 | if p info then
138 | [.node info (children.toList.map (filter p m)).flatten.toPArray']
139 | else
140 | (children.toList.map (filter p m)).flatten
141 | | .hole mvar => if m mvar then [.hole mvar] else []
142 |
143 | /-- Discard all nodes besides `.context` nodes and `TacticInfo` nodes. -/
144 | partial def retainTacticInfo (tree : InfoTree) : List InfoTree :=
145 | tree.filter fun | .ofTacticInfo _ => true | _ => false
146 |
147 | /-- Retain only nodes with "original" syntax. -/
148 | partial def retainOriginal (tree : InfoTree) : List InfoTree :=
149 | tree.filter Info.isOriginal
150 |
151 | /-- Discard all TacticInfo nodes that are tactic combinators or structuring tactics. -/
152 | -- There is considerable grey area here: what to do with `classical`?
153 | partial def retainSubstantive (tree : InfoTree) : List InfoTree :=
154 | tree.filter fun | .ofTacticInfo i => i.isSubstantive | _ => true
155 |
156 | /-- Analogue of `Lean.Elab.InfoTree.findInfo?`, but that returns all results. -/
157 | partial def findAllInfo (t : InfoTree) (ctx? : Option ContextInfo) (p : Info → Bool)
158 | (stop : Info → Bool := fun _ => false) :
159 | List (Info × Option ContextInfo) :=
160 | match t with
161 | | context ctx t => t.findAllInfo (ctx.mergeIntoOuter? ctx?) p stop
162 | | node i ts =>
163 | let info := if p i then [(i, ctx?)] else []
164 | let rest := if stop i then [] else ts.toList.flatMap (fun t => t.findAllInfo ctx? p stop)
165 | info ++ rest
166 | | _ => []
167 |
168 | /-- Analogue of `findAllInfo` but specialized to tactics. It additionally returns root goals. -/
169 | partial def findAllInfoTactics (t : InfoTree) (ctx? : Option ContextInfo) (rootGoals : List MVarId := []) :
170 | List (Info × Option ContextInfo × List MVarId) :=
171 | match t with
172 | | .context ctx t => t.findAllInfoTactics (ctx.mergeIntoOuter? ctx?) rootGoals
173 | | .node info ts =>
174 | let tInfo := match info with
175 | | .ofTacticInfo i =>
176 | if info.isOriginal && i.isSubstantive then
177 | let rootGoals := if rootGoals.isEmpty then i.goalsBefore else rootGoals
178 | [(info, ctx?, rootGoals)]
179 | else
180 | []
181 | | _ => []
182 | tInfo ++ ts.toList.flatMap (fun t => t.findAllInfoTactics ctx? rootGoals)
183 | | _ => []
184 |
185 | /-- Return all `TacticInfo` nodes in an `InfoTree` with "original" syntax,
186 | each equipped with its relevant `ContextInfo`. -/
187 | def findTacticNodes (t : InfoTree) : List (TacticInfo × ContextInfo × List MVarId) :=
188 | let infos := t.findAllInfoTactics none
189 | infos.filterMap fun p => match p with
190 | | (.ofTacticInfo i, some ctx, rootGoals) => (i, ctx, rootGoals)
191 | | _ => none
192 |
193 | /-- Returns the root goals for a given `InfoTree`. -/
194 | partial def findRootGoals (t : InfoTree) (ctx? : Option ContextInfo := none) :
195 | List (TacticInfo × ContextInfo × List MVarId) :=
196 | match t with
197 | | .context ctx t => t.findRootGoals (ctx.mergeIntoOuter? ctx?)
198 | | .node info ts =>
199 | match info with
200 | | .ofTacticInfo i =>
201 | match ctx? with
202 | | some ctx => [(i, ctx, i.goalsBefore)]
203 | | _ => []
204 | | _ => ts.toList.flatMap (fun t => t.findRootGoals ctx?)
205 | | _ => []
206 |
207 | /-- Return all `TacticInfo` nodes in an `InfoTree`
208 | corresponding to explicit invocations of the `sorry` tactic,
209 | each equipped with its relevant `ContextInfo`. -/
210 | def findSorryTacticNodes (t : InfoTree) : List (TacticInfo × ContextInfo) :=
211 | let infos := t.findAllInfo none fun i => match i with
212 | | .ofTacticInfo i => i.stx.isSorryTactic && !i.goalsBefore.isEmpty
213 | | _ => false
214 | infos.filterMap fun p => match p with
215 | | (.ofTacticInfo i, some ctx) => (i, ctx)
216 | | _ => none
217 |
218 | /-- Return all `TermInfo` nodes in an `InfoTree`
219 | corresponding to explicit `sorry` terms,
220 | each equipped with its relevant `ContextInfo`. -/
221 | def findSorryTermNodes (t : InfoTree) : List (TermInfo × ContextInfo) :=
222 | let infos := t.findAllInfo none
223 | (fun i => match i with | .ofTermInfo i => i.stx.isSorryTerm | _ => false)
224 | (fun i => match i with | .ofTacticInfo i => i.stx.isSorryTactic | _ => false)
225 | infos.filterMap fun p => match p with
226 | | (.ofTermInfo i, some ctx) => (i, ctx)
227 | | _ => none
228 |
229 | inductive SorryType
230 | | tactic : MVarId → SorryType
231 | | term : LocalContext → Option Expr → SorryType
232 | deriving Inhabited
233 |
234 | /--
235 | Finds all appearances of `sorry` in an `InfoTree`, reporting
236 | * the `ContextInfo` at that point,
237 | * the `MVarId` for a goal that was closed by `sorry`,
238 | or the `Option Expr` expected type for a term supplied by `sorry`
239 | * and the start and end positions of the `sorry` in the file.
240 | -/
241 | def sorries (t : InfoTree) : List (ContextInfo × SorryType × Position × Position) :=
242 | (t.findSorryTacticNodes.map fun ⟨i, ctx⟩ =>
243 | -- HACK: creating a child ngen
244 | ({ ctx with mctx := i.mctxBefore, ngen := ctx.ngen.mkChild.1 }, .tactic i.goalsBefore.head!,
245 | stxRange ctx.fileMap i.stx)) ++
246 | (t.findSorryTermNodes.map fun ⟨i, ctx⟩ =>
247 | (ctx, .term i.lctx i.expectedType?, stxRange ctx.fileMap i.stx))
248 |
249 | def tactics (t : InfoTree) : List (ContextInfo × Syntax × List MVarId × List MVarId × Position × Position × Array Name) :=
250 | -- HACK: creating a child ngen
251 | t.findTacticNodes.map fun ⟨i, ctx, rootGoals⟩ =>
252 | let range := stxRange ctx.fileMap i.stx
253 | ( { ctx with mctx := i.mctxBefore, ngen := ctx.ngen.mkChild.1 },
254 | i.stx,
255 | rootGoals,
256 | i.goalsBefore,
257 | range.fst,
258 | range.snd,
259 | i.getUsedConstantsAsSet.toArray )
260 |
261 | def rootGoals (t : InfoTree) : List (ContextInfo × List MVarId × Position) :=
262 | t.findRootGoals.map fun ⟨i, ctx, rootGoals⟩ =>
263 | let range := stxRange ctx.fileMap i.stx
264 | ( { ctx with mctx := i.mctxBefore, ngen := ctx.ngen.mkChild.1 },
265 | rootGoals,
266 | range.fst )
267 |
268 | end Lean.Elab.InfoTree
269 |
270 | namespace Lean.Elab.TacticInfo
271 |
272 | /-- Return the range of the tactic, as a pair of file positions. -/
273 | def range (info : TacticInfo) (ctx : ContextInfo) : Position × Position := ctx.fileMap.stxRange info.stx
274 |
275 | /-- Pretty print a tactic. -/
276 | def pp (info : TacticInfo) (ctx : ContextInfo) : IO Format :=
277 | ctx.runMetaM {} try
278 | Lean.PrettyPrinter.ppTactic ⟨info.stx⟩
279 | catch _ =>
280 | pure ""
281 |
282 | open Meta
283 |
284 | /-- Run a tactic on the goals stored in a `TacticInfo`. -/
285 | def runMetaMGoalsBefore (info : TacticInfo) (ctx : ContextInfo) (x : List MVarId → MetaM α) : IO α := do
286 | ctx.runMetaM {} <| Meta.withMCtx info.mctxBefore <| x info.goalsBefore
287 |
288 | /-- Run a tactic on the after goals stored in a `TacticInfo`. -/
289 | def runMetaMGoalsAfter (info : TacticInfo) (ctx : ContextInfo) (x : List MVarId → MetaM α) : IO α := do
290 | ctx.runMetaM {} <| Meta.withMCtx info.mctxAfter <| x info.goalsAfter
291 |
292 | /-- Run a tactic on the main goal stored in a `TacticInfo`. -/
293 | def runMetaM (info : TacticInfo) (ctx : ContextInfo) (x : MVarId → MetaM α) : IO α := do
294 | match info.goalsBefore.head? with
295 | | none => throw <| IO.userError s!"No goals at {← info.pp ctx}"
296 | | some g => info.runMetaMGoalsBefore ctx fun _ => do g.withContext <| x g
297 |
298 | def mainGoal (info : TacticInfo) (ctx : ContextInfo) : IO Expr :=
299 | info.runMetaM ctx (fun g => do instantiateMVars (← g.getType))
300 |
301 | def formatMainGoal (info : TacticInfo) (ctx : ContextInfo) : IO Format :=
302 | info.runMetaM ctx (fun g => do ppExpr (← instantiateMVars (← g.getType)))
303 |
304 | def goalState (info : TacticInfo) (ctx : ContextInfo) : IO (List Format) := do
305 | info.runMetaMGoalsBefore ctx (fun gs => gs.mapM fun g => do Meta.ppGoal g)
306 |
307 | def goalStateAfter (info : TacticInfo) (ctx : ContextInfo) : IO (List Format) := do
308 | info.runMetaMGoalsAfter ctx (fun gs => gs.mapM fun g => do Meta.ppGoal g)
309 |
310 | def ppExpr (info : TacticInfo) (ctx : ContextInfo) (e : Expr) : IO Format :=
311 | info.runMetaM ctx (fun _ => do Meta.ppExpr (← instantiateMVars e))
312 |
313 | end Lean.Elab.TacticInfo
314 |
315 | namespace Lean.Elab.InfoTree
316 |
317 | /--
318 | Finds all tactic invocations in an `InfoTree`,
319 | ignoring structuring tactics (e.g. `by`, `;`, multiline tactics, parenthesized tactics).
320 | -/
321 | def substantiveTactics (t : InfoTree) : List (TacticInfo × ContextInfo × List MVarId) :=
322 | t.findTacticNodes.filter fun i => i.1.isSubstantive
323 |
324 | end Lean.Elab.InfoTree
325 |
--------------------------------------------------------------------------------
/REPL/Lean/InfoTree/ToJson.lean:
--------------------------------------------------------------------------------
1 | import REPL.Lean.InfoTree
2 | import REPL.Lean.ContextInfo
3 |
4 | /-!
5 | # Exporting an `InfoTree` as Json
6 |
7 | -/
8 |
9 | namespace Lean.Elab
10 |
11 | structure InfoTreeNode (α : Type) where
12 | kind : String
13 | node : Option α
14 | children : List Json
15 | deriving ToJson
16 |
17 | deriving instance ToJson for Lean.Position
18 |
19 | structure Syntax.Range where
20 | synthetic : Bool
21 | start : Lean.Position
22 | finish : Lean.Position
23 | deriving ToJson
24 |
25 | structure Syntax.Json where
26 | pp : Option String
27 | -- raw : String
28 | range : Range
29 | deriving ToJson
30 |
31 | def _root_.Lean.Syntax.toRange (stx : Syntax) (ctx : ContextInfo) : Syntax.Range :=
32 | let pos := stx.getPos?.getD 0
33 | let endPos := stx.getTailPos?.getD pos
34 | { start := ctx.fileMap.toPosition pos
35 | finish := ctx.fileMap.toPosition endPos
36 | synthetic := match stx.getHeadInfo with
37 | | .original .. => false
38 | | _ => true }
39 |
40 | def _root_.Lean.Syntax.toJson (stx : Syntax) (ctx : ContextInfo) (lctx : LocalContext) : IO Syntax.Json := do
41 | return {
42 | pp := match (← ctx.ppSyntax lctx stx).pretty with
43 | | "failed to pretty print term (use 'set_option pp.rawOnError true' for raw representation)" => none
44 | | pp => some pp
45 | -- raw := toString stx
46 | range := stx.toRange ctx }
47 |
48 | structure TacticInfo.Json where
49 | name : Option Name
50 | stx : Syntax.Json
51 | goalsBefore : List String
52 | goalsAfter : List String
53 | deriving ToJson
54 |
55 | -- Note: this is not responsible for converting the children to Json.
56 | def TacticInfo.toJson (i : TacticInfo) (ctx : ContextInfo) : IO TacticInfo.Json := do
57 | return {
58 | name := i.name?
59 | stx :=
60 | { pp := Format.pretty (← i.pp ctx),
61 | -- raw := toString i.info.stx,
62 | range := i.stx.toRange ctx },
63 | goalsBefore := (← i.goalState ctx).map Format.pretty,
64 | goalsAfter := (← i.goalStateAfter ctx).map Format.pretty }
65 |
66 | structure CommandInfo.Json where
67 | elaborator : Option Name
68 | stx : Syntax.Json
69 | deriving ToJson
70 |
71 | def CommandInfo.toJson (info : CommandInfo) (ctx : ContextInfo) : IO CommandInfo.Json := do
72 | return {
73 | elaborator := match info.elaborator with | .anonymous => none | n => some n,
74 | stx := ← info.stx.toJson ctx {} }
75 |
76 | structure TermInfo.Json where
77 | elaborator : Option Name
78 | stx : Syntax.Json
79 | expectedType? : Option String
80 | expr : String
81 | isBinder : Bool
82 | deriving ToJson
83 |
84 | def TermInfo.toJson (info : TermInfo) (ctx : ContextInfo) : IO TermInfo.Json := do
85 | return {
86 | elaborator := match info.elaborator with | .anonymous => none | n => some n,
87 | stx := ← info.stx.toJson ctx info.lctx,
88 | expectedType? := ← info.expectedType?.mapM fun ty => do
89 | pure (← ctx.ppExpr info.lctx ty).pretty
90 | expr := (← ctx.ppExpr info.lctx info.expr).pretty
91 | isBinder := info.isBinder }
92 |
93 | structure InfoTree.HoleJson where
94 | goalState : String
95 | deriving ToJson
96 |
97 | partial def InfoTree.toJson (t : InfoTree) (ctx? : Option ContextInfo) : IO Json := do
98 | match t with
99 | | .context ctx t => t.toJson (ctx.mergeIntoOuter? ctx?)
100 | | .node info children =>
101 | if let some ctx := ctx? then
102 | let node : Option Json ← match info with
103 | | .ofTermInfo info => some <$> (do pure <| Lean.toJson (← info.toJson ctx))
104 | | .ofCommandInfo info => some <$> (do pure <| Lean.toJson (← info.toJson ctx))
105 | | .ofTacticInfo info => some <$> (do pure <| Lean.toJson (← info.toJson ctx))
106 | | _ => pure none
107 | return Lean.toJson (InfoTreeNode.mk info.kind node (← children.toList.mapM fun t' => t'.toJson ctx))
108 | else throw <| IO.userError "No `ContextInfo` available."
109 | | .hole mvarId =>
110 | if let some ctx := ctx? then
111 | return Lean.toJson (InfoTree.HoleJson.mk (← ctx.runMetaM {} (do Meta.ppGoal mvarId)).pretty)
112 | else throw <| IO.userError "No `ContextInfo` available."
113 |
114 | end Lean.Elab
115 |
--------------------------------------------------------------------------------
/REPL/Main.lean:
--------------------------------------------------------------------------------
1 | /-
2 | Copyright (c) 2023 Scott Morrison. All rights reserved.
3 | Released under Apache 2.0 license as described in the file LICENSE.
4 | Authors: Scott Morrison
5 | -/
6 | import REPL.JSON
7 | import REPL.Frontend
8 | import REPL.Util.Path
9 | import REPL.Lean.ContextInfo
10 | import REPL.Lean.Environment
11 | import REPL.Lean.InfoTree
12 | import REPL.Lean.InfoTree.ToJson
13 | import REPL.Snapshots
14 |
15 | /-!
16 | # A REPL for Lean.
17 |
18 | Communicates via JSON on stdin and stdout. Commands should be separated by blank lines.
19 |
20 | Commands may be of the form
21 | ```
22 | { "cmd" : "import Mathlib.Data.List.Basic\ndef f := 2" }
23 | ```
24 | or
25 | ```
26 | { "cmd" : "example : f = 2 := rfl", "env" : 3 }
27 | ```
28 |
29 | The `env` field, if present,
30 | must contain a number received in the `env` field of a previous response,
31 | and causes the command to be run in the existing environment.
32 |
33 | If there is no `env` field, a new environment is created.
34 |
35 | You can only use `import` commands when you do not specify the `env` field.
36 |
37 | You can backtrack simply by using earlier values for `env`.
38 |
39 | The results are of the form
40 | ```
41 | {"sorries":
42 | [{"pos": {"line": 1, "column": 18},
43 | "endPos": {"line": 1, "column": 23},
44 | "goal": "\n⊢ Nat"}],
45 | "messages":
46 | [{"severity": "error",
47 | "pos": {"line": 1, "column": 23},
48 | "endPos": {"line": 1, "column": 26},
49 | "data":
50 | "type mismatch\n rfl\nhas type\n f = f : Prop\nbut is expected to have type\n f = 2 : Prop"}],
51 | "env": 6}
52 | ```
53 | showing any messages generated, or sorries with their goal states.
54 | Information is generated for tactic mode sorries, but not for term mode sorries.
55 | -/
56 |
57 | open Lean Elab
58 |
59 | namespace REPL
60 |
61 | /-- The monadic state for the Lean REPL. -/
62 | structure State where
63 | /--
64 | Environment snapshots after complete declarations.
65 | The user can run a declaration in a given environment using `{"cmd": "def f := 37", "env": 17}`.
66 | -/
67 | cmdStates : Array CommandSnapshot := #[]
68 | /--
69 | Proof states after individual tactics.
70 | The user can run a tactic in a given proof state using `{"tactic": "exact 42", "proofState": 5}`.
71 | Declarations with containing `sorry` record a proof state at each sorry,
72 | and report the numerical index for the recorded state at each sorry.
73 | -/
74 | proofStates : Array ProofSnapshot := #[]
75 |
76 | /--
77 | The Lean REPL monad.
78 |
79 | We only use this with `m := IO`, but it is set up as a monad transformer for flexibility.
80 | -/
81 | abbrev M (m : Type → Type) := StateT State m
82 |
83 | variable [Monad m] [MonadLiftT IO m]
84 |
85 | /-- Record an `CommandSnapshot` into the REPL state, returning its index for future use. -/
86 | def recordCommandSnapshot (state : CommandSnapshot) : M m Nat := do
87 | let id := (← get).cmdStates.size
88 | modify fun s => { s with cmdStates := s.cmdStates.push state }
89 | return id
90 |
91 | /-- Record a `ProofSnapshot` into the REPL state, returning its index for future use. -/
92 | def recordProofSnapshot (proofState : ProofSnapshot) : M m Nat := do
93 | let id := (← get).proofStates.size
94 | modify fun s => { s with proofStates := s.proofStates.push proofState }
95 | return id
96 |
97 | def sorries (trees : List InfoTree) (env? : Option Environment) (rootGoals? : Option (List MVarId))
98 | : M m (List Sorry) :=
99 | trees.flatMap InfoTree.sorries |>.filter (fun t => match t.2.1 with
100 | | .term _ none => false
101 | | _ => true ) |>.mapM
102 | fun ⟨ctx, g, pos, endPos⟩ => do
103 | let (goal, proofState) ← match g with
104 | | .tactic g => do
105 | let lctx ← ctx.runMetaM {} do
106 | match ctx.mctx.findDecl? g with
107 | | some decl => return decl.lctx
108 | | none => throwError "unknown metavariable '{g}'"
109 | let s ← ProofSnapshot.create ctx lctx env? [g] rootGoals?
110 | pure ("\n".intercalate <| (← s.ppGoals).map fun s => s!"{s}", some s)
111 | | .term lctx (some t) => do
112 | let s ← ProofSnapshot.create ctx lctx env? [] rootGoals? [t]
113 | pure ("\n".intercalate <| (← s.ppGoals).map fun s => s!"{s}", some s)
114 | | .term _ none => unreachable!
115 | let proofStateId ← proofState.mapM recordProofSnapshot
116 | return Sorry.of goal pos endPos proofStateId
117 |
118 | def ppTactic (ctx : ContextInfo) (stx : Syntax) : IO Format :=
119 | ctx.runMetaM {} try
120 | Lean.PrettyPrinter.ppTactic ⟨stx⟩
121 | catch _ =>
122 | pure ""
123 |
124 | def tactics (trees : List InfoTree) (env? : Option Environment) : M m (List Tactic) :=
125 | trees.flatMap InfoTree.tactics |>.mapM
126 | fun ⟨ctx, stx, rootGoals, goals, pos, endPos, ns⟩ => do
127 | let proofState := some (← ProofSnapshot.create ctx none env? goals rootGoals)
128 | let goals := s!"{(← ctx.ppGoals goals)}".trim
129 | let tactic := Format.pretty (← ppTactic ctx stx)
130 | let proofStateId ← proofState.mapM recordProofSnapshot
131 | return Tactic.of goals tactic pos endPos proofStateId ns
132 |
133 | def collectRootGoalsAsSorries (trees : List InfoTree) (env? : Option Environment) : M m (List Sorry) := do
134 | trees.flatMap InfoTree.rootGoals |>.mapM
135 | fun ⟨ctx, goals, pos⟩ => do
136 | let proofState := some (← ProofSnapshot.create ctx none env? goals goals)
137 | let goals := s!"{(← ctx.ppGoals goals)}".trim
138 | let proofStateId ← proofState.mapM recordProofSnapshot
139 | return Sorry.of goals pos pos proofStateId
140 |
141 |
142 | private def collectFVarsAux : Expr → NameSet
143 | | .fvar fvarId => NameSet.empty.insert fvarId.name
144 | | .app fm arg => (collectFVarsAux fm).union $ collectFVarsAux arg
145 | | .lam _ binderType body _ => (collectFVarsAux binderType).union $ collectFVarsAux body
146 | | .forallE _ binderType body _ => (collectFVarsAux binderType).union $ collectFVarsAux body
147 | | .letE _ type value body _ => ((collectFVarsAux type).union $ collectFVarsAux value).union $ collectFVarsAux body
148 | | .mdata _ expr => collectFVarsAux expr
149 | | .proj _ _ struct => collectFVarsAux struct
150 | | _ => NameSet.empty
151 |
152 | /-- Collect all fvars in the expression, and return their names. -/
153 | private def collectFVars (e : Expr) : MetaM (Array Expr) := do
154 | let names := collectFVarsAux e
155 | let mut fvars := #[]
156 | for ldecl in ← getLCtx do
157 | if ldecl.isImplementationDetail then
158 | continue
159 | if names.contains ldecl.fvarId.name then
160 | fvars := fvars.push $ .fvar ldecl.fvarId
161 | return fvars
162 |
163 |
164 | private def abstractAllLambdaFVars (e : Expr) : MetaM Expr := do
165 | let mut e' := e
166 | while e'.hasFVar do
167 | let fvars ← collectFVars e'
168 | if fvars.isEmpty then
169 | break
170 | e' ← Meta.mkLambdaFVars fvars e'
171 | return e'
172 |
173 | /--
174 | Evaluates the current status of a proof, returning a string description.
175 | Main states include:
176 | - "Completed": Proof is complete and type checks successfully
177 | - "Incomplete": When goals remain, or proof contains sorry/metavariables
178 | - "Error": When kernel type checking errors occur
179 |
180 | Inspired by LeanDojo REPL's status tracking.
181 | -/
182 | def getProofStatus (proofState : ProofSnapshot) : M m String := do
183 | match proofState.tacticState.goals with
184 | | [] =>
185 | let res := proofState.runMetaM do
186 | match proofState.rootGoals with
187 | | [goalId] =>
188 | match proofState.metaState.mctx.getExprAssignmentCore? goalId with
189 | | none => return "Error: Goal not assigned"
190 | | some pf => do
191 | let pf ← instantiateMVars pf
192 |
193 | -- First check that the proof has the expected type
194 | let pft ← Meta.inferType pf >>= instantiateMVars
195 | let expectedType ← Meta.inferType (mkMVar goalId) >>= instantiateMVars
196 | unless (← Meta.isDefEq pft expectedType) do
197 | return s!"Error: proof has type {pft} but root goal has type {expectedType}"
198 |
199 | let pf ← goalId.withContext $ abstractAllLambdaFVars pf
200 | let pft ← Meta.inferType pf >>= instantiateMVars
201 |
202 | if pf.hasExprMVar then
203 | return "Incomplete: contains metavariable(s)"
204 |
205 | -- Find all level parameters
206 | let usedLevels := collectLevelParams {} pft
207 | let usedLevels := collectLevelParams usedLevels pf
208 |
209 | let decl := Declaration.defnDecl {
210 | name := Name.anonymous,
211 | type := pft,
212 | value := pf,
213 | levelParams := usedLevels.params.toList,
214 | hints := ReducibilityHints.opaque,
215 | safety := DefinitionSafety.safe
216 | }
217 |
218 | try
219 | let _ ← addDecl decl
220 | catch ex =>
221 | return s!"Error: kernel type check failed: {← ex.toMessageData.toString}"
222 |
223 | if pf.hasSorry then
224 | return "Incomplete: contains sorry"
225 | return "Completed"
226 |
227 | | _ => return "Not verified: more than one initial goal"
228 | return (← res).fst
229 |
230 | | _ => return "Incomplete: open goals remain"
231 |
232 | /-- Record a `ProofSnapshot` and generate a JSON response for it. -/
233 | def createProofStepReponse (proofState : ProofSnapshot) (old? : Option ProofSnapshot := none) :
234 | M m ProofStepResponse := do
235 | let messages := proofState.newMessages old?
236 | let messages ← messages.mapM fun m => Message.of m
237 | let traces ← proofState.newTraces old?
238 | let trees := proofState.newInfoTrees old?
239 | let trees ← match old? with
240 | | some old => do
241 | let (ctx, _) ← old.runMetaM do return { ← CommandContextInfo.save with }
242 | let ctx := PartialContextInfo.commandCtx ctx
243 | pure <| trees.map fun t => InfoTree.context ctx t
244 | | none => pure trees
245 | -- For debugging purposes, sometimes we print out the trees here:
246 | -- trees.forM fun t => do IO.println (← t.format)
247 | let sorries ← sorries trees none (some proofState.rootGoals)
248 | let id ← recordProofSnapshot proofState
249 | return {
250 | proofState := id
251 | goals := (← proofState.ppGoals).map fun s => s!"{s}"
252 | messages
253 | sorries
254 | traces
255 | proofStatus := (← getProofStatus proofState) }
256 |
257 | /-- Pickle a `CommandSnapshot`, generating a JSON response. -/
258 | def pickleCommandSnapshot (n : PickleEnvironment) : M m (CommandResponse ⊕ Error) := do
259 | match (← get).cmdStates[n.env]? with
260 | | none => return .inr ⟨"Unknown environment."⟩
261 | | some env =>
262 | discard <| env.pickle n.pickleTo
263 | return .inl { env := n.env }
264 |
265 | /-- Unpickle a `CommandSnapshot`, generating a JSON response. -/
266 | def unpickleCommandSnapshot (n : UnpickleEnvironment) : M IO CommandResponse := do
267 | let (env, _) ← CommandSnapshot.unpickle n.unpickleEnvFrom
268 | let env ← recordCommandSnapshot env
269 | return { env }
270 |
271 | /-- Pickle a `ProofSnapshot`, generating a JSON response. -/
272 | -- This generates a new identifier, which perhaps is not what we want?
273 | def pickleProofSnapshot (n : PickleProofState) : M m (ProofStepResponse ⊕ Error) := do
274 | match (← get).proofStates[n.proofState]? with
275 | | none => return .inr ⟨"Unknown proof State."⟩
276 | | some proofState =>
277 | discard <| proofState.pickle n.pickleTo
278 | return .inl (← createProofStepReponse proofState)
279 |
280 | /-- Unpickle a `ProofSnapshot`, generating a JSON response. -/
281 | def unpickleProofSnapshot (n : UnpickleProofState) : M IO (ProofStepResponse ⊕ Error) := do
282 | let (cmdSnapshot?, notFound) ← do match n.env with
283 | | none => pure (none, false)
284 | | some i => do match (← get).cmdStates[i]? with
285 | | some env => pure (some env, false)
286 | | none => pure (none, true)
287 | if notFound then
288 | return .inr ⟨"Unknown environment."⟩
289 | let (proofState, _) ← ProofSnapshot.unpickle n.unpickleProofStateFrom cmdSnapshot?
290 | Sum.inl <$> createProofStepReponse proofState
291 |
292 | /--
293 | Run a command, returning the id of the new environment, and any messages and sorries.
294 | -/
295 | def runCommand (s : Command) : M IO (CommandResponse ⊕ Error) := do
296 | let (cmdSnapshot?, notFound) ← do match s.env with
297 | | none => pure (none, false)
298 | | some i => do match (← get).cmdStates[i]? with
299 | | some env => pure (some env, false)
300 | | none => pure (none, true)
301 | if notFound then
302 | return .inr ⟨"Unknown environment."⟩
303 | let initialCmdState? := cmdSnapshot?.map fun c => c.cmdState
304 | let (initialCmdState, cmdState, messages, trees) ← try
305 | IO.processInput s.cmd initialCmdState?
306 | catch ex =>
307 | return .inr ⟨ex.toString⟩
308 | let messages ← messages.mapM fun m => Message.of m
309 | -- For debugging purposes, sometimes we print out the trees here:
310 | -- trees.forM fun t => do IO.println (← t.format)
311 | let sorries ← sorries trees initialCmdState.env none
312 | let sorries ← match s.rootGoals with
313 | | some true => pure (sorries ++ (← collectRootGoalsAsSorries trees initialCmdState.env))
314 | | _ => pure sorries
315 | let tactics ← match s.allTactics with
316 | | some true => tactics trees initialCmdState.env
317 | | _ => pure []
318 | let cmdSnapshot :=
319 | { cmdState
320 | cmdContext := (cmdSnapshot?.map fun c => c.cmdContext).getD
321 | { fileName := "",
322 | fileMap := default,
323 | snap? := none,
324 | cancelTk? := none } }
325 | let env ← recordCommandSnapshot cmdSnapshot
326 | let jsonTrees := match s.infotree with
327 | | some "full" => trees
328 | | some "tactics" => trees.flatMap InfoTree.retainTacticInfo
329 | | some "original" => trees.flatMap InfoTree.retainTacticInfo |>.flatMap InfoTree.retainOriginal
330 | | some "substantive" => trees.flatMap InfoTree.retainTacticInfo |>.flatMap InfoTree.retainSubstantive
331 | | _ => []
332 | let infotree ← if jsonTrees.isEmpty then
333 | pure none
334 | else
335 | pure <| some <| Json.arr (← jsonTrees.toArray.mapM fun t => t.toJson none)
336 | return .inl
337 | { env,
338 | messages,
339 | sorries,
340 | tactics
341 | infotree }
342 |
343 | def processFile (s : File) : M IO (CommandResponse ⊕ Error) := do
344 | try
345 | let cmd ← IO.FS.readFile s.path
346 | runCommand { s with env := none, cmd }
347 | catch e =>
348 | pure <| .inr ⟨e.toString⟩
349 |
350 | /--
351 | Run a single tactic, returning the id of the new proof statement, and the new goals.
352 | -/
353 | -- TODO detect sorries?
354 | def runProofStep (s : ProofStep) : M IO (ProofStepResponse ⊕ Error) := do
355 | match (← get).proofStates[s.proofState]? with
356 | | none => return .inr ⟨"Unknown proof state."⟩
357 | | some proofState =>
358 | try
359 | let proofState' ← proofState.runString s.tactic
360 | return .inl (← createProofStepReponse proofState' proofState)
361 | catch ex =>
362 | return .inr ⟨"Lean error:\n" ++ ex.toString⟩
363 |
364 | end REPL
365 |
366 | open REPL
367 |
368 | /-- Get lines from stdin until a blank line is entered. -/
369 | partial def getLines : IO String := do
370 | let line ← (← IO.getStdin).getLine
371 | if line.trim.isEmpty then
372 | return line
373 | else
374 | return line ++ (← getLines)
375 |
376 | instance [ToJson α] [ToJson β] : ToJson (α ⊕ β) where
377 | toJson x := match x with
378 | | .inl a => toJson a
379 | | .inr b => toJson b
380 |
381 | /-- Commands accepted by the REPL. -/
382 | inductive Input
383 | | command : REPL.Command → Input
384 | | file : REPL.File → Input
385 | | proofStep : REPL.ProofStep → Input
386 | | pickleEnvironment : REPL.PickleEnvironment → Input
387 | | unpickleEnvironment : REPL.UnpickleEnvironment → Input
388 | | pickleProofSnapshot : REPL.PickleProofState → Input
389 | | unpickleProofSnapshot : REPL.UnpickleProofState → Input
390 |
391 | /-- Parse a user input string to an input command. -/
392 | def parse (query : String) : IO Input := do
393 | let json := Json.parse query
394 | match json with
395 | | .error e => throw <| IO.userError <| toString <| toJson <|
396 | (⟨"Could not parse JSON:\n" ++ e⟩ : Error)
397 | | .ok j => match fromJson? j with
398 | | .ok (r : REPL.ProofStep) => return .proofStep r
399 | | .error _ => match fromJson? j with
400 | | .ok (r : REPL.PickleEnvironment) => return .pickleEnvironment r
401 | | .error _ => match fromJson? j with
402 | | .ok (r : REPL.UnpickleEnvironment) => return .unpickleEnvironment r
403 | | .error _ => match fromJson? j with
404 | | .ok (r : REPL.PickleProofState) => return .pickleProofSnapshot r
405 | | .error _ => match fromJson? j with
406 | | .ok (r : REPL.UnpickleProofState) => return .unpickleProofSnapshot r
407 | | .error _ => match fromJson? j with
408 | | .ok (r : REPL.Command) => return .command r
409 | | .error _ => match fromJson? j with
410 | | .ok (r : REPL.File) => return .file r
411 | | .error e => throw <| IO.userError <| toString <| toJson <|
412 | (⟨"Could not parse as a valid JSON command:\n" ++ e⟩ : Error)
413 |
414 | /-- Avoid buffering the output. -/
415 | def printFlush [ToString α] (s : α) : IO Unit := do
416 | let out ← IO.getStdout
417 | out.putStr (toString s)
418 | out.flush -- Flush the output
419 |
420 | /-- Read-eval-print loop for Lean. -/
421 | unsafe def repl : IO Unit :=
422 | StateT.run' loop {}
423 | where loop : M IO Unit := do
424 | let query ← getLines
425 | if query = "" then
426 | return ()
427 | if query.startsWith "#" || query.startsWith "--" then loop else
428 | IO.println <| toString <| ← match ← parse query with
429 | | .command r => return toJson (← runCommand r)
430 | | .file r => return toJson (← processFile r)
431 | | .proofStep r => return toJson (← runProofStep r)
432 | | .pickleEnvironment r => return toJson (← pickleCommandSnapshot r)
433 | | .unpickleEnvironment r => return toJson (← unpickleCommandSnapshot r)
434 | | .pickleProofSnapshot r => return toJson (← pickleProofSnapshot r)
435 | | .unpickleProofSnapshot r => return toJson (← unpickleProofSnapshot r)
436 | printFlush "\n" -- easier to parse the output if there are blank lines
437 | loop
438 |
439 | /-- Main executable function, run as `lake exe repl`. -/
440 | unsafe def main (_ : List String) : IO Unit := do
441 | initSearchPath (← Lean.findSysroot)
442 | repl
443 |
--------------------------------------------------------------------------------
/REPL/Snapshots.lean:
--------------------------------------------------------------------------------
1 | /-
2 | Copyright (c) 2023 Lean FRO, LLC. All rights reserved.
3 | Released under Apache 2.0 license as described in the file LICENSE.
4 | Authors: Scott Morrison
5 | -/
6 | import Lean.Replay
7 | import Lean.Elab.Command
8 | import REPL.Util.Pickle
9 |
10 | open Lean Elab
11 |
12 | namespace Lean.Elab.Command
13 |
14 | @[inline] def CommandElabM.run (x : CommandElabM α) (ctx : Context) (s : State) : EIO Exception (α × State) :=
15 | (x ctx).run s
16 |
17 | @[inline] def CommandElabM.run' (x : CommandElabM α) (ctx : Context) (s : State) : EIO Exception α :=
18 | Prod.fst <$> x.run ctx s
19 |
20 | @[inline] def CommandElabM.toIO (x : CommandElabM α) (ctx : Context) (s : State) : IO (α × State) := do
21 | match (← (x.run ctx s).toIO') with
22 | | Except.error (Exception.error _ msg) => throw <| IO.userError (← msg.toString)
23 | | Except.error (Exception.internal id _) => throw <| IO.userError <| "internal exception #" ++ toString id.idx
24 | | Except.ok a => return a
25 |
26 | end Lean.Elab.Command
27 |
28 | namespace REPL
29 |
30 | /--
31 | Bundled structure for the `State` and `Context` objects
32 | for the `CommandElabM` monad.
33 | -/
34 | structure CommandSnapshot where
35 | cmdState : Command.State
36 | cmdContext : Command.Context
37 |
38 | namespace CommandSnapshot
39 |
40 | open Lean.Elab.Command
41 |
42 | /-- A copy of `Command.State` with the `Environment`, caches, and logging omitted. -/
43 | structure CompactableCommandSnapshot where
44 | -- env : Environment
45 | scopes : List Scope := [{ header := "" }]
46 | nextMacroScope : Nat := firstFrontendMacroScope + 1
47 | maxRecDepth : Nat
48 | nextInstIdx : Nat := 1 -- for generating anonymous instance names
49 | ngen : NameGenerator := {}
50 | -- infoState : InfoState := {}
51 | -- traceState : TraceState := {}
52 | -- messages : MessageLog := {}
53 |
54 | open System (FilePath)
55 |
56 | /--
57 | Run a `CommandElabM` monadic function in the current `ProofSnapshot`,
58 | updating the `Command.State`.
59 | -/
60 | def runCommandElabM (p : CommandSnapshot) (t : CommandElabM α) : IO (α × CommandSnapshot) := do
61 | let (a, cmdState) ← (CommandElabM.toIO · p.cmdContext p.cmdState) do t
62 | return (a, { p with cmdState })
63 |
64 |
65 | /--
66 | Pickle a `CommandSnapshot`, discarding closures and non-essential caches.
67 |
68 | When pickling the `Environment`, we do so relative to its imports.
69 | -/
70 | def pickle (p : CommandSnapshot) (path : FilePath) : IO Unit := do
71 | let env := p.cmdState.env
72 | let p' := { p with cmdState := { p.cmdState with env := ← mkEmptyEnvironment }}
73 | _root_.pickle path
74 | (env.header.imports,
75 | env.constants.map₂,
76 | ({ p'.cmdState with } : CompactableCommandSnapshot),
77 | p'.cmdContext)
78 |
79 | /--
80 | Unpickle a `CommandSnapshot`.
81 | -/
82 | def unpickle (path : FilePath) : IO (CommandSnapshot × CompactedRegion) := unsafe do
83 | let ((imports, map₂, cmdState, cmdContext), region) ←
84 | _root_.unpickle (Array Import × PHashMap Name ConstantInfo × CompactableCommandSnapshot ×
85 | Command.Context) path
86 | let env ← (← importModules imports {} 0 (loadExts := true)).replay (Std.HashMap.ofList map₂.toList)
87 | let p' : CommandSnapshot :=
88 | { cmdState := { cmdState with env }
89 | cmdContext }
90 | let (_, p'') ← p'.runCommandElabM do
91 | for o in ← getOpenDecls do
92 | if let .simple ns _ := o then do
93 | activateScoped ns
94 | return (p'', region)
95 |
96 | end CommandSnapshot
97 |
98 | /--
99 | Bundled structure for the `State` and `Context` objects
100 | for the `CoreM`, `MetaM`, `TermElabM`, and `TacticM` monads.
101 | -/
102 | structure ProofSnapshot where
103 | coreState : Core.State
104 | coreContext : Core.Context
105 | metaState : Meta.State
106 | metaContext : Meta.Context
107 | termState : Term.State
108 | termContext : Term.Context
109 | tacticState : Tactic.State
110 | tacticContext : Tactic.Context
111 | rootGoals : List MVarId
112 |
113 | namespace ProofSnapshot
114 |
115 | open Lean Elab Tactic
116 |
117 | /-- New messages in a `ProofSnapshot`, relative to an optional previous `ProofSnapshot`. -/
118 | def newMessages (new : ProofSnapshot) (old? : Option ProofSnapshot := none) : List Lean.Message :=
119 | match old? with
120 | | none => new.coreState.messages.toList
121 | | some old => new.coreState.messages.toList.drop (old.coreState.messages.toList.length)
122 |
123 | /-- New info trees in a `ProofSnapshot`, relative to an optional previous `ProofSnapshot`. -/
124 | def newInfoTrees (new : ProofSnapshot) (old? : Option ProofSnapshot := none) : List InfoTree :=
125 | let infoState := new.coreState.infoState
126 | let trees := match old? with
127 | | none => infoState.trees.toList
128 | | some old => infoState.trees.toList.drop (old.coreState.infoState.trees.size)
129 | trees.map fun t => t.substitute infoState.assignment
130 |
131 | /-- Run a `CoreM` monadic function in the current `ProofSnapshot`, updating the `Core.State`. -/
132 | def runCoreM (p : ProofSnapshot) (t : CoreM α) : IO (α × ProofSnapshot) := do
133 | let (a, coreState) ← (Lean.Core.CoreM.toIO · p.coreContext p.coreState) do t
134 | return (a, { p with coreState })
135 |
136 | /-- Run a `MetaM` monadic function in the current `ProofSnapshot`, updating the `Meta.State`. -/
137 | def runMetaM (p : ProofSnapshot) (t : MetaM α) : IO (α × ProofSnapshot) := do
138 | let ((a, metaState), p') ←
139 | p.runCoreM (Lean.Meta.MetaM.run (ctx := p.metaContext) (s := p.metaState) do t)
140 | return (a, { p' with metaState })
141 |
142 | /-- Run a `TermElabM` monadic function in the current `ProofSnapshot`, updating the `Term.State`. -/
143 | def runTermElabM (p : ProofSnapshot) (t : TermElabM α) : IO (α × ProofSnapshot) := do
144 | let ((a, termState), p') ← p.runMetaM (Lean.Elab.Term.TermElabM.run (s := p.termState)
145 | (do let r ← t; Term.synthesizeSyntheticMVarsNoPostponing; pure r))
146 | return (a, { p' with termState })
147 |
148 | /-- Run a `TacticM` monadic function in the current `ProofSnapshot`, updating the `Tactic.State`. -/
149 | def runTacticM (p : ProofSnapshot) (t : TacticM α) : IO (α × ProofSnapshot) := do
150 | let ((a, tacticState), p') ← p.runTermElabM (t p.tacticContext |>.run p.tacticState)
151 | return (a, { p' with tacticState })
152 |
153 | /--
154 | Run a `TacticM` monadic function in the current `ProofSnapshot`, updating the `Tactic.State`,
155 | and discarding the return value.
156 | -/
157 | def runTacticM' (p : ProofSnapshot) (t : TacticM α) : IO ProofSnapshot :=
158 | Prod.snd <$> p.runTacticM t
159 |
160 | /-- New traces in a `ProofSnapshot`, relative to an optional previous `ProofSnapshot`. -/
161 | def newTraces (new : ProofSnapshot) (old? : Option ProofSnapshot := none) : IO (List String) :=
162 | match old? with
163 | | none => (·.1) <$> new.runCoreM (do
164 | (← getTraces).toList.mapM fun t => do pure (← t.msg.toString).trim)
165 | | some old => do
166 | let oldCount ← (·.1) <$> old.runCoreM (return (← getTraces).size)
167 | (·.1) <$> new.runCoreM (do
168 | ((← getTraces).toList.drop oldCount).mapM fun t => do pure (← t.msg.toString).trim)
169 |
170 | /--
171 | Evaluate a `Syntax` into a `TacticM` tactic, and run it in the current `ProofSnapshot`.
172 | -/
173 | def runSyntax (p : ProofSnapshot) (t : Syntax) : IO ProofSnapshot :=
174 | Prod.snd <$> p.runTacticM (evalTactic t)
175 |
176 | /--
177 | Parse a string into a `Syntax`, evaluate it as a `TacticM` tactic,
178 | and run it in the current `ProofSnapshot`.
179 | -/
180 | def runString (p : ProofSnapshot) (t : String) : IO ProofSnapshot :=
181 | match Parser.runParserCategory p.coreState.env `tactic t with
182 | | .error e => throw (IO.userError e)
183 | | .ok stx => p.runSyntax stx
184 |
185 | /-- Pretty print the current goals in the `ProofSnapshot`. -/
186 | def ppGoals (p : ProofSnapshot) : IO (List Format) :=
187 | Prod.fst <$> p.runMetaM do p.tacticState.goals.mapM (Meta.ppGoal ·)
188 | /--
189 | Construct a `ProofSnapshot` from a `ContextInfo` and optional `LocalContext`, and a list of goals.
190 |
191 | For convenience, we also allow a list of `Expr`s, and these are appended to the goals
192 | as fresh metavariables with the given types.
193 | -/
194 | def create (ctx : ContextInfo) (lctx? : Option LocalContext) (env? : Option Environment)
195 | (goals : List MVarId) (rootGoals? : Option (List MVarId)) (types : List Expr := [])
196 | : IO ProofSnapshot := do
197 | ctx.runMetaM (lctx?.getD {}) do
198 | let goals := goals ++ (← types.mapM fun t => Expr.mvarId! <$> Meta.mkFreshExprMVar (some t))
199 | let s ← getThe Core.State
200 | let s := match env? with
201 | | none => s
202 | | some env => { s with env }
203 | pure <|
204 | { coreState := s
205 | coreContext := ← readThe Core.Context
206 | metaState := ← getThe Meta.State
207 | metaContext := ← readThe Meta.Context
208 | termState := {}
209 | termContext := {}
210 | tacticState := { goals }
211 | tacticContext := { elaborator := .anonymous }
212 | rootGoals := match rootGoals? with
213 | | none => goals
214 | | some gs => gs }
215 |
216 | open Lean.Core in
217 | /-- A copy of `Core.State` with the `Environment`, caches, and logging omitted. -/
218 | structure CompactableCoreState where
219 | -- env : Environment
220 | nextMacroScope : MacroScope := firstFrontendMacroScope + 1
221 | ngen : NameGenerator := {}
222 | -- traceState : TraceState := {}
223 | -- cache : Core.Cache := {}
224 | -- messages : MessageLog := {}
225 | -- infoState : Elab.InfoState := {}
226 |
227 | open Lean.Meta in
228 | /-- A copy of `Meta.Context` with closures omitted. -/
229 | structure CompactableMetaContext where
230 | config : Config := {}
231 | lctx : LocalContext := {}
232 | localInstances : LocalInstances := #[]
233 | defEqCtx? : Option DefEqContext := none
234 | synthPendingDepth : Nat := 0
235 | -- canUnfold? : Option (Config → ConstantInfo → CoreM Bool) := none
236 |
237 | /-- A copy of `Term.Context` with closures and a cache omitted. -/
238 | structure CompactableTermContext where
239 | declName? : Option Name := none
240 | auxDeclToFullName : FVarIdMap Name := {}
241 | macroStack : MacroStack := []
242 | mayPostpone : Bool := true
243 | errToSorry : Bool := true
244 | autoBoundImplicit : Bool := false
245 | autoBoundImplicits : PArray Expr := {}
246 | -- autoBoundImplicitForbidden : Name → Bool := fun _ => false
247 | sectionVars : NameMap Name := {}
248 | sectionFVars : NameMap Expr := {}
249 | implicitLambda : Bool := true
250 | isNoncomputableSection : Bool := false
251 | ignoreTCFailures : Bool := false
252 | inPattern : Bool := false
253 | -- tacticCache? : Option (IO.Ref Tactic.Cache) := none
254 | saveRecAppSyntax : Bool := true
255 | holesAsSyntheticOpaque : Bool := false
256 |
257 | open System (FilePath)
258 |
259 | /--
260 | Pickle a `ProofSnapshot`, discarding closures and non-essential caches.
261 |
262 | When pickling the `Environment`, we do so relative to its imports.
263 | -/
264 | def pickle (p : ProofSnapshot) (path : FilePath) : IO Unit := do
265 | let env := p.coreState.env
266 | let p' := { p with coreState := { p.coreState with env := ← mkEmptyEnvironment }}
267 | let (cfg, _) ← Lean.Meta.getConfig.toIO p'.coreContext p'.coreState p'.metaContext p'.metaState
268 | _root_.pickle path
269 | (env.header.imports,
270 | env.constants.map₂,
271 | ({ p'.coreState with } : CompactableCoreState),
272 | p'.coreContext,
273 | p'.metaState,
274 | ({ p'.metaContext with config := cfg } : CompactableMetaContext),
275 | p'.termState,
276 | ({ p'.termContext with } : CompactableTermContext),
277 | p'.tacticState,
278 | p'.tacticContext,
279 | p'.rootGoals)
280 |
281 | /--
282 | Unpickle a `ProofSnapshot`.
283 | -/
284 | def unpickle (path : FilePath) (cmd? : Option CommandSnapshot) :
285 | IO (ProofSnapshot × CompactedRegion) := unsafe do
286 | let ((imports, map₂, coreState, coreContext, metaState, metaContext, termState, termContext,
287 | tacticState, tacticContext, rootGoals), region) ←
288 | _root_.unpickle (Array Import × PHashMap Name ConstantInfo × CompactableCoreState ×
289 | Core.Context × Meta.State × CompactableMetaContext × Term.State × CompactableTermContext ×
290 | Tactic.State × Tactic.Context × List MVarId) path
291 | let env ← match cmd? with
292 | | none =>
293 | enableInitializersExecution
294 | (← importModules imports {} 0 (loadExts := true)).replay (Std.HashMap.ofList map₂.toList)
295 | | some cmd =>
296 | cmd.cmdState.env.replay (Std.HashMap.ofList map₂.toList)
297 | let p' : ProofSnapshot :=
298 | { coreState := { coreState with env }
299 | coreContext
300 | metaState
301 | metaContext := { metaContext with }
302 | termState
303 | termContext := { termContext with }
304 | tacticState
305 | tacticContext
306 | rootGoals }
307 | let (_, p'') ← p'.runCoreM do
308 | for o in ← getOpenDecls do
309 | if let .simple ns _ := o then
310 | activateScoped ns
311 | return (p'', region)
312 |
313 | end ProofSnapshot
314 |
--------------------------------------------------------------------------------
/REPL/Util/Path.lean:
--------------------------------------------------------------------------------
1 | /-
2 | Copyright (c) 2022 Gabriel Ebner. All rights reserved.
3 | Released under Apache 2.0 license as described in the file LICENSE.
4 | Authors: Gabriel Ebner
5 | -/
6 | import Lean
7 |
8 | -- This has been duplicated from Std4 to avoid a dependency.
9 |
10 | /-!
11 | # `compile_time_search_path%` term elaborator.
12 |
13 | Use this as `searchPathRef.set compile_time_search_path%`.
14 | -/
15 |
16 | open Lean System
17 |
18 | -- Ideally this instance would be constructed simply by `deriving instance ToExpr for FilePath`
19 | -- but for now we have decided not to upstream the `ToExpr` derive handler from `Mathlib`.
20 | -- https://leanprover.zulipchat.com/#narrow/stream/348111-std4/topic/ToExpr.20derive.20handler/near/386476438
21 | instance : ToExpr FilePath where
22 | toTypeExpr := mkConst ``FilePath
23 | toExpr path := mkApp (mkConst ``FilePath.mk) (toExpr path.1)
24 |
25 | /--
26 | Term elaborator that retrieves the current `SearchPath`.
27 |
28 | Typical usage is `searchPathRef.set compile_time_search_path%`.
29 |
30 | This must not be used in files that are potentially compiled on another machine and then
31 | imported.
32 | (That is, if used in an imported file it will embed the search path from whichever machine
33 | compiled the `.olean`.)
34 | -/
35 | elab "compile_time_search_path%" : term =>
36 | return toExpr (← searchPathRef.get)
37 |
--------------------------------------------------------------------------------
/REPL/Util/Pickle.lean:
--------------------------------------------------------------------------------
1 | /-
2 | Copyright (c) 2023 Mario Carneiro. All rights reserved.
3 | Released under Apache 2.0 license as described in the file LICENSE.
4 | Authors: Mario Carneiro
5 | -/
6 | import Lean.Environment
7 |
8 | /-!
9 | # Pickling and unpickling objects
10 |
11 | By abusing `saveModuleData` and `readModuleData` we can pickle and unpickle objects to disk.
12 | -/
13 |
14 | open Lean System
15 |
16 | /--
17 | Save an object to disk.
18 | If you need to write multiple objects from within a single declaration,
19 | you will need to provide a unique `key` for each.
20 | -/
21 | def pickle {α : Type} (path : FilePath) (x : α) (key : Name := by exact decl_name%) : IO Unit :=
22 | saveModuleData path key (unsafe unsafeCast x)
23 |
24 | /--
25 | Load an object from disk.
26 | Note: The returned `CompactedRegion` can be used to free the memory behind the value
27 | of type `α`, using `CompactedRegion.free` (which is only safe once all references to the `α` are
28 | released). Ignoring the `CompactedRegion` results in the data being leaked.
29 | Use `withUnpickle` to call `CompactedRegion.free` automatically.
30 |
31 | This function is unsafe because the data being loaded may not actually have type `α`, and this
32 | may cause crashes or other bad behavior.
33 | -/
34 | unsafe def unpickle (α : Type) (path : FilePath) : IO (α × CompactedRegion) := do
35 | let (x, region) ← readModuleData path
36 | pure (unsafeCast x, region)
37 |
38 | /-- Load an object from disk and run some continuation on it, freeing memory afterwards. -/
39 | unsafe def withUnpickle [Monad m] [MonadLiftT IO m] {α β : Type}
40 | (path : FilePath) (f : α → m β) : m β := do
41 | let (x, region) ← unpickle α path
42 | let r ← f x
43 | region.free
44 | pure r
45 |
--------------------------------------------------------------------------------
/lake-manifest.json:
--------------------------------------------------------------------------------
1 | {"version": "1.1.0",
2 | "packagesDir": ".lake/packages",
3 | "packages": [],
4 | "name": "REPL",
5 | "lakeDir": ".lake"}
6 |
--------------------------------------------------------------------------------
/lakefile.toml:
--------------------------------------------------------------------------------
1 | name = "REPL"
2 | defaultTargets = ["repl"]
3 |
4 | [[lean_lib]]
5 | name = "REPL"
6 |
7 | [[lean_exe]]
8 | name = "repl"
9 | root = "REPL.Main"
10 | supportInterpreter = true
11 |
--------------------------------------------------------------------------------
/lean-toolchain:
--------------------------------------------------------------------------------
1 | leanprover/lean4:v4.21.0-rc3
2 |
--------------------------------------------------------------------------------
/test.sh:
--------------------------------------------------------------------------------
1 | #!/bin/bash
2 |
3 | # Define the paths
4 | IN_DIR="test"
5 | EXPECTED_DIR="test"
6 |
7 | lake build
8 |
9 | # ignore locale to ensure test `bla` runs before `bla2`
10 | export LC_COLLATE=C
11 |
12 | # Iterate over each .in file in the test directory
13 | for infile in $IN_DIR/*.in; do
14 | # Extract the base filename without the extension
15 | base=$(basename "$infile" .in)
16 |
17 | # Define the path for the expected output file
18 | expectedfile="$EXPECTED_DIR/$base.expected.out"
19 |
20 | # Check if the expected output file exists
21 | if [[ ! -f $expectedfile ]]; then
22 | echo "Expected output file $expectedfile does not exist. Skipping $infile."
23 | continue
24 | fi
25 |
26 | # Run the command and store its output in a temporary file
27 | tmpfile=$(mktemp)
28 | .lake/build/bin/repl < "$infile" > "$tmpfile" 2>&1
29 |
30 | # Compare the output with the expected output
31 | if diff "$tmpfile" "$expectedfile"; then
32 | echo "$base: PASSED"
33 | # Remove the temporary file
34 | rm "$tmpfile"
35 | else
36 | echo "$base: FAILED"
37 | # Rename the temporary file instead of removing it
38 | mv "$tmpfile" "${expectedfile/.expected.out/.produced.out}"
39 | exit 1
40 | fi
41 |
42 | done
43 |
44 | # Run the Mathlib tests
45 | cp lean-toolchain test/Mathlib/
46 | cd test/Mathlib/ && ./test.sh
47 |
--------------------------------------------------------------------------------
/test/Mathlib/.gitignore:
--------------------------------------------------------------------------------
1 | /build
2 | /lake-packages/*
3 | /lakefile.olean
4 | /test/*.olean
5 | /test/*.olean.tmp
6 |
--------------------------------------------------------------------------------
/test/Mathlib/H20231110.sh:
--------------------------------------------------------------------------------
1 | (head -n 2 test/H20231110.in; for i in {1..100}; do tail -n +3 test/H20231110.in; done) |
2 | lake env ../../build/bin/repl
3 |
--------------------------------------------------------------------------------
/test/Mathlib/ReplMathlibTests.lean:
--------------------------------------------------------------------------------
1 | import Mathlib
2 |
--------------------------------------------------------------------------------
/test/Mathlib/lake-manifest.json:
--------------------------------------------------------------------------------
1 | {"version": "1.1.0",
2 | "packagesDir": ".lake/packages",
3 | "packages":
4 | [{"url": "https://github.com/leanprover-community/mathlib4",
5 | "type": "git",
6 | "subDir": null,
7 | "scope": "",
8 | "rev": "6455ba8ab6d25ef9f661dc663a524375d3984164",
9 | "name": "mathlib",
10 | "manifestFile": "lake-manifest.json",
11 | "inputRev": "v4.21.0-rc3",
12 | "inherited": false,
13 | "configFile": "lakefile.lean"},
14 | {"url": "https://github.com/leanprover-community/plausible",
15 | "type": "git",
16 | "subDir": null,
17 | "scope": "leanprover-community",
18 | "rev": "1603151ac0db4e822908e18094f16acc250acaff",
19 | "name": "plausible",
20 | "manifestFile": "lake-manifest.json",
21 | "inputRev": "main",
22 | "inherited": true,
23 | "configFile": "lakefile.toml"},
24 | {"url": "https://github.com/leanprover-community/LeanSearchClient",
25 | "type": "git",
26 | "subDir": null,
27 | "scope": "leanprover-community",
28 | "rev": "6c62474116f525d2814f0157bb468bf3a4f9f120",
29 | "name": "LeanSearchClient",
30 | "manifestFile": "lake-manifest.json",
31 | "inputRev": "main",
32 | "inherited": true,
33 | "configFile": "lakefile.toml"},
34 | {"url": "https://github.com/leanprover-community/import-graph",
35 | "type": "git",
36 | "subDir": null,
37 | "scope": "leanprover-community",
38 | "rev": "e25fe66cf13e902ba550533ef681cc35a9f18dc2",
39 | "name": "importGraph",
40 | "manifestFile": "lake-manifest.json",
41 | "inputRev": "main",
42 | "inherited": true,
43 | "configFile": "lakefile.toml"},
44 | {"url": "https://github.com/leanprover-community/ProofWidgets4",
45 | "type": "git",
46 | "subDir": null,
47 | "scope": "leanprover-community",
48 | "rev": "6980f6ca164de593cb77cd03d8eac549cc444156",
49 | "name": "proofwidgets",
50 | "manifestFile": "lake-manifest.json",
51 | "inputRev": "v0.0.62",
52 | "inherited": true,
53 | "configFile": "lakefile.lean"},
54 | {"url": "https://github.com/leanprover-community/aesop",
55 | "type": "git",
56 | "subDir": null,
57 | "scope": "leanprover-community",
58 | "rev": "f0424862c97fec5bae253f4f1e0ff001f78187c0",
59 | "name": "aesop",
60 | "manifestFile": "lake-manifest.json",
61 | "inputRev": "master",
62 | "inherited": true,
63 | "configFile": "lakefile.toml"},
64 | {"url": "https://github.com/leanprover-community/quote4",
65 | "type": "git",
66 | "subDir": null,
67 | "scope": "leanprover-community",
68 | "rev": "e1d2994e0acdee2f0c03c9d84d28a5df34aa0020",
69 | "name": "Qq",
70 | "manifestFile": "lake-manifest.json",
71 | "inputRev": "master",
72 | "inherited": true,
73 | "configFile": "lakefile.toml"},
74 | {"url": "https://github.com/leanprover-community/batteries",
75 | "type": "git",
76 | "subDir": null,
77 | "scope": "leanprover-community",
78 | "rev": "08681ddeb7536a50dea8026c6693cb9b07f01717",
79 | "name": "batteries",
80 | "manifestFile": "lake-manifest.json",
81 | "inputRev": "main",
82 | "inherited": true,
83 | "configFile": "lakefile.toml"},
84 | {"url": "https://github.com/leanprover/lean4-cli",
85 | "type": "git",
86 | "subDir": null,
87 | "scope": "leanprover",
88 | "rev": "a0abd472348dd725adbb26732e79b26e7e220913",
89 | "name": "Cli",
90 | "manifestFile": "lake-manifest.json",
91 | "inputRev": "main",
92 | "inherited": true,
93 | "configFile": "lakefile.toml"}],
94 | "name": "«repl-mathlib-tests»",
95 | "lakeDir": ".lake"}
96 |
--------------------------------------------------------------------------------
/test/Mathlib/lakefile.toml:
--------------------------------------------------------------------------------
1 | name = "«repl-mathlib-tests»"
2 | defaultTargets = ["ReplMathlibTests"]
3 |
4 | [[require]]
5 | name = "mathlib"
6 | git = "https://github.com/leanprover-community/mathlib4"
7 | rev = "v4.21.0-rc3"
8 |
9 | [[lean_lib]]
10 | name = "ReplMathlibTests"
11 | globs = ["test.+"]
12 |
--------------------------------------------------------------------------------
/test/Mathlib/lean-toolchain:
--------------------------------------------------------------------------------
1 | leanprover/lean4:v4.21.0-rc3
2 |
--------------------------------------------------------------------------------
/test/Mathlib/test.sh:
--------------------------------------------------------------------------------
1 | #!/bin/bash
2 |
3 | # Define the paths
4 | IN_DIR="test"
5 | EXPECTED_DIR="test"
6 |
7 | lake exe cache get > /dev/null
8 | lake build Mathlib
9 |
10 | # Iterate over each .in file in the test directory
11 | for infile in $IN_DIR/*.in; do
12 | # Extract the base filename without the extension
13 | base=$(basename "$infile" .in)
14 |
15 | # Define the path for the expected output file
16 | expectedfile="$EXPECTED_DIR/$base.expected.out"
17 |
18 | # Check if the expected output file exists
19 | if [[ ! -f $expectedfile ]]; then
20 | echo "Expected output file $expectedfile does not exist. Skipping $infile."
21 | continue
22 | fi
23 |
24 | # Run the command and store its output in a temporary file
25 | tmpfile=$(mktemp)
26 | lake env ../../.lake/build/bin/repl < "$infile" > "$tmpfile" 2>&1
27 |
28 | # Compare the output with the expected output
29 | if diff "$tmpfile" "$expectedfile"; then
30 | echo "$base: PASSED"
31 | # Remove the temporary file
32 | rm "$tmpfile"
33 | else
34 | echo "$base: FAILED"
35 | # Rename the temporary file instead of removing it
36 | mv "$tmpfile" "${expectedfile/.expected.out/.produced.out}"
37 | exit 1
38 | fi
39 |
40 | done
41 |
42 |
--------------------------------------------------------------------------------
/test/Mathlib/test/20240209.expected.out:
--------------------------------------------------------------------------------
1 | {"sorries":
2 | [{"proofState": 0,
3 | "pos": {"line": 1, "column": 22},
4 | "goal": "⊢ False",
5 | "endPos": {"line": 1, "column": 27}}],
6 | "messages":
7 | [{"severity": "warning",
8 | "pos": {"line": 1, "column": 0},
9 | "endPos": {"line": 1, "column": 7},
10 | "data": "declaration uses 'sorry'"}],
11 | "env": 0}
12 |
13 | {"proofStatus": "Incomplete: contains sorry",
14 | "proofState": 1,
15 | "messages":
16 | [{"severity": "error",
17 | "pos": {"line": 0, "column": 0},
18 | "endPos": {"line": 0, "column": 0},
19 | "data": "unsolved goals\n⊢ False"}],
20 | "goals": []}
21 |
22 |
--------------------------------------------------------------------------------
/test/Mathlib/test/20240209.in:
--------------------------------------------------------------------------------
1 | {"cmd": "example : False := by sorry"}
2 |
3 | {"tactic": "simpa using show False by done", "proofState": 0}
4 |
--------------------------------------------------------------------------------
/test/Mathlib/test/20240209.lean:
--------------------------------------------------------------------------------
1 | example : False := by
2 | simpa using show False by done
3 |
--------------------------------------------------------------------------------
/test/Mathlib/test/H20231020.expected.out:
--------------------------------------------------------------------------------
1 | {"env": 0}
2 |
3 | {"env": 1}
4 |
5 | {"env": 2}
6 |
7 | {"env": 3}
8 |
9 |
--------------------------------------------------------------------------------
/test/Mathlib/test/H20231020.in:
--------------------------------------------------------------------------------
1 | {"cmd": "import Mathlib.Algebra.BigOperators.Group.Finset.Basic\nimport Mathlib.Data.Real.Basic\nimport Mathlib.Data.Complex.Basic\nimport Mathlib.Data.Nat.Log\nimport Mathlib.Data.Complex.Exponential\nimport Mathlib.NumberTheory.Divisors\nimport Mathlib.Data.ZMod.Defs\nimport Mathlib.Data.ZMod.Basic\nimport Mathlib.Topology.Basic\nimport Mathlib.Data.Nat.Digits\nimport Mathlib.Tactic.NormNum.GCD\nopen BigOperators\nopen Real\nopen Nat\nopen Topology"}
2 |
3 | {"cmd": "theorem mathd_numbertheory_188 : Nat.gcd 180 168 = 12 := by norm_num", "env": 0}
4 |
5 | {"cmd": "theorem mathd_numbertheory_403 : ∑ k ∈ (Nat.properDivisors 198), k = 270 := by simp (config := {decide := true})", "env": 0}
6 |
7 | {"cmd": "theorem mathd_numbertheory_109 (v : ℕ → ℕ) (h₀ : ∀ n, v n = 2 * n - 1) : (∑ k ∈ Finset.Icc 1 100, v k) % 7 = 4 := by simp_rw (config := {decide := true}) [h₀]", "env": 0}
8 |
9 |
--------------------------------------------------------------------------------
/test/Mathlib/test/H20231020.lean:
--------------------------------------------------------------------------------
1 | import Mathlib.Algebra.BigOperators.Group.Finset.Basic
2 | import Mathlib.Data.Real.Basic
3 | import Mathlib.Data.Complex.Basic
4 | import Mathlib.Data.Nat.Log
5 | import Mathlib.Data.Complex.Exponential
6 | import Mathlib.NumberTheory.Divisors
7 | import Mathlib.Data.ZMod.Defs
8 | import Mathlib.Data.ZMod.Basic
9 | import Mathlib.Topology.Basic
10 | import Mathlib.Data.Nat.Digits
11 | import Mathlib.Tactic.NormNum.GCD
12 |
13 | open BigOperators
14 | open Real
15 | open Nat
16 | open Topology
17 |
18 | theorem mathd_numbertheory_188 : Nat.gcd 180 168 = 12 := by norm_num
19 |
20 | theorem mathd_numbertheory_403 : ∑ k ∈ (Nat.properDivisors 198), k = 270 := by simp (config := {decide := true})
21 |
22 | theorem mathd_numbertheory_109 (v : ℕ → ℕ) (h₀ : ∀ n, v n = 2 * n - 1) : (∑ k ∈ Finset.Icc 1 100, v k) % 7 = 4 := by simp_rw (config := {decide := true}) [h₀]
23 |
--------------------------------------------------------------------------------
/test/Mathlib/test/H20231110.expected.out:
--------------------------------------------------------------------------------
1 | {"env": 0}
2 |
3 | {"env": 1}
4 |
5 |
--------------------------------------------------------------------------------
/test/Mathlib/test/H20231110.in:
--------------------------------------------------------------------------------
1 | {"cmd": "import Mathlib\nopen Real\nlocal macro_rules | `($x ^ $y) => `(HPow.hPow $x $y)"}
2 |
3 | {"cmd": "example {n} {x a b : ℝ} (m : ℕ) (e₁ : n + 1 = m) (rm : ℝ) (er : ↑m = rm) (h : |x| ≤ 1) (e : |1 - a| ≤ b - |x| / rm * ((rm + 1) / rm)) : |exp x - expNear n x a| ≤ |x| ^ n / n.factorial * b := by apply Real.exp_approx_end' m e₁ rm er h e", "env": 0}
4 |
5 |
--------------------------------------------------------------------------------
/test/Mathlib/test/H20231115.expected.out:
--------------------------------------------------------------------------------
1 | {"env": 0}
2 |
3 | {"sorries":
4 | [{"proofState": 0,
5 | "pos": {"line": 1, "column": 36},
6 | "goal": "x : Nat\n⊢ x + 1 > x",
7 | "endPos": {"line": 1, "column": 41}}],
8 | "messages":
9 | [{"severity": "warning",
10 | "pos": {"line": 1, "column": 0},
11 | "endPos": {"line": 1, "column": 7},
12 | "data": "declaration uses 'sorry'"}],
13 | "env": 1}
14 |
15 | {"proofStatus": "Incomplete: open goals remain",
16 | "proofState": 1,
17 | "goals":
18 | ["case zero\n⊢ 0 + 1 > 0",
19 | "case succ\nx : Nat\nhx : x + 1 > x\n⊢ x + 1 + 1 > x + 1"]}
20 |
21 |
--------------------------------------------------------------------------------
/test/Mathlib/test/H20231115.in:
--------------------------------------------------------------------------------
1 | {"cmd": "import Mathlib.Tactic.Cases"}
2 |
3 | {"cmd": "example {x : Nat} : x + 1 > x := by sorry", "env": 0}
4 |
5 | {"tactic": "induction' x with x hx", "proofState": 0}
6 |
--------------------------------------------------------------------------------
/test/Mathlib/test/H20231115_2.expected.out:
--------------------------------------------------------------------------------
1 | {"env": 0}
2 |
3 | {"sorries":
4 | [{"proofState": 0,
5 | "pos": {"line": 3, "column": 12},
6 | "goal": "case zero\n⊢ 0 + 1 > 0",
7 | "endPos": {"line": 3, "column": 17}},
8 | {"proofState": 1,
9 | "pos": {"line": 3, "column": 12},
10 | "goal": "case succ\nx : Nat\nhx : x + 1 > x\n⊢ x + 1 + 1 > x + 1",
11 | "endPos": {"line": 3, "column": 17}}],
12 | "messages":
13 | [{"severity": "warning",
14 | "pos": {"line": 1, "column": 0},
15 | "endPos": {"line": 1, "column": 7},
16 | "data": "declaration uses 'sorry'"}],
17 | "env": 1}
18 |
19 |
--------------------------------------------------------------------------------
/test/Mathlib/test/H20231115_2.in:
--------------------------------------------------------------------------------
1 | {"cmd": "import Mathlib.Tactic.Cases"}
2 |
3 | {"cmd": "example {x : Nat} : x + 1 > x := by\n induction' x with x hx\n all_goals sorry", "env": 0}
4 |
5 |
--------------------------------------------------------------------------------
/test/Mathlib/test/H20231115_3.expected.out:
--------------------------------------------------------------------------------
1 | {"env": 0}
2 |
3 | {"messages":
4 | [{"severity": "error",
5 | "pos": {"line": 1, "column": 33},
6 | "endPos": {"line": 2, "column": 24},
7 | "data":
8 | "unsolved goals\ncase zero\n⊢ 0 + 1 > 0\n\ncase succ\nx : Nat\nhx : x + 1 > x\n⊢ x + 1 + 1 > x + 1"}],
9 | "env": 1}
10 |
11 |
--------------------------------------------------------------------------------
/test/Mathlib/test/H20231115_3.in:
--------------------------------------------------------------------------------
1 | {"cmd": "import Mathlib.Tactic.Cases"}
2 |
3 | {"cmd": "example {x : Nat} : x + 1 > x := by\n induction' x with x hx", "env": 0}
4 |
5 |
--------------------------------------------------------------------------------
/test/Mathlib/test/H20231214.in:
--------------------------------------------------------------------------------
1 | {"cmd": "import Mathlib\nopen Real\nopen Nat\nopen BigOperators"}
2 |
3 | {"cmd" : "theorem mathd_algebra_35\n (p q : ℝ → ℝ)\n (h₀ : ∀ x, p x = 2 - x^2)\n (h₁ : ∀ x, x ≠ 0 -> q x = 6 / x) :\n p (q 2) = -7 := by sorry", "env": 0}
4 |
5 | {"tactic": "have hQ : ∀ x, p x = 6 / x", "proofState": 0}
6 |
7 | {"tactic": " intro x", "proofState": 1}
8 |
9 | {"tactic": " calc p x = 6 / x * p x := h₀ (x)\n _ = (6/2) * 6 / x * (6 / x) := sub_sub_sub_cancel sq_ne_zero", "proofState": 2}
10 |
--------------------------------------------------------------------------------
/test/Mathlib/test/H20231214.lean:
--------------------------------------------------------------------------------
1 | import Mathlib
2 | open Real
3 | open Nat
4 | open BigOperators
5 |
6 | /--
7 | error: type mismatch
8 | h₀ x
9 | has type
10 | p x = 2 - x ^ 2 : Prop
11 | but is expected to have type
12 | p x = 6 / x * p x : Prop
13 | ---
14 | error: unsolved goals
15 | case calc.step
16 | p q : ℝ → ℝ
17 | h₀ : ∀ (x : ℝ), p x = 2 - x ^ 2
18 | h₁ : ∀ (x : ℝ), x ≠ 0 → q x = 6 / x
19 | x : ℝ
20 | ⊢ 6 / 2 * 6 / x * (6 / x) = 6 / x
21 | ---
22 | error: unsolved goals
23 | p q : ℝ → ℝ
24 | h₀ : ∀ (x : ℝ), p x = 2 - x ^ 2
25 | h₁ : ∀ (x : ℝ), x ≠ 0 → q x = 6 / x
26 | hQ : ∀ (x : ℝ), p x = 6 / x
27 | ⊢ p (q 2) = -7
28 | -/
29 | #guard_msgs in
30 | theorem mathd_algebra_35
31 | (p q : ℝ → ℝ)
32 | (h₀ : ∀ x, p x = 2 - x^2)
33 | (h₁ : ∀ x, x ≠ 0 -> q x = 6 / x) :
34 | p (q 2) = -7 := by
35 | have hQ : ∀ x, p x = 6 / x := by
36 | intro x
37 | calc p x = 6 / x * p x := h₀ (x)
38 | _ = (6/2) * 6 / x * (6 / x) := sorry
39 |
--------------------------------------------------------------------------------
/test/Mathlib/test/H20231215.expected.out:
--------------------------------------------------------------------------------
1 | {"env": 0}
2 |
3 | {"env": 0}
4 |
5 |
--------------------------------------------------------------------------------
/test/Mathlib/test/H20231215.in:
--------------------------------------------------------------------------------
1 | {"cmd": "import Mathlib\nopen BigOperators\nopen Real\nopen Nat"}
2 |
3 | {"pickleTo": "test/H20231215.olean", "env": 0}
4 |
5 |
--------------------------------------------------------------------------------
/test/Mathlib/test/H20231215_2.expected.out:
--------------------------------------------------------------------------------
1 | {"env": 0}
2 |
3 | {"sorries":
4 | [{"proofState": 0,
5 | "pos": {"line": 2, "column": 20},
6 | "goal": "⊢ 9! % 10 = 0",
7 | "endPos": {"line": 2, "column": 25}}],
8 | "messages":
9 | [{"severity": "warning",
10 | "pos": {"line": 1, "column": 8},
11 | "endPos": {"line": 1, "column": 30},
12 | "data": "declaration uses 'sorry'"}],
13 | "env": 1}
14 |
15 |
--------------------------------------------------------------------------------
/test/Mathlib/test/H20231215_2.in:
--------------------------------------------------------------------------------
1 | {"unpickleEnvFrom": "test/H20231215.olean"}
2 |
3 | {"cmd": "theorem mathd_numbertheory_739 :\n 9! % 10 = 0 := by sorry", "env": 0}
4 |
--------------------------------------------------------------------------------
/test/Mathlib/test/exact.expected.out:
--------------------------------------------------------------------------------
1 | {"env": 0}
2 |
3 | {"sorries":
4 | [{"proofState": 0,
5 | "pos": {"line": 1, "column": 27},
6 | "goal": "⊢ 0 < 1",
7 | "endPos": {"line": 1, "column": 32}}],
8 | "messages":
9 | [{"severity": "warning",
10 | "pos": {"line": 1, "column": 8},
11 | "endPos": {"line": 1, "column": 12},
12 | "data": "declaration uses 'sorry'"}],
13 | "env": 1}
14 |
15 | {"proofStatus": "Completed",
16 | "proofState": 1,
17 | "messages":
18 | [{"severity": "info",
19 | "pos": {"line": 0, "column": 0},
20 | "endPos": {"line": 0, "column": 0},
21 | "data": "Try this: exact Nat.one_pos"}],
22 | "goals": []}
23 |
24 | {"sorries":
25 | [{"proofState": 2,
26 | "pos": {"line": 1, "column": 27},
27 | "goal": "⊢ 3 = 7",
28 | "endPos": {"line": 1, "column": 32}}],
29 | "messages":
30 | [{"severity": "warning",
31 | "pos": {"line": 1, "column": 8},
32 | "endPos": {"line": 1, "column": 12},
33 | "data": "declaration uses 'sorry'"}],
34 | "env": 2}
35 |
36 | {"message":
37 | "Lean error:\n`exact?` could not close the goal. Try `apply?` to see partial suggestions."}
38 |
39 |
--------------------------------------------------------------------------------
/test/Mathlib/test/exact.in:
--------------------------------------------------------------------------------
1 | {"cmd": "import Mathlib"}
2 |
3 | {"cmd": "theorem test : 0 < 1 := by sorry", "env": 0}
4 |
5 | {"tactic": "exact?", "proofState": 0}
6 |
7 | {"cmd": "theorem test : 3 = 7 := by sorry", "env": 0}
8 |
9 | {"tactic": "exact?", "proofState": 2}
10 |
11 |
--------------------------------------------------------------------------------
/test/Mathlib/test/import_Mathlib.lean:
--------------------------------------------------------------------------------
1 | import Mathlib
2 |
--------------------------------------------------------------------------------
/test/Mathlib/test/induction.expected.out:
--------------------------------------------------------------------------------
1 | {"env": 0}
2 |
3 | {"sorries":
4 | [{"proofState": 0,
5 | "pos": {"line": 1, "column": 36},
6 | "goal": "x : ℕ\n⊢ x = x",
7 | "endPos": {"line": 1, "column": 41}}],
8 | "messages":
9 | [{"severity": "warning",
10 | "pos": {"line": 1, "column": 8},
11 | "endPos": {"line": 1, "column": 11},
12 | "data": "declaration uses 'sorry'"}],
13 | "env": 1}
14 |
15 | {"proofStatus": "Incomplete: open goals remain",
16 | "proofState": 1,
17 | "goals":
18 | ["case zero\n⊢ 0 = 0", "case succ\nn✝ : ℕ\na✝ : n✝ = n✝\n⊢ n✝ + 1 = n✝ + 1"]}
19 |
20 | {"proofStatus": "Incomplete: open goals remain",
21 | "proofState": 2,
22 | "goals": ["case succ\nn✝ : ℕ\na✝ : n✝ = n✝\n⊢ n✝ + 1 = n✝ + 1"]}
23 |
24 | {"sorries":
25 | [{"proofState": 3, "goal": "case zero\n⊢ 0 = 0"},
26 | {"proofState": 4, "goal": "case succ\nx : ℕ\na✝ : x = x\n⊢ x + 1 = x + 1"}],
27 | "proofStatus": "Incomplete: contains sorry",
28 | "proofState": 5,
29 | "goals": []}
30 |
31 |
--------------------------------------------------------------------------------
/test/Mathlib/test/induction.in:
--------------------------------------------------------------------------------
1 | {"cmd": "import Mathlib"}
2 |
3 | {"env" : 0, "cmd": "theorem foo (x : Nat) : x = x := by sorry"}
4 |
5 | {"proofState" : 0, "tactic": "induction x"}
6 |
7 | {"proofState" : 1, "tactic": "next => rfl"}
8 |
9 | {"proofState" : 0, "tactic": "induction x with\n| zero => sorry\n| succ x => sorry"}
10 |
11 |
--------------------------------------------------------------------------------
/test/Mathlib/test/induction.lean:
--------------------------------------------------------------------------------
1 | import Mathlib
2 |
3 | theorem bar (x : Nat) : x = x := by
4 | induction x with
5 | | zero => sorry
6 | | succ n ih => sorry
7 |
--------------------------------------------------------------------------------
/test/Mathlib/test/on_goal.expected.out:
--------------------------------------------------------------------------------
1 | {"env": 0}
2 |
3 | {"sorries":
4 | [{"proofState": 0,
5 | "pos": {"line": 1, "column": 60},
6 | "goal": "x : ℝ\nh0 : |x| > 1\n⊢ x < 0 ∨ 2 * x > 2",
7 | "endPos": {"line": 1, "column": 65}}],
8 | "messages":
9 | [{"severity": "warning",
10 | "pos": {"line": 1, "column": 0},
11 | "endPos": {"line": 1, "column": 7},
12 | "data": "declaration uses 'sorry'"}],
13 | "env": 1}
14 |
15 | {"sorries":
16 | [{"proofState": 1, "goal": "x : ℝ\nh0 : |x| > 1\n⊢ x = x"},
17 | {"proofState": 2, "goal": "x : ℝ\nh0 : |x| > 1\nh1 : x = x\n⊢ x = x"}],
18 | "proofStatus": "Incomplete: open goals remain",
19 | "proofState": 3,
20 | "goals":
21 | ["case pos\nx : ℝ\nh0 : |x| > 1\nh1 h2 : x = x\nh3 : x < 0\n⊢ x < 0 ∨ 2 * x > 2",
22 | "case neg\nx : ℝ\nh0 : |x| > 1\nh1 h2 : x = x\nh3 : ¬x < 0\n⊢ x < 0 ∨ 2 * x > 2"]}
23 |
24 |
--------------------------------------------------------------------------------
/test/Mathlib/test/on_goal.in:
--------------------------------------------------------------------------------
1 | {"cmd": "import Mathlib\nopen Real"}
2 |
3 | {"cmd": "example {x : ℝ} (h0: |x| > 1) : (x < 0) ∨ (2 * x > 2) := by sorry", "env": 0}
4 |
5 | {"tactic": "on_goal 1 =>\n have h1 : x = x := by sorry\n have h2 : x = x := by sorry\n by_cases h3 : x < 0", "proofState": 0}
6 |
--------------------------------------------------------------------------------
/test/Mathlib/test/pickle.expected.out:
--------------------------------------------------------------------------------
1 | {"env": 0}
2 |
3 | {"sorries":
4 | [{"proofState": 0,
5 | "pos": {"line": 4, "column": 14},
6 | "goal": "x : ℕ\nh : 2 * (2 * (2 * (2 * x))) = 48\n⊢ x = 3",
7 | "endPos": {"line": 4, "column": 19}}],
8 | "messages":
9 | [{"severity": "warning",
10 | "pos": {"line": 1, "column": 8},
11 | "endPos": {"line": 1, "column": 25},
12 | "data": "declaration uses 'sorry'"}],
13 | "env": 1}
14 |
15 | {"proofStatus": "Incomplete: open goals remain",
16 | "proofState": 1,
17 | "goals": ["x : ℕ\nh : 2 * (2 * (2 * (2 * x))) = 48\n⊢ x = 3"]}
18 |
19 |
--------------------------------------------------------------------------------
/test/Mathlib/test/pickle.in:
--------------------------------------------------------------------------------
1 | {"cmd": "import Mathlib\nopen BigOperators\nopen Real\nopen Nat"}
2 |
3 | {"cmd": "theorem mathd_algebra_455\n (x : Nat)\n (h : 2 * (2 * (2 * (2 * x))) = 48) :\n x = 3 := by sorry", "env": 0}
4 |
5 | {"pickleTo": "test/pickle.olean", "proofState": 0}
6 |
7 |
--------------------------------------------------------------------------------
/test/Mathlib/test/pickle_2.expected.out:
--------------------------------------------------------------------------------
1 | {"proofStatus": "Incomplete: open goals remain",
2 | "proofState": 0,
3 | "goals": ["x : ℕ\nh : 2 * (2 * (2 * (2 * x))) = 48\n⊢ x = 3"]}
4 |
5 | {"proofStatus": "Incomplete: open goals remain",
6 | "proofState": 1,
7 | "goals": ["x : ℕ\nh : 2 * (2 * (2 * (2 * x))) = 48\n⊢ x = 3"]}
8 |
9 |
--------------------------------------------------------------------------------
/test/Mathlib/test/pickle_2.in:
--------------------------------------------------------------------------------
1 | {"unpickleProofStateFrom": "test/pickle.olean"}
2 |
3 | {"tactic": "norm_num at h", "proofState": 0}
4 |
5 |
--------------------------------------------------------------------------------
/test/Mathlib/test/placeholder_synthesis.expected.out:
--------------------------------------------------------------------------------
1 | {"env": 0}
2 |
3 | {"messages":
4 | [{"severity": "error",
5 | "pos": {"line": 3, "column": 19},
6 | "endPos": {"line": 3, "column": 20},
7 | "data":
8 | "don't know how to synthesize placeholder\ncontext:\nn : ℕ\nh : n ≠ 2\n⊢ n = 2"}],
9 | "env": 1}
10 |
11 | {"sorries":
12 | [{"proofState": 0,
13 | "pos": {"line": 1, "column": 19},
14 | "goal": "⊢ 1 = 0",
15 | "endPos": {"line": 1, "column": 24}}],
16 | "messages":
17 | [{"severity": "warning",
18 | "pos": {"line": 1, "column": 0},
19 | "endPos": {"line": 1, "column": 7},
20 | "data": "declaration uses 'sorry'"}],
21 | "env": 2}
22 |
23 | {"proofStatus": "Incomplete: open goals remain",
24 | "proofState": 1,
25 | "goals": ["case intro\nn : ℕ\nh : n ≠ 2\n⊢ 1 = 0"]}
26 |
27 | {"proofStatus": "Incomplete: contains metavariable(s)",
28 | "proofState": 2,
29 | "goals": []}
30 |
31 |
--------------------------------------------------------------------------------
/test/Mathlib/test/placeholder_synthesis.in:
--------------------------------------------------------------------------------
1 | {"cmd": "import Mathlib"}
2 |
3 | {"cmd": "example : 1 = 0 := by\n cases' exists_ne 2 with n h\n simpa [] using h _", "env": 0}
4 |
5 | {"cmd": "example : 1 = 0 := sorry", "env": 0}
6 |
7 | {"tactic": "cases' exists_ne 2 with n h", "proofState": 0}
8 |
9 | {"tactic": "simpa [] using h _", "proofState": 1}
10 |
--------------------------------------------------------------------------------
/test/all_tactics.expected.out:
--------------------------------------------------------------------------------
1 | {"tactics":
2 | [{"usedConstants": ["instOfNatNat", "Nat", "OfNat.ofNat", "letFun"],
3 | "tactic": "have t := 37",
4 | "proofState": 0,
5 | "pos": {"line": 1, "column": 18},
6 | "goals": "⊢ Nat",
7 | "endPos": {"line": 1, "column": 30}},
8 | {"usedConstants": [],
9 | "tactic": "exact t",
10 | "proofState": 1,
11 | "pos": {"line": 1, "column": 32},
12 | "goals": "t : Nat\n⊢ Nat",
13 | "endPos": {"line": 1, "column": 39}}],
14 | "env": 0}
15 |
16 |
--------------------------------------------------------------------------------
/test/all_tactics.in:
--------------------------------------------------------------------------------
1 | {"cmd": "def f : Nat := by have t := 37; exact t", "allTactics": true}
2 |
--------------------------------------------------------------------------------
/test/app_type_mismatch.expected.out:
--------------------------------------------------------------------------------
1 | {"messages":
2 | [{"severity": "error",
3 | "pos": {"line": 1, "column": 0},
4 | "endPos": {"line": 1, "column": 7},
5 | "data": "(kernel) declaration has metavariables '_example'"}],
6 | "env": 0}
7 |
8 | {"sorries":
9 | [{"proofState": 0,
10 | "pos": {"line": 1, "column": 19},
11 | "goal": "⊢ 1 = 0",
12 | "endPos": {"line": 1, "column": 24}}],
13 | "messages":
14 | [{"severity": "warning",
15 | "pos": {"line": 1, "column": 0},
16 | "endPos": {"line": 1, "column": 7},
17 | "data": "declaration uses 'sorry'"}],
18 | "env": 1}
19 |
20 | {"proofStatus": "Incomplete: open goals remain",
21 | "proofState": 1,
22 | "goals": ["case zero\n⊢ 0 = 0", "case succ\nn✝ : Nat\n⊢ n✝ + 1 = 0"]}
23 |
24 | {"proofStatus": "Incomplete: open goals remain",
25 | "proofState": 2,
26 | "goals": ["case succ\nn✝ : Nat\n⊢ n✝ + 1 = 0"]}
27 |
28 | {"proofStatus": "Incomplete: contains metavariable(s)",
29 | "proofState": 3,
30 | "goals": []}
31 |
32 |
--------------------------------------------------------------------------------
/test/app_type_mismatch.in:
--------------------------------------------------------------------------------
1 | {"cmd": "example : 1 = 0 := by\n cases 1\n rfl\n apply ?succ"}
2 |
3 | {"cmd": "example : 1 = 0 := sorry"}
4 |
5 | {"tactic": "cases 1", "proofState": 0}
6 |
7 | {"tactic": "rfl", "proofState": 1}
8 |
9 | {"tactic": "apply ?succ", "proofState": 2}
10 |
11 |
--------------------------------------------------------------------------------
/test/app_type_mismatch2.expected.out:
--------------------------------------------------------------------------------
1 | {"messages":
2 | [{"severity": "error",
3 | "pos": {"line": 1, "column": 0},
4 | "endPos": {"line": 1, "column": 7},
5 | "data":
6 | "(kernel) application type mismatch\n have this := of_decide_eq_true (id (Eq.refl true));\n of_decide_eq_true (id (Eq.refl true))\nargument has type\n true = true → true = true\nbut function has type\n (∀ (x : true = true), (fun this => n✝ + 1 = 0) x) → (fun this => n✝ + 1 = 0) ⋯"}],
7 | "env": 0}
8 |
9 | {"sorries":
10 | [{"proofState": 0,
11 | "pos": {"line": 1, "column": 19},
12 | "goal": "⊢ 1 = 0",
13 | "endPos": {"line": 1, "column": 24}}],
14 | "messages":
15 | [{"severity": "warning",
16 | "pos": {"line": 1, "column": 0},
17 | "endPos": {"line": 1, "column": 7},
18 | "data": "declaration uses 'sorry'"}],
19 | "env": 1}
20 |
21 | {"proofStatus": "Incomplete: open goals remain",
22 | "proofState": 1,
23 | "goals": ["case zero\n⊢ 0 = 0", "case succ\nn✝ : Nat\n⊢ n✝ + 1 = 0"]}
24 |
25 | {"proofStatus": "Incomplete: open goals remain",
26 | "proofState": 2,
27 | "goals": ["case succ\nn✝ : Nat\n⊢ n✝ + 1 = 0"]}
28 |
29 | {"proofStatus":
30 | "Error: kernel type check failed: (kernel) application type mismatch\n have this := of_decide_eq_true (id (Eq.refl true));\n of_decide_eq_true (id (Eq.refl true))\nargument has type\n true = true → true = true\nbut function has type\n (∀ (x : true = true), (fun this => n✝ + 1 = 0) x) → (fun this => n✝ + 1 = 0) ⋯",
31 | "proofState": 3,
32 | "goals": []}
33 |
34 | {"sorries": [{"proofState": 4, "goal": "n✝ : Nat\n⊢ true = true"}],
35 | "proofStatus": "Incomplete: open goals remain",
36 | "proofState": 5,
37 | "goals": ["case succ\nn✝ : Nat\nthis : true = true\n⊢ n✝ + 1 = 0"]}
38 |
39 | {"message":
40 | "Lean error:\ntactic 'apply' failed, could not unify the type of `?succ`\n n✝ + 1 = 0\nwith the goal\n true = true\nn✝ : Nat\n⊢ true = true"}
41 |
42 |
--------------------------------------------------------------------------------
/test/app_type_mismatch2.in:
--------------------------------------------------------------------------------
1 | {"cmd": "example : 1 = 0 := by\n cases 1\n rfl\n have: true := by\n apply ?succ\n trivial"}
2 |
3 | {"cmd": "example : 1 = 0 := sorry"}
4 |
5 | {"tactic": "cases 1", "proofState": 0}
6 |
7 | {"tactic": "rfl", "proofState": 1}
8 |
9 | {"tactic": "have: true := by\n apply ?succ\n trivial", "proofState": 2}
10 |
11 | {"tactic": "have: true := sorry", "proofState": 2}
12 |
13 | {"tactic": "apply ?succ", "proofState": 4}
14 |
15 |
--------------------------------------------------------------------------------
/test/assumption_proof.expected.out:
--------------------------------------------------------------------------------
1 | {"sorries":
2 | [{"proofState": 0,
3 | "pos": {"line": 1, "column": 49},
4 | "goal": "x : Nat\nh1 : x = 2\n⊢ x = 2",
5 | "endPos": {"line": 1, "column": 54}}],
6 | "messages":
7 | [{"severity": "warning",
8 | "pos": {"line": 1, "column": 8},
9 | "endPos": {"line": 1, "column": 10},
10 | "data": "declaration uses 'sorry'"}],
11 | "env": 0}
12 |
13 | {"proofStatus": "Completed", "proofState": 1, "goals": []}
14 |
15 |
--------------------------------------------------------------------------------
/test/assumption_proof.in:
--------------------------------------------------------------------------------
1 | {"cmd": "theorem aa (x : Nat) (h1 : x = 2) : x = 2 := by sorry"}
2 |
3 | {"tactic": "assumption", "proofState": 0}
4 |
--------------------------------------------------------------------------------
/test/by_cases.expected.out:
--------------------------------------------------------------------------------
1 | {"sorries":
2 | [{"proofState": 0,
3 | "pos": {"line": 1, "column": 36},
4 | "goal": "x : Int\n⊢ x = x",
5 | "endPos": {"line": 1, "column": 41}}],
6 | "messages":
7 | [{"severity": "warning",
8 | "pos": {"line": 1, "column": 8},
9 | "endPos": {"line": 1, "column": 11},
10 | "data": "declaration uses 'sorry'"}],
11 | "env": 0}
12 |
13 | {"proofStatus": "Incomplete: open goals remain",
14 | "proofState": 1,
15 | "goals":
16 | ["case pos\nx : Int\nh : x < 0\n⊢ x = x",
17 | "case neg\nx : Int\nh : ¬x < 0\n⊢ x = x"]}
18 |
19 | {"proofStatus": "Incomplete: open goals remain",
20 | "proofState": 2,
21 | "goals": ["case neg\nx : Int\nh : ¬x < 0\n⊢ x = x"]}
22 |
23 | {"sorries":
24 | [{"proofState": 3, "goal": "case pos\nx : Int\nh : x < 0\n⊢ x = x"},
25 | {"proofState": 4, "goal": "case neg\nx : Int\nh : ¬x < 0\n⊢ x = x"}],
26 | "proofStatus": "Incomplete: contains sorry",
27 | "proofState": 5,
28 | "goals": []}
29 |
30 |
--------------------------------------------------------------------------------
/test/by_cases.in:
--------------------------------------------------------------------------------
1 | {"cmd": "theorem foo (x : Int) : x = x := by sorry"}
2 |
3 | {"proofState" : 0, "tactic": "by_cases h : x < 0"}
4 |
5 | {"proofState" : 1, "tactic": "case pos => rfl"}
6 |
7 | {"proofState" : 1, "tactic": "all_goals sorry"}
8 |
9 |
--------------------------------------------------------------------------------
/test/by_cases.lean:
--------------------------------------------------------------------------------
1 | theorem foo (x : Int) : x = x := by
2 | by_cases h : x < 0
3 | next => rfl
4 | next => rfl
5 |
--------------------------------------------------------------------------------
/test/calc.expected.out:
--------------------------------------------------------------------------------
1 | {"tactics":
2 | [{"usedConstants":
3 | ["Trans.trans",
4 | "sorryAx",
5 | "instOfNatNat",
6 | "Lean.Name.num",
7 | "Lean.Name.str",
8 | "Lean.Name.anonymous",
9 | "instTransEq",
10 | "Nat",
11 | "Lean.Name",
12 | "OfNat.ofNat",
13 | "Bool.false",
14 | "Eq"],
15 | "tactic": "calc\n 3 = 4 := by sorry\n 4 = 5 := by sorry",
16 | "proofState": 2,
17 | "pos": {"line": 1, "column": 22},
18 | "goals": "⊢ 3 = 5",
19 | "endPos": {"line": 3, "column": 19}},
20 | {"usedConstants": [],
21 | "tactic": "\n 3 = 4 := by sorry\n 4 = 5 := by sorry",
22 | "proofState": 3,
23 | "pos": {"line": 2, "column": 2},
24 | "goals": "no goals",
25 | "endPos": {"line": 3, "column": 19}},
26 | {"usedConstants":
27 | ["sorryAx",
28 | "instOfNatNat",
29 | "Lean.Name.num",
30 | "Lean.Name.str",
31 | "Lean.Name.anonymous",
32 | "Nat",
33 | "Lean.Name",
34 | "OfNat.ofNat",
35 | "Bool.false",
36 | "Eq"],
37 | "tactic": "sorry",
38 | "proofState": 4,
39 | "pos": {"line": 2, "column": 14},
40 | "goals": "⊢ 3 = 4",
41 | "endPos": {"line": 2, "column": 19}},
42 | {"usedConstants":
43 | ["sorryAx",
44 | "instOfNatNat",
45 | "Lean.Name.num",
46 | "Lean.Name.str",
47 | "Lean.Name.anonymous",
48 | "Nat",
49 | "Lean.Name",
50 | "OfNat.ofNat",
51 | "Bool.false",
52 | "Eq"],
53 | "tactic": "sorry",
54 | "proofState": 5,
55 | "pos": {"line": 3, "column": 14},
56 | "goals": "⊢ 4 = 5",
57 | "endPos": {"line": 3, "column": 19}}],
58 | "sorries":
59 | [{"proofState": 0,
60 | "pos": {"line": 2, "column": 14},
61 | "goal": "⊢ 3 = 4",
62 | "endPos": {"line": 2, "column": 19}},
63 | {"proofState": 1,
64 | "pos": {"line": 3, "column": 14},
65 | "goal": "⊢ 4 = 5",
66 | "endPos": {"line": 3, "column": 19}}],
67 | "messages":
68 | [{"severity": "warning",
69 | "pos": {"line": 1, "column": 0},
70 | "endPos": {"line": 1, "column": 7},
71 | "data": "declaration uses 'sorry'"}],
72 | "env": 0}
73 |
74 |
--------------------------------------------------------------------------------
/test/calc.in:
--------------------------------------------------------------------------------
1 | {"cmd": "example : 3 = 5 := by calc\n 3 = 4 := by sorry\n 4 = 5 := by sorry", "allTactics": true }
2 |
--------------------------------------------------------------------------------
/test/def_eval.expected.out:
--------------------------------------------------------------------------------
1 | {"env": 0}
2 |
3 | {"messages":
4 | [{"severity": "info",
5 | "pos": {"line": 1, "column": 0},
6 | "endPos": {"line": 1, "column": 6},
7 | "data": "rfl : f = f"}],
8 | "env": 1}
9 |
10 |
--------------------------------------------------------------------------------
/test/def_eval.in:
--------------------------------------------------------------------------------
1 | {"cmd": "def f := 37"}
2 |
3 | {"cmd": "#check (rfl : f = 37)", "env": 0}
4 |
--------------------------------------------------------------------------------
/test/dup_msg.expected.out:
--------------------------------------------------------------------------------
1 | {"env": 0}
2 |
3 | {"messages":
4 | [{"severity": "info",
5 | "pos": {"line": 1, "column": 0},
6 | "endPos": {"line": 1, "column": 6},
7 | "data": "f : Nat"}],
8 | "env": 1}
9 |
10 | {"messages":
11 | [{"severity": "error",
12 | "pos": {"line": 1, "column": 7},
13 | "endPos": {"line": 1, "column": 8},
14 | "data": "unknown identifier 'g'"}],
15 | "env": 2}
16 |
17 |
--------------------------------------------------------------------------------
/test/dup_msg.in:
--------------------------------------------------------------------------------
1 | {"cmd": "def f := 2"}
2 |
3 | {"cmd": "#check f", "env": 0}
4 |
5 | {"cmd": "#check g", "env": 1}
6 |
--------------------------------------------------------------------------------
/test/dup_sorries.expected.out:
--------------------------------------------------------------------------------
1 | {"sorries":
2 | [{"proofState": 0,
3 | "pos": {"line": 1, "column": 24},
4 | "goal": "⊢ 1 = 1",
5 | "endPos": {"line": 1, "column": 29}}],
6 | "messages":
7 | [{"severity": "warning",
8 | "pos": {"line": 1, "column": 8},
9 | "endPos": {"line": 1, "column": 12},
10 | "data": "declaration uses 'sorry'"}],
11 | "env": 0}
12 |
13 | {"sorries":
14 | [{"proofState": 1,
15 | "pos": {"line": 1, "column": 24},
16 | "goal": "⊢ 2 = 2",
17 | "endPos": {"line": 1, "column": 29}}],
18 | "messages":
19 | [{"severity": "warning",
20 | "pos": {"line": 1, "column": 8},
21 | "endPos": {"line": 1, "column": 12},
22 | "data": "declaration uses 'sorry'"}],
23 | "env": 1}
24 |
25 |
--------------------------------------------------------------------------------
/test/dup_sorries.in:
--------------------------------------------------------------------------------
1 | {"cmd": "theorem thm1 : 1 = 1 := sorry"}
2 |
3 | {"cmd": "theorem thm2 : 2 = 2 := sorry", "env": 0}
4 |
--------------------------------------------------------------------------------
/test/file.expected.out:
--------------------------------------------------------------------------------
1 | {"tactics":
2 | [{"usedConstants":
3 | ["f", "g", "instHAdd", "HAdd.hAdd", "Nat", "instAddNat", "rfl"],
4 | "tactic": "exact rfl",
5 | "proofState": 0,
6 | "pos": {"line": 5, "column": 29},
7 | "goals": "⊢ f + g = 39",
8 | "endPos": {"line": 5, "column": 38}}],
9 | "env": 0}
10 |
11 |
--------------------------------------------------------------------------------
/test/file.in:
--------------------------------------------------------------------------------
1 | {"path": "test/file.lean", "allTactics": true}
2 |
--------------------------------------------------------------------------------
/test/file.lean:
--------------------------------------------------------------------------------
1 | def f : Nat := 37
2 |
3 | def g := 2
4 |
5 | theorem h : f + g = 39 := by exact rfl
6 |
--------------------------------------------------------------------------------
/test/have_by_sorry.expected.out:
--------------------------------------------------------------------------------
1 | {"sorries":
2 | [{"proofState": 0,
3 | "pos": {"line": 2, "column": 23},
4 | "goal": "x : Int\n⊢ x = 1",
5 | "endPos": {"line": 2, "column": 28}}],
6 | "messages":
7 | [{"severity": "error",
8 | "pos": {"line": 1, "column": 33},
9 | "endPos": {"line": 2, "column": 28},
10 | "data": "unsolved goals\nx : Int\nh : x = 1\n⊢ x = x"}],
11 | "env": 0}
12 |
13 | {"sorries":
14 | [{"proofState": 1,
15 | "pos": {"line": 1, "column": 36},
16 | "goal": "x : Int\n⊢ x = x",
17 | "endPos": {"line": 1, "column": 41}}],
18 | "messages":
19 | [{"severity": "warning",
20 | "pos": {"line": 1, "column": 8},
21 | "endPos": {"line": 1, "column": 11},
22 | "data": "declaration uses 'sorry'"}],
23 | "env": 1}
24 |
25 | {"sorries": [{"proofState": 2, "goal": "x : Int\n⊢ x = 1"}],
26 | "proofStatus": "Incomplete: open goals remain",
27 | "proofState": 3,
28 | "goals": ["x : Int\nh : x = 1\n⊢ x = 1"]}
29 |
30 |
--------------------------------------------------------------------------------
/test/have_by_sorry.in:
--------------------------------------------------------------------------------
1 | {"cmd": "theorem foo (x : Int) : x = x := by\n have h : x = 1 := by sorry"}
2 |
3 | {"cmd": "theorem foo (x : Int) : x = x := by sorry"}
4 |
5 | {"proofState" : 0, "tactic": "have h : x = 1 := by sorry"}
6 |
7 |
--------------------------------------------------------------------------------
/test/import_lean.expected.out:
--------------------------------------------------------------------------------
1 | {"env": 0}
2 |
3 |
--------------------------------------------------------------------------------
/test/import_lean.in:
--------------------------------------------------------------------------------
1 | {"cmd" : "import Lean"}
2 |
--------------------------------------------------------------------------------
/test/incomplete.expected.out:
--------------------------------------------------------------------------------
1 | {"messages":
2 | [{"severity": "error",
3 | "pos": {"line": 1, "column": 15},
4 | "endPos": {"line": 1, "column": 32},
5 | "data": "unsolved goals\n⊢ Nat"}],
6 | "env": 0}
7 |
8 | {"messages":
9 | [{"severity": "error",
10 | "pos": {"line": 3, "column": 19},
11 | "endPos": {"line": 3, "column": 20},
12 | "data": "unsolved goals\ncase pos\nx : Bool\nh✝ : x = true\n⊢ Nat"},
13 | {"severity": "error",
14 | "pos": {"line": 1, "column": 26},
15 | "endPos": {"line": 3, "column": 20},
16 | "data": "unsolved goals\ncase neg\nx : Bool\nh✝ : ¬x = true\n⊢ Nat"}],
17 | "env": 1}
18 |
19 |
--------------------------------------------------------------------------------
/test/incomplete.in:
--------------------------------------------------------------------------------
1 | {"cmd": "def f : Nat := by apply Nat.succ"}
2 |
3 | {"cmd": "def f (x : Bool) : Nat := by\n by_cases x\n { apply Nat.succ }"}
4 |
--------------------------------------------------------------------------------
/test/incomplete.lean:
--------------------------------------------------------------------------------
https://raw.githubusercontent.com/leanprover-community/repl/ab9da2b867c54d709fa3720790ea7406606be28a/test/incomplete.lean
--------------------------------------------------------------------------------
/test/infotree.expected.out:
--------------------------------------------------------------------------------
1 | {"messages":
2 | [{"severity": "warning",
3 | "pos": {"line": 1, "column": 7},
4 | "endPos": {"line": 1, "column": 8},
5 | "data":
6 | "unused variable `h`\nnote: this linter can be disabled with `set_option linter.unusedVariables false`"}],
7 | "infotree":
8 | [{"node":
9 | {"stx":
10 | {"range":
11 | {"synthetic": false,
12 | "start": {"line": 1, "column": 28},
13 | "finish": {"line": 1, "column": 39}},
14 | "pp": "constructor"},
15 | "name": "Lean.Parser.Tactic.constructor",
16 | "goalsBefore": ["h : Int\n⊢ Nat"],
17 | "goalsAfter": []},
18 | "kind": "TacticInfo",
19 | "children": []}],
20 | "env": 0}
21 |
22 |
--------------------------------------------------------------------------------
/test/infotree.in:
--------------------------------------------------------------------------------
1 | {"cmd": "def f (h : Int) : Nat := by constructor", "infotree": "substantive"}
2 |
3 |
--------------------------------------------------------------------------------
/test/invalid_tactic.expected.out:
--------------------------------------------------------------------------------
1 | {"sorries":
2 | [{"proofState": 0,
3 | "pos": {"line": 1, "column": 43},
4 | "goal": "x : Nat\n⊢ x = x",
5 | "endPos": {"line": 1, "column": 48}}],
6 | "messages":
7 | [{"severity": "warning",
8 | "pos": {"line": 1, "column": 8},
9 | "endPos": {"line": 1, "column": 18},
10 | "data": "declaration uses 'sorry'"}],
11 | "env": 0}
12 |
13 | {"proofStatus": "Incomplete: contains sorry",
14 | "proofState": 1,
15 | "messages":
16 | [{"severity": "error",
17 | "pos": {"line": 0, "column": 0},
18 | "endPos": {"line": 0, "column": 0},
19 | "data": "unknown identifier 'my_fake_premise'"}],
20 | "goals": []}
21 |
22 |
--------------------------------------------------------------------------------
/test/invalid_tactic.in:
--------------------------------------------------------------------------------
1 | {"cmd": "theorem my_theorem (x : Nat) : x = x := by sorry"}
2 |
3 | {"tactic": "exact my_fake_premise", "proofState": 0}
4 |
--------------------------------------------------------------------------------
/test/name_generator.expected.out:
--------------------------------------------------------------------------------
1 | {"sorries":
2 | [{"proofState": 0,
3 | "pos": {"line": 2, "column": 45},
4 | "goal": "x : Int\nh0 : x > 0\n⊢ False",
5 | "endPos": {"line": 2, "column": 50}}],
6 | "messages":
7 | [{"severity": "warning",
8 | "pos": {"line": 2, "column": 0},
9 | "endPos": {"line": 2, "column": 7},
10 | "data": "declaration uses 'sorry'"}],
11 | "env": 0}
12 |
13 | {"sorries": [{"proofState": 1, "goal": "x : Int\nh0 : x > 0\n⊢ x > 0"}],
14 | "proofStatus": "Incomplete: open goals remain",
15 | "proofState": 2,
16 | "goals": ["x : Int\nh0 h : x > 0\n⊢ False"]}
17 |
18 | {"proofStatus": "Incomplete: contains metavariable(s)",
19 | "proofState": 3,
20 | "goals": []}
21 |
22 | {"proofStatus": "Incomplete: contains metavariable(s)",
23 | "proofState": 4,
24 | "goals": []}
25 |
26 | {"traces":
27 | ["[Meta.Tactic.simp.rewrite] of_eq_true (eq_true h0):1000:\n x > 0\n ==>\n True"],
28 | "proofStatus": "Incomplete: contains metavariable(s)",
29 | "proofState": 5,
30 | "goals": []}
31 |
32 | {"traces":
33 | ["[Meta.Tactic.simp.rewrite] gt_iff_lt:1000:\n x > 0\n ==>\n 0 < x",
34 | "[Meta.Tactic.simp.rewrite] h0:1000:\n 0 < x\n ==>\n True"],
35 | "proofStatus": "Incomplete: contains metavariable(s)",
36 | "proofState": 6,
37 | "goals": []}
38 |
39 | {"traces":
40 | ["[Meta.Tactic.simp.rewrite] gt_iff_lt:1000:\n x > 0\n ==>\n 0 < x",
41 | "[Meta.Tactic.simp.rewrite] h0:1000:\n 0 < x\n ==>\n True"],
42 | "proofStatus": "Incomplete: contains metavariable(s)",
43 | "proofState": 7,
44 | "goals": []}
45 |
46 | {"proofStatus": "Incomplete: contains metavariable(s)",
47 | "proofState": 8,
48 | "messages":
49 | [{"severity": "error",
50 | "pos": {"line": 0, "column": 0},
51 | "endPos": {"line": 0, "column": 0},
52 | "data": "unsolved goals\nx : Int\nh0 : x > 0\n⊢ x > 0"}],
53 | "goals": []}
54 |
55 | {"proofStatus": "Incomplete: contains metavariable(s)",
56 | "proofState": 9,
57 | "messages":
58 | [{"severity": "error",
59 | "pos": {"line": 0, "column": 0},
60 | "endPos": {"line": 0, "column": 0},
61 | "data": "unsolved goals\nx : Int\nh0 : x > 0\n⊢ x > 0"}],
62 | "goals": []}
63 |
64 |
--------------------------------------------------------------------------------
/test/name_generator.in:
--------------------------------------------------------------------------------
1 | {"cmd": "set_option trace.Meta.Tactic.simp true\nexample {x : Int} (h0 : x > 0) : False := by sorry"}
2 |
3 | {"tactic": "have h : x > 0 := by sorry", "proofState": 0}
4 |
5 | {"tactic": "exact h0", "proofState": 1}
6 |
7 | {"tactic": "assumption", "proofState": 1}
8 |
9 | {"tactic": "simp only [of_eq_true (eq_true h0)]", "proofState": 1}
10 |
11 | {"tactic": "{ simp [h0] }", "proofState": 1}
12 |
13 | {"tactic": "{ simp (config := {memoize := false}) [h0] }", "proofState": 1}
14 |
15 | {"tactic": "{ rw (config := {}) [(show x = x from rfl)] }", "proofState": 1}
16 |
17 | {"tactic": "{ rw [(show x = x from rfl)] }", "proofState": 1}
18 |
19 |
--------------------------------------------------------------------------------
/test/no_goal_sorry.expected.out:
--------------------------------------------------------------------------------
1 | {"messages":
2 | [{"severity": "error",
3 | "pos": {"line": 2, "column": 11},
4 | "endPos": {"line": 2, "column": 18},
5 | "data": "type expected, got\n (set Nat : ?m.8 PUnit)"},
6 | {"severity": "error",
7 | "pos": {"line": 2, "column": 22},
8 | "endPos": {"line": 2, "column": 24},
9 | "data": "Case tag 'body' not found.\n\nThere are no cases to select."}],
10 | "env": 0}
11 |
12 |
--------------------------------------------------------------------------------
/test/no_goal_sorry.in:
--------------------------------------------------------------------------------
1 | {"cmd" : "example : True := by\n have h : set Nat := by sorry\n sorry"}
2 |
--------------------------------------------------------------------------------
/test/no_goal_sorry_2.expected.out:
--------------------------------------------------------------------------------
1 | {"sorries":
2 | [{"proofState": 0,
3 | "pos": {"line": 2, "column": 2},
4 | "goal": "⊢ True",
5 | "endPos": {"line": 2, "column": 7}}],
6 | "messages":
7 | [{"severity": "error",
8 | "pos": {"line": 3, "column": 2},
9 | "endPos": {"line": 3, "column": 7},
10 | "data": "no goals to be solved"}],
11 | "env": 0}
12 |
13 |
--------------------------------------------------------------------------------
/test/no_goal_sorry_2.in:
--------------------------------------------------------------------------------
1 | {"cmd" : "example : True := by\n sorry\n sorry"}
2 |
--------------------------------------------------------------------------------
/test/options.expected.out:
--------------------------------------------------------------------------------
1 | {"messages":
2 | [{"severity": "info",
3 | "pos": {"line": 1, "column": 0},
4 | "endPos": {"line": 1, "column": 6},
5 | "data": "constructor List.cons.{u} : {α : Type u} → α → List α → List α"}],
6 | "env": 0}
7 |
8 | {"env": 1}
9 |
10 | {"messages":
11 | [{"severity": "info",
12 | "pos": {"line": 1, "column": 0},
13 | "endPos": {"line": 1, "column": 6},
14 | "data":
15 | "constructor List.cons.{u} : {α : Type u} → α → List.{u} α → List.{u} α"}],
16 | "env": 2}
17 |
18 |
--------------------------------------------------------------------------------
/test/options.in:
--------------------------------------------------------------------------------
1 | {"cmd": "#print List.cons"}
2 |
3 | {"cmd": "set_option pp.universes true"}
4 |
5 | {"cmd": "#print List.cons", "env": 1}
6 |
7 |
--------------------------------------------------------------------------------
/test/pickle_environment.expected.out:
--------------------------------------------------------------------------------
1 | {"env": 0}
2 |
3 | {"env": 0}
4 |
5 | {"env": 1}
6 |
7 | {"env": 2}
8 |
9 |
--------------------------------------------------------------------------------
/test/pickle_environment.in:
--------------------------------------------------------------------------------
1 | {"cmd": "def f := 42"}
2 |
3 | {"pickleTo": "test/a.olean", "env": 0}
4 |
5 | {"unpickleEnvFrom": "test/a.olean"}
6 |
7 | {"cmd": "example : f = 42 := by rfl", "env": 1}
8 |
--------------------------------------------------------------------------------
/test/pickle_environment_with_imports.expected.out:
--------------------------------------------------------------------------------
1 | {"env": 0}
2 |
3 | {"env": 1}
4 |
5 | {"env": 1}
6 |
7 | {"env": 2}
8 |
9 | {"env": 3}
10 |
11 |
--------------------------------------------------------------------------------
/test/pickle_environment_with_imports.in:
--------------------------------------------------------------------------------
1 | {"cmd": "import Lean"}
2 |
3 | {"cmd": "def f := 42", "env": 0}
4 |
5 | {"pickleTo": "test/b.olean", "env": 1}
6 |
7 | {"unpickleEnvFrom": "test/b.olean"}
8 |
9 | {"cmd": "example : f = 42 := by rfl", "env": 2}
10 |
--------------------------------------------------------------------------------
/test/pickle_open.expected.out:
--------------------------------------------------------------------------------
1 | {"env": 0}
2 |
3 | {"env": 1}
4 |
5 | {"env": 2}
6 |
7 | {"env": 1}
8 |
9 |
--------------------------------------------------------------------------------
/test/pickle_open.in:
--------------------------------------------------------------------------------
1 | {"cmd": "def X.Y : Nat := 37"}
2 |
3 | {"cmd": "open X", "env": 0}
4 |
5 | {"cmd": "example : Y = 37 := rfl", "env": 1}
6 |
7 | {"pickleTo": "test/e.olean", "env": 1}
8 |
--------------------------------------------------------------------------------
/test/pickle_open_2.expected.out:
--------------------------------------------------------------------------------
1 | {"env": 0}
2 |
3 | {"env": 1}
4 |
5 |
--------------------------------------------------------------------------------
/test/pickle_open_2.in:
--------------------------------------------------------------------------------
1 | {"unpickleEnvFrom": "test/e.olean"}
2 |
3 | {"cmd": "example : Y = 37 := rfl", "env": 0}
4 |
--------------------------------------------------------------------------------
/test/pickle_open_scoped.expected.out:
--------------------------------------------------------------------------------
1 | {"env": 0}
2 |
3 | {"env": 1}
4 |
5 | {"sorries":
6 | [{"proofState": 0,
7 | "pos": {"line": 1, "column": 22},
8 | "goal": "⊢ ◾",
9 | "endPos": {"line": 1, "column": 27}}],
10 | "messages":
11 | [{"severity": "warning",
12 | "pos": {"line": 1, "column": 7},
13 | "endPos": {"line": 1, "column": 14},
14 | "data": "declaration uses 'sorry'"}],
15 | "env": 2}
16 |
17 | {"env": 1}
18 |
19 |
--------------------------------------------------------------------------------
/test/pickle_open_scoped.in:
--------------------------------------------------------------------------------
1 | {"cmd": "import Lean"}
2 |
3 | {"cmd": "open Lean.Compiler", "env": 0}
4 |
5 | {"cmd": "unsafe example : ◾ := sorry", "env": 1}
6 |
7 | {"pickleTo": "test/f.olean", "env": 1}
8 |
9 |
--------------------------------------------------------------------------------
/test/pickle_open_scoped_2.expected.out:
--------------------------------------------------------------------------------
1 | {"env": 0}
2 |
3 | {"sorries":
4 | [{"proofState": 0,
5 | "pos": {"line": 1, "column": 22},
6 | "goal": "⊢ ◾",
7 | "endPos": {"line": 1, "column": 27}}],
8 | "messages":
9 | [{"severity": "warning",
10 | "pos": {"line": 1, "column": 7},
11 | "endPos": {"line": 1, "column": 14},
12 | "data": "declaration uses 'sorry'"}],
13 | "env": 1}
14 |
15 |
--------------------------------------------------------------------------------
/test/pickle_open_scoped_2.in:
--------------------------------------------------------------------------------
1 | {"unpickleEnvFrom": "test/f.olean"}
2 |
3 | {"cmd": "unsafe example : ◾ := sorry", "env": 0}
4 |
--------------------------------------------------------------------------------
/test/pickle_proof_state_1.expected.out:
--------------------------------------------------------------------------------
1 | {"env": 0}
2 |
3 | {"sorries":
4 | [{"proofState": 0,
5 | "pos": {"line": 1, "column": 18},
6 | "goal": "⊢ Nat",
7 | "endPos": {"line": 1, "column": 23}}],
8 | "messages":
9 | [{"severity": "warning",
10 | "pos": {"line": 1, "column": 4},
11 | "endPos": {"line": 1, "column": 5},
12 | "data": "declaration uses 'sorry'"}],
13 | "env": 1}
14 |
15 | {"proofStatus": "Incomplete: open goals remain",
16 | "proofState": 1,
17 | "goals": ["⊢ Nat"]}
18 |
19 | {"proofStatus": "Incomplete: open goals remain",
20 | "proofState": 2,
21 | "goals": ["⊢ Nat"]}
22 |
23 | {"proofStatus": "Incomplete: open goals remain",
24 | "proofState": 3,
25 | "goals": ["t : Nat\n⊢ Nat"]}
26 |
27 | {"proofStatus": "Incomplete: open goals remain",
28 | "proofState": 4,
29 | "goals": ["t : Nat\n⊢ Nat"]}
30 |
31 | {"proofStatus": "Incomplete: open goals remain",
32 | "proofState": 5,
33 | "goals": ["t : Nat\n⊢ Nat"]}
34 |
35 | {"proofStatus": "Completed", "proofState": 6, "goals": []}
36 |
37 |
--------------------------------------------------------------------------------
/test/pickle_proof_state_1.in:
--------------------------------------------------------------------------------
1 | {"cmd" : "import Lean"}
2 |
3 | {"cmd" : "def f : Nat := by sorry", "env": 0}
4 |
5 | {"pickleTo": "test/c.olean", "proofState": 0}
6 |
7 | {"unpickleProofStateFrom": "test/c.olean"}
8 |
9 | {"tactic": "have t : Nat := 42", "proofState": 2}
10 |
11 | {"pickleTo": "test/d.olean", "proofState": 3}
12 |
13 | {"unpickleProofStateFrom": "test/d.olean"}
14 |
15 | {"tactic": "exact t", "proofState": 5}
16 |
--------------------------------------------------------------------------------
/test/pickle_proof_state_2.expected.out:
--------------------------------------------------------------------------------
1 | {"proofStatus": "Incomplete: open goals remain",
2 | "proofState": 0,
3 | "goals": ["t : Nat\n⊢ Nat"]}
4 |
5 | {"proofStatus": "Completed", "proofState": 1, "goals": []}
6 |
7 |
--------------------------------------------------------------------------------
/test/pickle_proof_state_2.in:
--------------------------------------------------------------------------------
1 | {"unpickleProofStateFrom": "test/d.olean"}
2 |
3 | {"tactic": "exact t", "proofState": 0}
4 |
--------------------------------------------------------------------------------
/test/pickle_proof_state_env.expected.out:
--------------------------------------------------------------------------------
1 | {"env": 0}
2 |
3 | {"sorries":
4 | [{"proofState": 0,
5 | "pos": {"line": 1, "column": 18},
6 | "goal": "⊢ Nat",
7 | "endPos": {"line": 1, "column": 23}}],
8 | "messages":
9 | [{"severity": "warning",
10 | "pos": {"line": 1, "column": 4},
11 | "endPos": {"line": 1, "column": 5},
12 | "data": "declaration uses 'sorry'"}],
13 | "env": 1}
14 |
15 | {"proofStatus": "Incomplete: open goals remain",
16 | "proofState": 1,
17 | "goals": ["⊢ Nat"]}
18 |
19 | {"proofStatus": "Incomplete: open goals remain",
20 | "proofState": 2,
21 | "goals": ["⊢ Nat"]}
22 |
23 | {"proofStatus": "Incomplete: open goals remain",
24 | "proofState": 3,
25 | "goals": ["t : Nat\n⊢ Nat"]}
26 |
27 | {"proofStatus": "Incomplete: open goals remain",
28 | "proofState": 4,
29 | "goals": ["t : Nat\n⊢ Nat"]}
30 |
31 | {"proofStatus": "Incomplete: open goals remain",
32 | "proofState": 5,
33 | "goals": ["t : Nat\n⊢ Nat"]}
34 |
35 | {"proofStatus": "Completed", "proofState": 6, "goals": []}
36 |
37 |
--------------------------------------------------------------------------------
/test/pickle_proof_state_env.in:
--------------------------------------------------------------------------------
1 | {"cmd" : "import Lean"}
2 |
3 | {"cmd" : "def f : Nat := by sorry", "env": 0}
4 |
5 | {"pickleTo": "test/c.olean", "proofState": 0}
6 |
7 | {"unpickleProofStateFrom": "test/c.olean", "env": 0}
8 |
9 | {"tactic": "have t : Nat := 42", "proofState": 2}
10 |
11 | {"pickleTo": "test/d.olean", "proofState": 3}
12 |
13 | {"unpickleProofStateFrom": "test/d.olean", "env": 0}
14 |
15 | {"tactic": "exact t", "proofState": 5}
16 |
--------------------------------------------------------------------------------
/test/pickle_scoped_notation.in:
--------------------------------------------------------------------------------
1 | {"cmd": "import Lean"}
2 |
3 | {"cmd": "namespace X", "env": 0}
4 |
5 | {"cmd": "def f (x : Nat) := x + 37", "env": 1}
6 |
7 | {"cmd": "scoped notation:10000 n \"!\" => f n", "env": 2}
8 |
9 | {"cmd": "end X", "env": 3}
10 |
11 | {"cmd": "open X", "env": 4}
12 |
13 | {"cmd": "example : 39 = 2 ! := rfl", "env": 5}
14 |
15 | {"pickleTo": "test/f.olean", "env": 5}
16 |
17 |
--------------------------------------------------------------------------------
/test/pickle_scoped_notation_2.in:
--------------------------------------------------------------------------------
1 | {"unpickleEnvFrom": "test/f.olean"}
2 |
3 | {"cmd": "example : 39 = 2 ! := rfl", "env": 0}
4 |
--------------------------------------------------------------------------------
/test/proof_branching.expected.out:
--------------------------------------------------------------------------------
1 | {"sorries":
2 | [{"proofState": 0,
3 | "pos": {"line": 1, "column": 75},
4 | "goal": "p q r : Prop\nh1 : p ∧ q\nh2 : q → r\n⊢ p ∧ r",
5 | "endPos": {"line": 1, "column": 80}}],
6 | "messages":
7 | [{"severity": "warning",
8 | "pos": {"line": 1, "column": 8},
9 | "endPos": {"line": 1, "column": 19},
10 | "data": "declaration uses 'sorry'"}],
11 | "env": 0}
12 |
13 | {"proofStatus": "Incomplete: open goals remain",
14 | "proofState": 1,
15 | "goals":
16 | ["case left\np q r : Prop\nh1 : p ∧ q\nh2 : q → r\n⊢ p",
17 | "case right\np q r : Prop\nh1 : p ∧ q\nh2 : q → r\n⊢ r"]}
18 |
19 | {"proofStatus": "Incomplete: open goals remain",
20 | "proofState": 2,
21 | "goals": ["case right\np q r : Prop\nh1 : p ∧ q\nh2 : q → r\n⊢ r"]}
22 |
23 | {"proofStatus": "Incomplete: open goals remain",
24 | "proofState": 3,
25 | "goals": ["case right\np q r : Prop\nh1 : p ∧ q\nh2 : q → r\n⊢ q"]}
26 |
27 | {"proofStatus": "Completed", "proofState": 4, "goals": []}
28 |
29 |
--------------------------------------------------------------------------------
/test/proof_branching.in:
--------------------------------------------------------------------------------
1 | {"cmd": "theorem complex_and (p q r : Prop) (h1 : p ∧ q) (h2 : q → r) : p ∧ r := by sorry"}
2 |
3 | {"tactic": "apply And.intro", "proofState": 0}
4 |
5 | {"tactic": "exact h1.left", "proofState": 1}
6 |
7 | {"tactic": "apply h2", "proofState": 2}
8 |
9 | {"tactic": "exact h1.right", "proofState": 3}
10 |
11 |
--------------------------------------------------------------------------------
/test/proof_branching2.expected.out:
--------------------------------------------------------------------------------
1 | {"sorries":
2 | [{"proofState": 0,
3 | "pos": {"line": 1, "column": 43},
4 | "goal": "p q : Prop\n⊢ p ∧ q → q ∧ p",
5 | "endPos": {"line": 1, "column": 48}}],
6 | "messages":
7 | [{"severity": "warning",
8 | "pos": {"line": 1, "column": 0},
9 | "endPos": {"line": 1, "column": 7},
10 | "data": "declaration uses 'sorry'"}],
11 | "env": 0}
12 |
13 | {"proofStatus": "Incomplete: open goals remain",
14 | "proofState": 1,
15 | "goals": ["p q : Prop\nh : p ∧ q\n⊢ q ∧ p"]}
16 |
17 | {"proofStatus": "Incomplete: open goals remain",
18 | "proofState": 2,
19 | "goals": ["p q : Prop\nh : p ∧ q\nhp : p\n⊢ q ∧ p"]}
20 |
21 | {"proofStatus": "Incomplete: open goals remain",
22 | "proofState": 3,
23 | "goals": ["p q : Prop\nh : p ∧ q\nhp : p\nhq : q\n⊢ q ∧ p"]}
24 |
25 | {"proofStatus": "Incomplete: open goals remain",
26 | "proofState": 4,
27 | "goals":
28 | ["case left\np q : Prop\nh : p ∧ q\nhp : p\nhq : q\n⊢ q",
29 | "case right\np q : Prop\nh : p ∧ q\nhp : p\nhq : q\n⊢ p"]}
30 |
31 | {"proofStatus": "Incomplete: open goals remain",
32 | "proofState": 5,
33 | "goals": ["case right\np q : Prop\nh : p ∧ q\nhp : p\nhq : q\n⊢ p"]}
34 |
35 | {"proofStatus": "Completed", "proofState": 6, "goals": []}
36 |
37 |
--------------------------------------------------------------------------------
/test/proof_branching2.in:
--------------------------------------------------------------------------------
1 | {"cmd": "example (p q : Prop) : p ∧ q → q ∧ p := by sorry"}
2 |
3 | {"tactic": "intro h", "proofState": 0}
4 |
5 | {"tactic": "have hp : p := h.left", "proofState": 1}
6 |
7 | {"tactic": "have hq : q := h.right", "proofState": 2}
8 |
9 | {"tactic": "apply And.intro", "proofState": 3}
10 |
11 | {"tactic": "exact hq", "proofState": 4}
12 |
13 | {"tactic": "exact hp", "proofState": 5}
14 |
15 |
--------------------------------------------------------------------------------
/test/proof_step.expected.out:
--------------------------------------------------------------------------------
1 | {"sorries":
2 | [{"proofState": 0,
3 | "pos": {"line": 1, "column": 18},
4 | "goal": "⊢ Nat",
5 | "endPos": {"line": 1, "column": 23}}],
6 | "messages":
7 | [{"severity": "warning",
8 | "pos": {"line": 1, "column": 4},
9 | "endPos": {"line": 1, "column": 5},
10 | "data": "declaration uses 'sorry'"}],
11 | "env": 0}
12 |
13 | {"proofStatus": "Incomplete: open goals remain",
14 | "proofState": 1,
15 | "goals": ["⊢ Int"]}
16 |
17 | {"proofStatus": "Incomplete: open goals remain",
18 | "proofState": 2,
19 | "goals": ["t : Nat\n⊢ Nat"]}
20 |
21 | {"proofStatus": "Completed", "proofState": 3, "goals": []}
22 |
23 |
--------------------------------------------------------------------------------
/test/proof_step.in:
--------------------------------------------------------------------------------
1 | {"cmd" : "def f : Nat := by sorry"}
2 |
3 | {"tactic": "apply Int.natAbs", "proofState": 0}
4 |
5 | {"tactic": "have t : Nat := 42", "proofState": 0}
6 |
7 | {"tactic": "exact t", "proofState": 2}
8 |
--------------------------------------------------------------------------------
/test/proof_transitivity.expected.out:
--------------------------------------------------------------------------------
1 | {"sorries":
2 | [{"proofState": 0,
3 | "pos": {"line": 1, "column": 62},
4 | "goal": "x y z : Nat\nh1 : x = y\nh2 : y = z\n⊢ x = z",
5 | "endPos": {"line": 1, "column": 67}}],
6 | "messages":
7 | [{"severity": "warning",
8 | "pos": {"line": 1, "column": 0},
9 | "endPos": {"line": 1, "column": 7},
10 | "data": "declaration uses 'sorry'"}],
11 | "env": 0}
12 |
13 | {"proofStatus": "Completed", "proofState": 1, "goals": []}
14 |
15 | {"sorries":
16 | [{"proofState": 2,
17 | "pos": {"line": 1, "column": 64},
18 | "goal": "f : Nat → Nat\nn : Nat\nh : n = 3\n⊢ f n = f 3",
19 | "endPos": {"line": 1, "column": 69}}],
20 | "messages":
21 | [{"severity": "warning",
22 | "pos": {"line": 1, "column": 0},
23 | "endPos": {"line": 1, "column": 7},
24 | "data": "declaration uses 'sorry'"}],
25 | "env": 1}
26 |
27 | {"proofStatus": "Completed", "proofState": 3, "goals": []}
28 |
29 |
--------------------------------------------------------------------------------
/test/proof_transitivity.in:
--------------------------------------------------------------------------------
1 | {"cmd": "example (x y z : Nat) (h1 : x = y) (h2 : y = z) : x = z := by sorry"}
2 |
3 | {"tactic": "exact Eq.trans h1 h2", "proofState": 0}
4 |
5 | {"cmd": "example (f : Nat → Nat) (n : Nat) (h : n = 3) : f n = f 3 := by sorry"}
6 |
7 | {"tactic": "exact congrArg f h", "proofState": 2}
8 |
9 |
--------------------------------------------------------------------------------
/test/readme.expected.out:
--------------------------------------------------------------------------------
1 | {"sorries":
2 | [{"proofState": 0,
3 | "pos": {"line": 1, "column": 29},
4 | "goal": "x : Unit\n⊢ Nat",
5 | "endPos": {"line": 1, "column": 34}}],
6 | "messages":
7 | [{"severity": "warning",
8 | "pos": {"line": 1, "column": 4},
9 | "endPos": {"line": 1, "column": 5},
10 | "data": "declaration uses 'sorry'"}],
11 | "env": 0}
12 |
13 | {"proofStatus": "Incomplete: open goals remain",
14 | "proofState": 1,
15 | "goals": ["x : Unit\n⊢ Int"]}
16 |
17 | {"proofStatus": "Completed", "proofState": 2, "goals": []}
18 |
19 |
--------------------------------------------------------------------------------
/test/readme.in:
--------------------------------------------------------------------------------
1 | {"cmd" : "def f (x : Unit) : Nat := by sorry"}
2 |
3 | {"tactic": "apply Int.natAbs", "proofState": 0}
4 |
5 | {"tactic": "exact -37", "proofState": 1}
6 |
--------------------------------------------------------------------------------
/test/root_goals.expected.out:
--------------------------------------------------------------------------------
1 | {"sorries":
2 | [{"proofState": 0,
3 | "pos": {"line": 1, "column": 17},
4 | "goal": "⊢ 1 = 1",
5 | "endPos": {"line": 1, "column": 17}},
6 | {"proofState": 1,
7 | "pos": {"line": 5, "column": 22},
8 | "goal": "⊢ 1 + 1 = 2",
9 | "endPos": {"line": 5, "column": 22}}],
10 | "messages":
11 | [{"severity": "info",
12 | "pos": {"line": 3, "column": 2},
13 | "endPos": {"line": 3, "column": 8},
14 | "data": "Try this: exact rfl"}],
15 | "env": 0}
16 |
17 | {"proofStatus": "Completed", "proofState": 2, "goals": []}
18 |
19 | {"proofStatus": "Completed", "proofState": 3, "goals": []}
20 |
21 |
--------------------------------------------------------------------------------
/test/root_goals.in:
--------------------------------------------------------------------------------
1 | {"cmd": "example : 1=1 := by\n skip\n exact?\n\ntheorem bb : 1+1=2 := by rfl", "rootGoals": true}
2 |
3 | {"proofState" : 0, "tactic": "rfl"}
4 |
5 | {"proofState" : 1, "tactic": "exact rfl"}
6 |
7 |
--------------------------------------------------------------------------------
/test/self_proof_apply_check.expected.out:
--------------------------------------------------------------------------------
1 | {"messages":
2 | [{"severity": "error",
3 | "pos": {"line": 1, "column": 8},
4 | "endPos": {"line": 1, "column": 10},
5 | "data":
6 | "fail to show termination for\n ex\nwith errors\nfailed to infer structural recursion:\nno parameters suitable for structural recursion\n\nwell-founded recursion cannot be used, 'ex' does not take any (non-fixed) arguments"}],
7 | "env": 0}
8 |
9 | {"sorries":
10 | [{"proofState": 0,
11 | "pos": {"line": 1, "column": 22},
12 | "goal": "⊢ False",
13 | "endPos": {"line": 1, "column": 27}}],
14 | "messages":
15 | [{"severity": "warning",
16 | "pos": {"line": 1, "column": 8},
17 | "endPos": {"line": 1, "column": 10},
18 | "data": "declaration uses 'sorry'"}],
19 | "env": 1}
20 |
21 | {"proofStatus":
22 | "Error: kernel type check failed: (kernel) declaration has free variables '[anonymous]'",
23 | "proofState": 1,
24 | "goals": []}
25 |
26 |
--------------------------------------------------------------------------------
/test/self_proof_apply_check.in:
--------------------------------------------------------------------------------
1 | {"cmd": "theorem ex : False := by apply ex"}
2 |
3 | {"cmd": "theorem ex : False := sorry"}
4 |
5 | {"proofState": 0, "tactic": "apply ex"}
6 |
--------------------------------------------------------------------------------
/test/self_proof_check.expected.out:
--------------------------------------------------------------------------------
1 | {"messages":
2 | [{"severity": "error",
3 | "pos": {"line": 1, "column": 25},
4 | "endPos": {"line": 1, "column": 31},
5 | "data":
6 | "`exact?` could not close the goal. Try `apply?` to see partial suggestions."}],
7 | "env": 0}
8 |
9 | {"sorries":
10 | [{"proofState": 0,
11 | "pos": {"line": 1, "column": 22},
12 | "goal": "⊢ False",
13 | "endPos": {"line": 1, "column": 27}}],
14 | "messages":
15 | [{"severity": "warning",
16 | "pos": {"line": 1, "column": 8},
17 | "endPos": {"line": 1, "column": 10},
18 | "data": "declaration uses 'sorry'"}],
19 | "env": 1}
20 |
21 | {"message":
22 | "Lean error:\n`exact?` could not close the goal. Try `apply?` to see partial suggestions."}
23 |
24 | {"sorries":
25 | [{"proofState": 1,
26 | "pos": {"line": 1, "column": 25},
27 | "goal": "⊢ False",
28 | "endPos": {"line": 1, "column": 30}}],
29 | "messages":
30 | [{"severity": "warning",
31 | "pos": {"line": 1, "column": 8},
32 | "endPos": {"line": 1, "column": 10},
33 | "data": "declaration uses 'sorry'"}],
34 | "env": 2}
35 |
36 | {"message":
37 | "Lean error:\n`exact?` could not close the goal. Try `apply?` to see partial suggestions."}
38 |
39 | {"tactics":
40 | [{"usedConstants":
41 | ["False",
42 | "sorryAx",
43 | "instOfNatNat",
44 | "Lean.Name.num",
45 | "Lean.Name.str",
46 | "Lean.Name.anonymous",
47 | "Nat",
48 | "Lean.Name",
49 | "OfNat.ofNat",
50 | "Bool.false"],
51 | "tactic": "sorry",
52 | "proofState": 3,
53 | "pos": {"line": 1, "column": 25},
54 | "goals": "⊢ False",
55 | "endPos": {"line": 1, "column": 30}}],
56 | "sorries":
57 | [{"proofState": 2,
58 | "pos": {"line": 1, "column": 25},
59 | "goal": "⊢ False",
60 | "endPos": {"line": 1, "column": 30}}],
61 | "messages":
62 | [{"severity": "warning",
63 | "pos": {"line": 1, "column": 8},
64 | "endPos": {"line": 1, "column": 10},
65 | "data": "declaration uses 'sorry'"}],
66 | "env": 3}
67 |
68 | {"message":
69 | "Lean error:\n`exact?` could not close the goal. Try `apply?` to see partial suggestions."}
70 |
71 | {"sorries":
72 | [{"proofState": 4,
73 | "pos": {"line": 1, "column": 25},
74 | "goal": "⊢ False",
75 | "endPos": {"line": 1, "column": 30}},
76 | {"proofState": 5,
77 | "pos": {"line": 1, "column": 22},
78 | "goal": "⊢ False",
79 | "endPos": {"line": 1, "column": 22}}],
80 | "messages":
81 | [{"severity": "warning",
82 | "pos": {"line": 1, "column": 8},
83 | "endPos": {"line": 1, "column": 10},
84 | "data": "declaration uses 'sorry'"}],
85 | "env": 4}
86 |
87 | {"message":
88 | "Lean error:\n`exact?` could not close the goal. Try `apply?` to see partial suggestions."}
89 |
90 |
--------------------------------------------------------------------------------
/test/self_proof_check.in:
--------------------------------------------------------------------------------
1 | {"cmd": "theorem ex : False := by exact?"}
2 |
3 | {"cmd": "theorem ex : False := sorry"}
4 |
5 | {"proofState": 0, "tactic": "exact?"}
6 |
7 | {"cmd": "theorem ex : False := by sorry"}
8 |
9 | {"proofState": 1, "tactic": "exact?"}
10 |
11 | {"cmd": "theorem ex : False := by sorry", "allTactics": true}
12 |
13 | {"proofState": 3, "tactic": "exact?"}
14 |
15 | {"cmd": "theorem ex : False := by sorry", "rootGoals": true}
16 |
17 | {"proofState": 5, "tactic": "exact?"}
18 |
19 |
--------------------------------------------------------------------------------
/test/self_proof_exact_check.expected.out:
--------------------------------------------------------------------------------
1 | {"messages":
2 | [{"severity": "error",
3 | "pos": {"line": 1, "column": 8},
4 | "endPos": {"line": 1, "column": 10},
5 | "data":
6 | "fail to show termination for\n ex\nwith errors\nfailed to infer structural recursion:\nno parameters suitable for structural recursion\n\nwell-founded recursion cannot be used, 'ex' does not take any (non-fixed) arguments"}],
7 | "env": 0}
8 |
9 | {"sorries":
10 | [{"proofState": 0,
11 | "pos": {"line": 1, "column": 22},
12 | "goal": "⊢ False",
13 | "endPos": {"line": 1, "column": 27}}],
14 | "messages":
15 | [{"severity": "warning",
16 | "pos": {"line": 1, "column": 8},
17 | "endPos": {"line": 1, "column": 10},
18 | "data": "declaration uses 'sorry'"}],
19 | "env": 1}
20 |
21 | {"proofStatus":
22 | "Error: kernel type check failed: (kernel) declaration has free variables '[anonymous]'",
23 | "proofState": 1,
24 | "goals": []}
25 |
26 |
--------------------------------------------------------------------------------
/test/self_proof_exact_check.in:
--------------------------------------------------------------------------------
1 | {"cmd": "theorem ex : False := by exact ex"}
2 |
3 | {"cmd": "theorem ex : False := sorry"}
4 |
5 | {"proofState": 0, "tactic": "exact ex"}
6 |
--------------------------------------------------------------------------------
/test/self_proof_rw.expected.out:
--------------------------------------------------------------------------------
1 | {"messages":
2 | [{"severity": "error",
3 | "pos": {"line": 1, "column": 8},
4 | "endPos": {"line": 1, "column": 24},
5 | "data":
6 | "fail to show termination for\n self_application\nwith errors\nfailed to infer structural recursion:\nno parameters suitable for structural recursion\n\nwell-founded recursion cannot be used, 'self_application' does not take any (non-fixed) arguments"}],
7 | "env": 0}
8 |
9 | {"sorries":
10 | [{"proofState": 0,
11 | "pos": {"line": 1, "column": 39},
12 | "goal": "⊢ 1 = 0",
13 | "endPos": {"line": 1, "column": 44}}],
14 | "messages":
15 | [{"severity": "warning",
16 | "pos": {"line": 1, "column": 8},
17 | "endPos": {"line": 1, "column": 24},
18 | "data": "declaration uses 'sorry'"}],
19 | "env": 1}
20 |
21 | {"proofStatus":
22 | "Error: kernel type check failed: (kernel) declaration has free variables '[anonymous]'",
23 | "proofState": 1,
24 | "goals": []}
25 |
26 |
--------------------------------------------------------------------------------
/test/self_proof_rw.in:
--------------------------------------------------------------------------------
1 | {"cmd": "theorem self_application : 1 = 0 := by rw [self_application]"}
2 |
3 | {"cmd": "theorem self_application : 1 = 0 := by sorry"}
4 |
5 | {"proofState": 0, "tactic": "rw [self_application]"}
6 |
--------------------------------------------------------------------------------
/test/sorry_hypotheses.expected.out:
--------------------------------------------------------------------------------
1 | {"sorries":
2 | [{"proofState": 0,
3 | "pos": {"line": 1, "column": 36},
4 | "goal": "x : Int\n⊢ x = x",
5 | "endPos": {"line": 1, "column": 41}}],
6 | "messages":
7 | [{"severity": "warning",
8 | "pos": {"line": 1, "column": 8},
9 | "endPos": {"line": 1, "column": 11},
10 | "data": "declaration uses 'sorry'"}],
11 | "env": 0}
12 |
13 | {"sorries": [{"proofState": 1, "goal": "x : Int\n⊢ x = 1"}],
14 | "proofStatus": "Incomplete: open goals remain",
15 | "proofState": 2,
16 | "goals": ["x : Int\nh : x = 1\n⊢ x = x"]}
17 |
18 |
--------------------------------------------------------------------------------
/test/sorry_hypotheses.in:
--------------------------------------------------------------------------------
1 | {"cmd": "theorem foo (x : Int) : x = x := by sorry"}
2 |
3 | {"proofState" : 0, "tactic": "have h : x = 1 := sorry"}
4 |
5 |
--------------------------------------------------------------------------------
/test/synthesize_placeholder.expected.out:
--------------------------------------------------------------------------------
1 | {"messages":
2 | [{"severity": "error",
3 | "pos": {"line": 1, "column": 15},
4 | "endPos": {"line": 1, "column": 16},
5 | "data": "don't know how to synthesize placeholder\ncontext:\n⊢ Nat"}],
6 | "env": 0}
7 |
8 |
--------------------------------------------------------------------------------
/test/synthesize_placeholder.in:
--------------------------------------------------------------------------------
1 | {"cmd": "def f : Nat := _"}
2 |
--------------------------------------------------------------------------------
/test/tactic_mode_sorry.expected.out:
--------------------------------------------------------------------------------
1 | {"sorries":
2 | [{"proofState": 0,
3 | "pos": {"line": 1, "column": 18},
4 | "goal": "⊢ Nat",
5 | "endPos": {"line": 1, "column": 23}}],
6 | "messages":
7 | [{"severity": "warning",
8 | "pos": {"line": 1, "column": 4},
9 | "endPos": {"line": 1, "column": 5},
10 | "data": "declaration uses 'sorry'"}],
11 | "env": 0}
12 |
13 | {"sorries": [{"proofState": 1, "goal": "⊢ Nat"}],
14 | "proofStatus": "Incomplete: contains sorry",
15 | "proofState": 2,
16 | "goals": []}
17 |
18 |
--------------------------------------------------------------------------------
/test/tactic_mode_sorry.in:
--------------------------------------------------------------------------------
1 | {"cmd": "def f : Nat := by sorry"}
2 |
3 | {"proofState": 0, "tactic": "sorry"}
4 |
--------------------------------------------------------------------------------
/test/tactic_sorry.expected.out:
--------------------------------------------------------------------------------
1 | {"sorries":
2 | [{"proofState": 0,
3 | "pos": {"line": 1, "column": 18},
4 | "goal": "⊢ Nat",
5 | "endPos": {"line": 1, "column": 23}}],
6 | "messages":
7 | [{"severity": "warning",
8 | "pos": {"line": 1, "column": 4},
9 | "endPos": {"line": 1, "column": 5},
10 | "data": "declaration uses 'sorry'"}],
11 | "env": 0}
12 |
13 |
--------------------------------------------------------------------------------
/test/tactic_sorry.in:
--------------------------------------------------------------------------------
1 | {"cmd": "def f : Nat := by sorry"}
2 |
--------------------------------------------------------------------------------
/test/term_sorry.expected.out:
--------------------------------------------------------------------------------
1 | {"sorries":
2 | [{"proofState": 0,
3 | "pos": {"line": 1, "column": 15},
4 | "goal": "⊢ Nat",
5 | "endPos": {"line": 1, "column": 20}}],
6 | "messages":
7 | [{"severity": "warning",
8 | "pos": {"line": 1, "column": 4},
9 | "endPos": {"line": 1, "column": 5},
10 | "data": "declaration uses 'sorry'"}],
11 | "env": 0}
12 |
13 |
--------------------------------------------------------------------------------
/test/term_sorry.in:
--------------------------------------------------------------------------------
1 | {"cmd": "def f : Nat := sorry"}
2 |
--------------------------------------------------------------------------------
/test/trace_simp.expected.out:
--------------------------------------------------------------------------------
1 | {"env": 0}
2 |
3 | {"env": 1}
4 |
5 | {"env": 2}
6 |
7 | {"messages":
8 | [{"severity": "info",
9 | "pos": {"line": 1, "column": 23},
10 | "endPos": {"line": 1, "column": 27},
11 | "data":
12 | "[Meta.Tactic.simp.rewrite] f_def:1000:\n f\n ==>\n 37\n[Meta.Tactic.simp.rewrite] eq_self:1000:\n 37 = 37\n ==>\n True"}],
13 | "env": 3}
14 |
15 | {"sorries":
16 | [{"proofState": 0,
17 | "pos": {"line": 1, "column": 23},
18 | "goal": "⊢ f = 37",
19 | "endPos": {"line": 1, "column": 28}}],
20 | "messages":
21 | [{"severity": "warning",
22 | "pos": {"line": 1, "column": 0},
23 | "endPos": {"line": 1, "column": 7},
24 | "data": "declaration uses 'sorry'"}],
25 | "env": 4}
26 |
27 | {"traces": ["37"],
28 | "proofStatus": "Incomplete: open goals remain",
29 | "proofState": 1,
30 | "goals": ["⊢ f = 37"]}
31 |
32 | {"traces":
33 | ["[Meta.Tactic.simp.rewrite] f_def:1000:\n f\n ==>\n 37",
34 | "[Meta.Tactic.simp.rewrite] eq_self:1000:\n 37 = 37\n ==>\n True"],
35 | "proofStatus": "Completed",
36 | "proofState": 2,
37 | "goals": []}
38 |
39 | {"messages":
40 | [{"severity": "info",
41 | "pos": {"line": 2, "column": 2},
42 | "endPos": {"line": 2, "column": 7},
43 | "data": "!"}],
44 | "env": 5}
45 |
46 |
--------------------------------------------------------------------------------
/test/trace_simp.in:
--------------------------------------------------------------------------------
1 | {"cmd": "def f := 37"}
2 |
3 | {"cmd": "@[simp] theorem f_def : f = 37 := by rfl", "env": 0}
4 |
5 | {"cmd": "set_option trace.Meta.Tactic.simp true", "env": 1}
6 |
7 | {"cmd": "example : f = 37 := by simp", "env": 2}
8 |
9 | {"cmd": "example : f = 37 := by sorry", "env": 2}
10 |
11 | {"tactic": "trace \"37\"", "proofState": 0}
12 |
13 | {"tactic": "simp", "proofState": 0}
14 |
15 | {"cmd": "example : True := by\n trace \"!\"\n trivial"}
16 |
--------------------------------------------------------------------------------
/test/unfinished_tactic_block.expected.out:
--------------------------------------------------------------------------------
1 | {"messages":
2 | [{"severity": "error",
3 | "pos": {"line": 1, "column": 17},
4 | "endPos": null,
5 | "data": "unexpected end of input; expected '{'"},
6 | {"severity": "error",
7 | "pos": {"line": 1, "column": 15},
8 | "endPos": {"line": 1, "column": 17},
9 | "data": "unsolved goals\n⊢ Nat"}],
10 | "env": 0}
11 |
12 |
--------------------------------------------------------------------------------
/test/unfinished_tactic_block.in:
--------------------------------------------------------------------------------
1 | {"cmd": "def f : Nat := by"}
2 |
--------------------------------------------------------------------------------
/test/unknown_environment.expected.out:
--------------------------------------------------------------------------------
1 | {"message": "Unknown environment."}
2 |
3 |
--------------------------------------------------------------------------------
/test/unknown_environment.in:
--------------------------------------------------------------------------------
1 | {"cmd": "def f : Nat := by sorry", "env": 5}
2 |
--------------------------------------------------------------------------------
/test/unknown_proof_state.expected.out:
--------------------------------------------------------------------------------
1 | {"sorries":
2 | [{"proofState": 0,
3 | "pos": {"line": 1, "column": 18},
4 | "goal": "⊢ Nat",
5 | "endPos": {"line": 1, "column": 23}}],
6 | "messages":
7 | [{"severity": "warning",
8 | "pos": {"line": 1, "column": 4},
9 | "endPos": {"line": 1, "column": 5},
10 | "data": "declaration uses 'sorry'"}],
11 | "env": 0}
12 |
13 | {"message": "Unknown proof state."}
14 |
15 |
--------------------------------------------------------------------------------
/test/unknown_proof_state.in:
--------------------------------------------------------------------------------
1 | {"cmd" : "def f : Nat := by sorry"}
2 |
3 | {"tactic": "exact 42", "proofState": 1}
4 |
--------------------------------------------------------------------------------
/test/unknown_tactic.expected.out:
--------------------------------------------------------------------------------
1 | {"sorries":
2 | [{"proofState": 0,
3 | "pos": {"line": 1, "column": 18},
4 | "goal": "⊢ Nat",
5 | "endPos": {"line": 1, "column": 23}}],
6 | "messages":
7 | [{"severity": "warning",
8 | "pos": {"line": 1, "column": 4},
9 | "endPos": {"line": 1, "column": 5},
10 | "data": "declaration uses 'sorry'"}],
11 | "env": 0}
12 |
13 | {"message": "Lean error:\n:1:1: unknown tactic"}
14 |
15 |
--------------------------------------------------------------------------------
/test/unknown_tactic.in:
--------------------------------------------------------------------------------
1 | {"cmd" : "def f : Nat := by sorry"}
2 |
3 | {"tactic": "exat 42", "proofState": 0}
4 |
--------------------------------------------------------------------------------
/test/variables.expected.out:
--------------------------------------------------------------------------------
1 | {"env": 0}
2 |
3 | {"messages":
4 | [{"severity": "warning",
5 | "pos": {"line": 1, "column": 12},
6 | "endPos": {"line": 1, "column": 13},
7 | "data":
8 | "unused variable `y`\nnote: this linter can be disabled with `set_option linter.unusedVariables false`"}],
9 | "env": 1}
10 |
11 | {"sorries":
12 | [{"proofState": 0,
13 | "pos": {"line": 3, "column": 2},
14 | "goal":
15 | "x y : Nat\nf : Nat → Nat\nh0 : f 5 = 3\nh1 : f (4 * x * y) = 2 * y * (f (x + y) + f (x - y))\n⊢ ∃ k, f 2015 = k",
16 | "endPos": {"line": 3, "column": 7}}],
17 | "messages":
18 | [{"severity": "warning",
19 | "pos": {"line": 1, "column": 8},
20 | "endPos": {"line": 1, "column": 15},
21 | "data": "declaration uses 'sorry'"}],
22 | "env": 2}
23 |
24 |
--------------------------------------------------------------------------------
/test/variables.in:
--------------------------------------------------------------------------------
1 | {"cmd": "variable (x y : Nat)"}
2 |
3 | {"cmd": "variable (f : Nat → Nat)", "env": 0}
4 |
5 | {"cmd": "theorem problem (h0 : f 5 = 3) (h1 : f (4 * x * y) = 2 * y * (f (x + y) + f (x - y))) :\n ∃ (k : Nat), f 2015 = k := by\n sorry", "env": 1}
6 |
--------------------------------------------------------------------------------