Merge branch 'dev' into bug/expr-proj-to-app-panic

This commit is contained in:
Leni Aniva 2025-01-24 15:03:13 -08:00
commit 6f792b0657
Signed by: aniva
GPG Key ID: 4D9B1C8D10EA4C50
8 changed files with 92 additions and 20 deletions

View File

@ -43,6 +43,22 @@ def toFilteredSymbol (n: Lean.Name) (info: Lean.ConstantInfo): Option String :=
if isNameInternal n || info.isUnsafe if isNameInternal n || info.isUnsafe
then Option.none then Option.none
else Option.some <| toCompactSymbolName n info else Option.some <| toCompactSymbolName n info
def describe (_: Protocol.EnvDescribe): CoreM Protocol.EnvDescribeResult := do
let env ← Lean.MonadEnv.getEnv
return {
imports := env.header.imports.map toString,
modules := env.header.moduleNames.map (·.toString),
}
def moduleRead (args: Protocol.EnvModuleRead): CoreM (Protocol.CR Protocol.EnvModuleReadResult) := do
let env ← Lean.MonadEnv.getEnv
let .some i := env.header.moduleNames.findIdx? (· == args.module.toName) |
return .error $ Protocol.errorIndex s!"Module not found {args.module}"
let data := env.header.moduleData[i]!
return .ok {
imports := data.imports.map toString,
constNames := data.constNames.map (·.toString),
extraConstNames := data.extraConstNames.map (·.toString),
}
def catalog (_: Protocol.EnvCatalog): CoreM Protocol.EnvCatalogResult := do def catalog (_: Protocol.EnvCatalog): CoreM Protocol.EnvCatalogResult := do
let env ← Lean.MonadEnv.getEnv let env ← Lean.MonadEnv.getEnv
let names := env.constants.fold (init := #[]) (λ acc name info => let names := env.constants.fold (init := #[]) (λ acc name info =>

View File

@ -68,7 +68,8 @@ private partial def translateExpr (srcExpr: Expr) : MetaTranslateM Expr := do
match e with match e with
| .fvar fvarId => | .fvar fvarId =>
let .some fvarId' := state.fvarMap[fvarId]? | panic! s!"FVar id not registered: {fvarId.name}" let .some fvarId' := state.fvarMap[fvarId]? | panic! s!"FVar id not registered: {fvarId.name}"
assert! (← getLCtx).contains fvarId' -- Delegating this to `Meta.check` later
--assert! (← getLCtx).contains fvarId'
return .done $ .fvar fvarId' return .done $ .fvar fvarId'
| .mvar mvarId => do | .mvar mvarId => do
-- Must not be assigned -- Must not be assigned

View File

@ -207,6 +207,14 @@ protected def GoalState.tryNoConfuse (state: GoalState) (goal: MVarId) (eq: Stri
| .ok syn => pure syn | .ok syn => pure syn
| .error error => return .parseError error | .error error => return .parseError error
state.tryTacticM goal (tacticM := Tactic.evalNoConfuse eq) state.tryTacticM goal (tacticM := Tactic.evalNoConfuse eq)
@[export pantograph_goal_try_draft_m]
protected def GoalState.tryDraft (state: GoalState) (goal: MVarId) (expr: String): CoreM TacticResult := do
let expr ← match (← parseTermM expr) with
| .ok syn => pure syn
| .error error => return .parseError error
runTermElabM do
state.restoreElabM
state.tryTacticM goal (Tactic.evalDraft expr)
@[export pantograph_goal_let_m] @[export pantograph_goal_let_m]
def goalLet (state: GoalState) (goal: MVarId) (binderName: String) (type: String): CoreM TacticResult := def goalLet (state: GoalState) (goal: MVarId) (binderName: String) (type: String): CoreM TacticResult :=
runTermElabM <| state.tryLet goal binderName type runTermElabM <| state.tryLet goal binderName type

View File

@ -112,6 +112,24 @@ structure ExprEchoResult where
type: Expression type: Expression
deriving Lean.ToJson deriving Lean.ToJson
-- Describe the current state of the environment
structure EnvDescribe where
deriving Lean.FromJson
structure EnvDescribeResult where
imports : Array String
modules : Array String
deriving Lean.ToJson
-- Describe a module
structure EnvModuleRead where
module : String
deriving Lean.FromJson
structure EnvModuleReadResult where
imports: Array String
constNames: Array String
extraConstNames: Array String
deriving Lean.ToJson
-- Print all symbols in environment -- Print all symbols in environment
structure EnvCatalog where structure EnvCatalog where
deriving Lean.FromJson deriving Lean.FromJson
@ -237,6 +255,7 @@ structure GoalTactic where
calc?: Option String := .none calc?: Option String := .none
-- true to enter `conv`, `false` to exit. In case of exit the `goalId` is ignored. -- true to enter `conv`, `false` to exit. In case of exit the `goalId` is ignored.
conv?: Option Bool := .none conv?: Option Bool := .none
draft?: Option String := .none
-- In case of the `have` tactic, the new free variable name is provided here -- In case of the `have` tactic, the new free variable name is provided here
binderName?: Option String := .none binderName?: Option String := .none

View File

@ -59,6 +59,12 @@ and then add the new constants.
def environmentPickle (env : Environment) (path : System.FilePath) : IO Unit := def environmentPickle (env : Environment) (path : System.FilePath) : IO Unit :=
Pantograph.pickle path (env.header.imports, env.constants.map₂) Pantograph.pickle path (env.header.imports, env.constants.map₂)
def resurrectEnvironment
(imports : Array Import)
(map₂ : PHashMap Name ConstantInfo)
: IO Environment := do
let env ← importModules imports {} 0
env.replay (Std.HashMap.ofList map₂.toList)
/-- /--
Unpickle an `Environment` from disk. Unpickle an `Environment` from disk.
@ -68,8 +74,7 @@ and then replace the new constants.
@[export pantograph_env_unpickle_m] @[export pantograph_env_unpickle_m]
def environmentUnpickle (path : System.FilePath) : IO (Environment × CompactedRegion) := unsafe do def environmentUnpickle (path : System.FilePath) : IO (Environment × CompactedRegion) := unsafe do
let ((imports, map₂), region) ← Pantograph.unpickle (Array Import × PHashMap Name ConstantInfo) path let ((imports, map₂), region) ← Pantograph.unpickle (Array Import × PHashMap Name ConstantInfo) path
let env ← importModules imports {} 0 return (← resurrectEnvironment imports map₂, region)
return (← env.replay (Std.HashMap.ofList map₂.toList), region)
open Lean.Core in open Lean.Core in
@ -88,7 +93,9 @@ def goalStatePickle (goalState : GoalState) (path : System.FilePath) : IO Unit :
savedState := { savedState := {
term := { term := {
meta := { meta := {
core, core := {
env, nextMacroScope, ngen, ..
},
meta, meta,
} }
«elab», «elab»,
@ -100,9 +107,10 @@ def goalStatePickle (goalState : GoalState) (path : System.FilePath) : IO Unit :
convMVar?, convMVar?,
calcPrevRhs?, calcPrevRhs?,
} := goalState } := goalState
--let env := core.env
Pantograph.pickle path ( Pantograph.pickle path (
({ core with } : CompactCoreState), env.constants.map₂,
({ nextMacroScope, ngen } : CompactCoreState),
meta, meta,
«elab», «elab»,
tactic, tactic,
@ -117,6 +125,8 @@ def goalStatePickle (goalState : GoalState) (path : System.FilePath) : IO Unit :
def goalStateUnpickle (path : System.FilePath) (env : Environment) def goalStateUnpickle (path : System.FilePath) (env : Environment)
: IO (GoalState × CompactedRegion) := unsafe do : IO (GoalState × CompactedRegion) := unsafe do
let (( let ((
map₂,
compactCore, compactCore,
meta, meta,
«elab», «elab»,
@ -127,6 +137,8 @@ def goalStateUnpickle (path : System.FilePath) (env : Environment)
convMVar?, convMVar?,
calcPrevRhs?, calcPrevRhs?,
), region) ← Pantograph.unpickle ( ), region) ← Pantograph.unpickle (
PHashMap Name ConstantInfo ×
CompactCoreState × CompactCoreState ×
Meta.State × Meta.State ×
Elab.Term.State × Elab.Term.State ×
@ -137,6 +149,7 @@ def goalStateUnpickle (path : System.FilePath) (env : Environment)
Option (MVarId × MVarId × List MVarId) × Option (MVarId × MVarId × List MVarId) ×
Option (MVarId × Expr) Option (MVarId × Expr)
) path ) path
let env ← env.replay (Std.HashMap.ofList map₂.toList)
let goalState := { let goalState := {
savedState := { savedState := {
term := { term := {

View File

@ -48,6 +48,8 @@ def execute (command: Protocol.Command): MainM Lean.Json := do
| "reset" => run reset | "reset" => run reset
| "stat" => run stat | "stat" => run stat
| "expr.echo" => run expr_echo | "expr.echo" => run expr_echo
| "env.describe" => run env_describe
| "env.module_read" => run env_module_read
| "env.catalog" => run env_catalog | "env.catalog" => run env_catalog
| "env.inspect" => run env_inspect | "env.inspect" => run env_inspect
| "env.add" => run env_add | "env.add" => run env_add
@ -85,6 +87,11 @@ def execute (command: Protocol.Command): MainM Lean.Json := do
let state ← get let state ← get
let nGoals := state.goalStates.size let nGoals := state.goalStates.size
return .ok { nGoals } return .ok { nGoals }
env_describe (args: Protocol.EnvDescribe): MainM (CR Protocol.EnvDescribeResult) := do
let result ← Environment.describe args
return .ok result
env_module_read (args: Protocol.EnvModuleRead): MainM (CR Protocol.EnvModuleReadResult) := do
Environment.moduleRead args
env_catalog (args: Protocol.EnvCatalog): MainM (CR Protocol.EnvCatalogResult) := do env_catalog (args: Protocol.EnvCatalog): MainM (CR Protocol.EnvCatalogResult) := do
let result ← Environment.catalog args let result ← Environment.catalog args
return .ok result return .ok result
@ -145,24 +152,27 @@ def execute (command: Protocol.Command): MainM Lean.Json := do
let .some goal := goalState.goals.get? args.goalId | let .some goal := goalState.goals.get? args.goalId |
return .error $ errorIndex s!"Invalid goal index {args.goalId}" return .error $ errorIndex s!"Invalid goal index {args.goalId}"
let nextGoalState?: Except _ TacticResult ← runTermElabInMainM do let nextGoalState?: Except _ TacticResult ← runTermElabInMainM do
match args.tactic?, args.expr?, args.have?, args.let?, args.calc?, args.conv? with -- NOTE: Should probably use a macro to handle this...
| .some tactic, .none, .none, .none, .none, .none => do match args.tactic?, args.expr?, args.have?, args.let?, args.calc?, args.conv?, args.draft? with
| .some tactic, .none, .none, .none, .none, .none, .none => do
pure <| Except.ok <| ← goalState.tryTactic goal tactic pure <| Except.ok <| ← goalState.tryTactic goal tactic
| .none, .some expr, .none, .none, .none, .none => do | .none, .some expr, .none, .none, .none, .none, .none => do
pure <| Except.ok <| ← goalState.tryAssign goal expr pure <| Except.ok <| ← goalState.tryAssign goal expr
| .none, .none, .some type, .none, .none, .none => do | .none, .none, .some type, .none, .none, .none, .none => do
let binderName := args.binderName?.getD "" let binderName := args.binderName?.getD ""
pure <| Except.ok <| ← goalState.tryHave goal binderName type pure <| Except.ok <| ← goalState.tryHave goal binderName type
| .none, .none, .none, .some type, .none, .none => do | .none, .none, .none, .some type, .none, .none, .none => do
let binderName := args.binderName?.getD "" let binderName := args.binderName?.getD ""
pure <| Except.ok <| ← goalState.tryLet goal binderName type pure <| Except.ok <| ← goalState.tryLet goal binderName type
| .none, .none, .none, .none, .some pred, .none => do | .none, .none, .none, .none, .some pred, .none, .none => do
pure <| Except.ok <| ← goalState.tryCalc goal pred pure <| Except.ok <| ← goalState.tryCalc goal pred
| .none, .none, .none, .none, .none, .some true => do | .none, .none, .none, .none, .none, .some true, .none => do
pure <| Except.ok <| ← goalState.conv goal pure <| Except.ok <| ← goalState.conv goal
| .none, .none, .none, .none, .none, .some false => do | .none, .none, .none, .none, .none, .some false, .none => do
pure <| Except.ok <| ← goalState.convExit pure <| Except.ok <| ← goalState.convExit
| _, _, _, _, _, _ => | .none, .none, .none, .none, .none, .none, .some draft => do
pure <| Except.ok <| ← goalState.tryDraft goal draft
| _, _, _, _, _, _, _ =>
let error := errorI "arguments" "Exactly one of {tactic, expr, have, calc, conv} must be supplied" let error := errorI "arguments" "Exactly one of {tactic, expr, have, calc, conv} must be supplied"
pure $ Except.error $ error pure $ Except.error $ error
match nextGoalState? with match nextGoalState? with

View File

@ -110,6 +110,9 @@ def test_symbol_location : TestT IO Unit := do
checkTrue "file" result.sourceUri?.isNone checkTrue "file" result.sourceUri?.isNone
checkEq "pos" (result.sourceStart?.map (·.column)) <| .some 0 checkEq "pos" (result.sourceStart?.map (·.column)) <| .some 0
checkEq "pos" (result.sourceEnd?.map (·.column)) <| .some 88 checkEq "pos" (result.sourceEnd?.map (·.column)) <| .some 88
let .ok { imports, constNames, extraConstNames } ← Environment.moduleRead ⟨"Init.Data.Nat.Basic"⟩ | fail "Module read failed"
checkEq "imports" imports #["Init.SimpLemmas", "Init.Data.NeZero"]
checkTrue "constNames" $ constNames.contains "Nat.succ_add"
def suite: List (String × IO LSpec.TestSeq) := def suite: List (String × IO LSpec.TestSeq) :=
[ [

View File

@ -33,6 +33,7 @@ See `Pantograph/Protocol.lean` for a description of the parameters and return va
to the previous `rhs`. to the previous `rhs`.
- `{ "conv": <bool> }`: Enter or exit conversion tactic mode. In the case of - `{ "conv": <bool> }`: Enter or exit conversion tactic mode. In the case of
exit, the goal id is ignored. exit, the goal id is ignored.
- `{ "draft": <expr> }`: Draft an expression with `sorry`s, turning them into goals. Coupling is not allowed.
* `goal.continue {"stateId": <id>, ["branch": <id>], ["goals": <names>]}`: * `goal.continue {"stateId": <id>, ["branch": <id>], ["goals": <names>]}`:
Execute continuation/resumption Execute continuation/resumption
- `{ "branch": <id> }`: Continue on branch state. The current state must have no goals. - `{ "branch": <id> }`: Continue on branch state. The current state must have no goals.
@ -44,11 +45,12 @@ See `Pantograph/Protocol.lean` for a description of the parameters and return va
state. The user is responsible to ensure the sender/receiver instances share state. The user is responsible to ensure the sender/receiver instances share
the same environment. the same environment.
* `frontend.process { ["fileName": <fileName>,] ["file": <str>], invocations: * `frontend.process { ["fileName": <fileName>,] ["file": <str>], invocations:
<bool>, sorrys: <bool>, newConstants: <bool> }`: Executes the Lean frontend on <bool>, sorrys: <bool>, typeErrorsAsGoals: <bool>, newConstants: <bool> }`:
a file, collecting the tactic invocations (`"invocations": true`), the Executes the Lean frontend on a file, collecting the tactic invocations
sorrys and type errors into goal states (`"sorrys": true`), and new constants (`"invocations": true`), the sorrys and type errors into goal states
(`"newConstants": true`). In the case of `sorrys`, this command additionally (`"sorrys": true`), and new constants (`"newConstants": true`). In the case of
outputs the position of each captured `sorry`. `sorrys`, this command additionally outputs the position of each captured
`sorry`.
## Errors ## Errors