refactor: Move all frontend functions to `Frontend`
This commit is contained in:
parent
860344f9c5
commit
8e3241c02a
|
@ -7,7 +7,7 @@ import Pantograph
|
||||||
import Repl
|
import Repl
|
||||||
|
|
||||||
-- Main IO functions
|
-- Main IO functions
|
||||||
open Pantograph
|
open Pantograph.Repl
|
||||||
|
|
||||||
/-- Parse a command either in `{ "cmd": ..., "payload": ... }` form or `cmd { ... }` form. -/
|
/-- Parse a command either in `{ "cmd": ..., "payload": ... }` form or `cmd { ... }` form. -/
|
||||||
def parseCommand (s: String): Except String Protocol.Command := do
|
def parseCommand (s: String): Except String Protocol.Command := do
|
||||||
|
|
|
@ -1,6 +1,6 @@
|
||||||
import Pantograph.Compile
|
|
||||||
import Pantograph.Condensed
|
import Pantograph.Condensed
|
||||||
import Pantograph.Environment
|
import Pantograph.Environment
|
||||||
|
import Pantograph.Frontend
|
||||||
import Pantograph.Goal
|
import Pantograph.Goal
|
||||||
import Pantograph.Library
|
import Pantograph.Library
|
||||||
import Pantograph.Protocol
|
import Pantograph.Protocol
|
||||||
|
|
|
@ -1,25 +0,0 @@
|
||||||
/- Adapted from lean-training-data by semorrison -/
|
|
||||||
import Pantograph.Protocol
|
|
||||||
import Pantograph.Compile.Frontend
|
|
||||||
import Pantograph.Compile.Elab
|
|
||||||
import Pantograph.Compile.Parse
|
|
||||||
|
|
||||||
open Lean
|
|
||||||
|
|
||||||
namespace Pantograph.Compile
|
|
||||||
|
|
||||||
def collectTacticsFromCompilation (steps : List CompilationStep) : IO (List Protocol.InvokedTactic) := do
|
|
||||||
let infoTrees := steps.bind (·.trees)
|
|
||||||
let tacticInfoTrees := infoTrees.bind λ tree => tree.filter λ
|
|
||||||
| info@(.ofTacticInfo _) => info.isOriginal
|
|
||||||
| _ => false
|
|
||||||
let tactics := tacticInfoTrees.bind collectTactics
|
|
||||||
tactics.mapM λ invocation => do
|
|
||||||
let goalBefore := (Format.joinSep (← invocation.goalState) "\n").pretty
|
|
||||||
let goalAfter := (Format.joinSep (← invocation.goalStateAfter) "\n").pretty
|
|
||||||
let tactic ← invocation.ctx.runMetaM {} do
|
|
||||||
let t ← Lean.PrettyPrinter.ppTactic ⟨invocation.info.stx⟩
|
|
||||||
return t.pretty
|
|
||||||
return { goalBefore, goalAfter, tactic }
|
|
||||||
|
|
||||||
end Pantograph.Compile
|
|
|
@ -1,14 +0,0 @@
|
||||||
import Lean
|
|
||||||
|
|
||||||
open Lean
|
|
||||||
|
|
||||||
namespace Pantograph.Compile
|
|
||||||
|
|
||||||
def parseTermM [Monad m] [MonadEnv m] (s: String): m (Except String Syntax) := do
|
|
||||||
return Parser.runParserCategory
|
|
||||||
(env := ← MonadEnv.getEnv)
|
|
||||||
(catName := `term)
|
|
||||||
(input := s)
|
|
||||||
(fileName := "<stdin>")
|
|
||||||
|
|
||||||
end Pantograph.Compile
|
|
|
@ -2,7 +2,6 @@
|
||||||
import Lean
|
import Lean
|
||||||
import Pantograph.Goal
|
import Pantograph.Goal
|
||||||
import Pantograph.Expr
|
import Pantograph.Expr
|
||||||
import Pantograph.Protocol
|
|
||||||
|
|
||||||
open Lean
|
open Lean
|
||||||
|
|
||||||
|
|
|
@ -0,0 +1,4 @@
|
||||||
|
/- Adapted from lean-training-data by semorrison -/
|
||||||
|
import Pantograph.Protocol
|
||||||
|
import Pantograph.Frontend.Basic
|
||||||
|
import Pantograph.Frontend.Elab
|
|
@ -27,8 +27,7 @@ protected def drop [Inhabited α] (t : PersistentArray α) (n : Nat) : List α :
|
||||||
end Lean.PersistentArray
|
end Lean.PersistentArray
|
||||||
|
|
||||||
|
|
||||||
namespace Pantograph.Compile
|
namespace Pantograph.Frontend
|
||||||
|
|
||||||
|
|
||||||
abbrev FrontendM := Elab.Frontend.FrontendM
|
abbrev FrontendM := Elab.Frontend.FrontendM
|
||||||
|
|
||||||
|
@ -47,6 +46,7 @@ structure CompilationStep where
|
||||||
Process one command, returning a `CompilationStep` and
|
Process one command, returning a `CompilationStep` and
|
||||||
`done : Bool`, indicating whether this was the last command.
|
`done : Bool`, indicating whether this was the last command.
|
||||||
-/
|
-/
|
||||||
|
@[export pantograph_frontend_process_one_command_m]
|
||||||
def processOneCommand: FrontendM (CompilationStep × Bool) := do
|
def processOneCommand: FrontendM (CompilationStep × Bool) := do
|
||||||
let s := (← get).commandState
|
let s := (← get).commandState
|
||||||
let before := s.env
|
let before := s.env
|
||||||
|
@ -67,26 +67,38 @@ partial def collectCompilationSteps : FrontendM (List CompilationStep) := do
|
||||||
else
|
else
|
||||||
return cmd :: (← collectCompilationSteps)
|
return cmd :: (← collectCompilationSteps)
|
||||||
|
|
||||||
|
|
||||||
def findSourcePath (module : Name) : IO System.FilePath := do
|
def findSourcePath (module : Name) : IO System.FilePath := do
|
||||||
return System.FilePath.mk ((← findOLean module).toString.replace ".lake/build/lib/" "") |>.withExtension "lean"
|
return System.FilePath.mk ((← findOLean module).toString.replace ".lake/build/lib/" "") |>.withExtension "lean"
|
||||||
|
|
||||||
def runFrontendMInFile { α } (module : Name) (opts : Options := {}) (m : FrontendM α): IO α := unsafe do
|
@[export pantograph_create_frontend_context_state_from_file_m]
|
||||||
|
unsafe def createFrontendContextStateFromFile (module : Name) (opts : Options := {})
|
||||||
|
: IO (Elab.Frontend.Context × Elab.Frontend.State) := do
|
||||||
let file ← IO.FS.readFile (← findSourcePath module)
|
let file ← IO.FS.readFile (← findSourcePath module)
|
||||||
let inputCtx := Parser.mkInputContext file module.toString
|
let inputCtx := Parser.mkInputContext file module.toString
|
||||||
|
|
||||||
let (header, parserState, messages) ← Parser.parseHeader inputCtx
|
let (header, parserState, messages) ← Parser.parseHeader inputCtx
|
||||||
let (env, messages) ← Elab.processHeader header opts messages inputCtx
|
let (env, messages) ← Elab.processHeader header opts messages inputCtx
|
||||||
let commandState := Elab.Command.mkState env messages opts
|
let commandState := Elab.Command.mkState env messages opts
|
||||||
m.run { inputCtx }
|
let context: Elab.Frontend.Context := { inputCtx }
|
||||||
|>.run' {
|
let state: Elab.Frontend.State := {
|
||||||
commandState := { commandState with infoState.enabled := true },
|
commandState := { commandState with infoState.enabled := true },
|
||||||
parserState,
|
parserState,
|
||||||
cmdPos := parserState.pos
|
cmdPos := parserState.pos
|
||||||
}
|
}
|
||||||
|
return (context, state)
|
||||||
|
|
||||||
def processSource (module : Name) (opts : Options := {}) : IO (List CompilationStep) :=
|
partial def mapCompilationSteps { α } (f: CompilationStep → IO α) : FrontendM (List α) := do
|
||||||
runFrontendMInFile module opts collectCompilationSteps
|
let (cmd, done) ← processOneCommand
|
||||||
|
let result ← f cmd
|
||||||
|
if done then
|
||||||
|
return [result]
|
||||||
|
else
|
||||||
|
return result :: (← mapCompilationSteps f)
|
||||||
|
|
||||||
|
def runFrontendMInFile { α } (module : Name) (opts : Options := {}) (m : FrontendM α): IO α := unsafe do
|
||||||
|
let (context, state) ← createFrontendContextStateFromFile module opts
|
||||||
|
m.run context |>.run' state
|
||||||
|
|
||||||
|
|
||||||
end Pantograph.Compile
|
|
||||||
|
end Pantograph.Frontend
|
|
@ -3,7 +3,8 @@ import Lean.Elab.Import
|
||||||
import Lean.Elab.Command
|
import Lean.Elab.Command
|
||||||
import Lean.Elab.InfoTree
|
import Lean.Elab.InfoTree
|
||||||
|
|
||||||
import Pantograph.Compile.Frontend
|
import Pantograph.Protocol
|
||||||
|
import Pantograph.Frontend.Basic
|
||||||
|
|
||||||
open Lean
|
open Lean
|
||||||
|
|
||||||
|
@ -75,7 +76,7 @@ partial def filter (p : Info → Bool) (m : MVarId → Bool := fun _ => false) :
|
||||||
end Lean.Elab.InfoTree
|
end Lean.Elab.InfoTree
|
||||||
|
|
||||||
|
|
||||||
namespace Pantograph.Compile
|
namespace Pantograph.Frontend
|
||||||
|
|
||||||
-- Info tree filtering functions
|
-- Info tree filtering functions
|
||||||
|
|
||||||
|
@ -142,5 +143,18 @@ def collectTacticNodes (t : Elab.InfoTree) : List TacticInvocation :=
|
||||||
def collectTactics (t : Elab.InfoTree) : List TacticInvocation :=
|
def collectTactics (t : Elab.InfoTree) : List TacticInvocation :=
|
||||||
collectTacticNodes t |>.filter fun i => i.info.isSubstantive
|
collectTacticNodes t |>.filter fun i => i.info.isSubstantive
|
||||||
|
|
||||||
|
@[export pantograph_frontend_collect_tactics_from_compilation_step_m]
|
||||||
|
def collectTacticsFromCompilationStep (step : CompilationStep) : IO (List Protocol.InvokedTactic) := do
|
||||||
|
let tacticInfoTrees := step.trees.bind λ tree => tree.filter λ
|
||||||
|
| info@(.ofTacticInfo _) => info.isOriginal
|
||||||
|
| _ => false
|
||||||
|
let tactics := tacticInfoTrees.bind collectTactics
|
||||||
|
tactics.mapM λ invocation => do
|
||||||
|
let goalBefore := (Format.joinSep (← invocation.goalState) "\n").pretty
|
||||||
|
let goalAfter := (Format.joinSep (← invocation.goalStateAfter) "\n").pretty
|
||||||
|
let tactic ← invocation.ctx.runMetaM {} do
|
||||||
|
let t ← Lean.PrettyPrinter.ppTactic ⟨invocation.info.stx⟩
|
||||||
|
return t.pretty
|
||||||
|
return { goalBefore, goalAfter, tactic }
|
||||||
|
|
||||||
end Pantograph.Compile
|
end Pantograph.Frontend
|
|
@ -3,9 +3,7 @@ Functions for handling metavariables
|
||||||
|
|
||||||
All the functions starting with `try` resume their inner monadic state.
|
All the functions starting with `try` resume their inner monadic state.
|
||||||
-/
|
-/
|
||||||
import Pantograph.Protocol
|
|
||||||
import Pantograph.Tactic
|
import Pantograph.Tactic
|
||||||
import Pantograph.Compile.Parse
|
|
||||||
import Lean
|
import Lean
|
||||||
|
|
||||||
|
|
||||||
|
@ -385,20 +383,4 @@ protected def GoalState.tryCalc (state: GoalState) (goal: MVarId) (pred: String)
|
||||||
catch exception =>
|
catch exception =>
|
||||||
return .failure #[← exception.toMessageData.toString]
|
return .failure #[← exception.toMessageData.toString]
|
||||||
|
|
||||||
|
|
||||||
protected def GoalState.tryMotivatedApply (state: GoalState) (goal: MVarId) (recursor: String):
|
|
||||||
Elab.TermElabM TacticResult := do
|
|
||||||
state.restoreElabM
|
|
||||||
let recursor ← match (← Compile.parseTermM recursor) with
|
|
||||||
| .ok syn => pure syn
|
|
||||||
| .error error => return .parseError error
|
|
||||||
state.tryTacticM goal (tacticM := Tactic.evalMotivatedApply recursor)
|
|
||||||
protected def GoalState.tryNoConfuse (state: GoalState) (goal: MVarId) (eq: String):
|
|
||||||
Elab.TermElabM TacticResult := do
|
|
||||||
state.restoreElabM
|
|
||||||
let eq ← match (← Compile.parseTermM eq) with
|
|
||||||
| .ok syn => pure syn
|
|
||||||
| .error error => return .parseError error
|
|
||||||
state.tryTacticM goal (tacticM := Tactic.evalNoConfuse eq)
|
|
||||||
|
|
||||||
end Pantograph
|
end Pantograph
|
||||||
|
|
|
@ -159,7 +159,7 @@ def goalAssign (state: GoalState) (goal: MVarId) (expr: String): CoreM TacticRes
|
||||||
runTermElabM <| state.tryAssign goal expr
|
runTermElabM <| state.tryAssign goal expr
|
||||||
@[export pantograph_goal_have_m]
|
@[export pantograph_goal_have_m]
|
||||||
protected def GoalState.tryHave (state: GoalState) (goal: MVarId) (binderName: String) (type: String): CoreM TacticResult := do
|
protected def GoalState.tryHave (state: GoalState) (goal: MVarId) (binderName: String) (type: String): CoreM TacticResult := do
|
||||||
let type ← match (← Compile.parseTermM type) with
|
let type ← match (← parseTermM type) with
|
||||||
| .ok syn => pure syn
|
| .ok syn => pure syn
|
||||||
| .error error => return .parseError error
|
| .error error => return .parseError error
|
||||||
runTermElabM do
|
runTermElabM do
|
||||||
|
@ -167,12 +167,28 @@ protected def GoalState.tryHave (state: GoalState) (goal: MVarId) (binderName: S
|
||||||
state.tryTacticM goal $ Tactic.evalHave binderName.toName type
|
state.tryTacticM goal $ Tactic.evalHave binderName.toName type
|
||||||
@[export pantograph_goal_try_define_m]
|
@[export pantograph_goal_try_define_m]
|
||||||
protected def GoalState.tryDefine (state: GoalState) (goal: MVarId) (binderName: String) (expr: String): CoreM TacticResult := do
|
protected def GoalState.tryDefine (state: GoalState) (goal: MVarId) (binderName: String) (expr: String): CoreM TacticResult := do
|
||||||
let expr ← match (← Compile.parseTermM expr) with
|
let expr ← match (← parseTermM expr) with
|
||||||
| .ok syn => pure syn
|
| .ok syn => pure syn
|
||||||
| .error error => return .parseError error
|
| .error error => return .parseError error
|
||||||
runTermElabM do
|
runTermElabM do
|
||||||
state.restoreElabM
|
state.restoreElabM
|
||||||
state.tryTacticM goal (Tactic.evalDefine binderName.toName expr)
|
state.tryTacticM goal (Tactic.evalDefine binderName.toName expr)
|
||||||
|
@[export pantograph_goal_try_motivated_apply_m]
|
||||||
|
protected def GoalState.tryMotivatedApply (state: GoalState) (goal: MVarId) (recursor: String):
|
||||||
|
Elab.TermElabM TacticResult := do
|
||||||
|
state.restoreElabM
|
||||||
|
let recursor ← match (← parseTermM recursor) with
|
||||||
|
| .ok syn => pure syn
|
||||||
|
| .error error => return .parseError error
|
||||||
|
state.tryTacticM goal (tacticM := Tactic.evalMotivatedApply recursor)
|
||||||
|
@[export pantograph_goal_try_no_confuse_m]
|
||||||
|
protected def GoalState.tryNoConfuse (state: GoalState) (goal: MVarId) (eq: String):
|
||||||
|
Elab.TermElabM TacticResult := do
|
||||||
|
state.restoreElabM
|
||||||
|
let eq ← match (← parseTermM eq) with
|
||||||
|
| .ok syn => pure syn
|
||||||
|
| .error error => return .parseError error
|
||||||
|
state.tryTacticM goal (tacticM := Tactic.evalNoConfuse eq)
|
||||||
@[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
|
||||||
|
|
|
@ -286,8 +286,6 @@ structure GoalDiag where
|
||||||
/-- Executes the Lean compiler on a single file -/
|
/-- Executes the Lean compiler on a single file -/
|
||||||
structure CompileUnit where
|
structure CompileUnit where
|
||||||
module: String
|
module: String
|
||||||
-- If set to true, query the string boundaries of compilation units
|
|
||||||
compilationUnits: Bool := false
|
|
||||||
-- If set to true, collect tactic invocations
|
-- If set to true, collect tactic invocations
|
||||||
invocations: Bool := false
|
invocations: Bool := false
|
||||||
deriving Lean.FromJson
|
deriving Lean.FromJson
|
||||||
|
@ -297,7 +295,8 @@ structure InvokedTactic where
|
||||||
tactic: String
|
tactic: String
|
||||||
deriving Lean.ToJson
|
deriving Lean.ToJson
|
||||||
structure CompileUnitResult where
|
structure CompileUnitResult where
|
||||||
units?: Option $ List (Nat × Nat)
|
-- String boundaries of compilation units
|
||||||
|
units: List (Nat × Nat)
|
||||||
invocations?: Option $ List InvokedTactic
|
invocations?: Option $ List InvokedTactic
|
||||||
deriving Lean.ToJson
|
deriving Lean.ToJson
|
||||||
|
|
||||||
|
|
|
@ -26,6 +26,13 @@ def parseTerm (env: Environment) (s: String): Except String Syntax :=
|
||||||
(input := s)
|
(input := s)
|
||||||
(fileName := "<stdin>")
|
(fileName := "<stdin>")
|
||||||
|
|
||||||
|
def parseTermM [Monad m] [MonadEnv m] (s: String): m (Except String Syntax) := do
|
||||||
|
return Parser.runParserCategory
|
||||||
|
(env := ← MonadEnv.getEnv)
|
||||||
|
(catName := `term)
|
||||||
|
(input := s)
|
||||||
|
(fileName := "<stdin>")
|
||||||
|
|
||||||
/-- Parse a syntax object. May generate additional metavariables! -/
|
/-- Parse a syntax object. May generate additional metavariables! -/
|
||||||
def elabType (syn: Syntax): Elab.TermElabM (Except String Expr) := do
|
def elabType (syn: Syntax): Elab.TermElabM (Except String Expr) := do
|
||||||
try
|
try
|
||||||
|
|
23
Repl.lean
23
Repl.lean
|
@ -1,7 +1,7 @@
|
||||||
import Lean.Data.HashMap
|
import Lean.Data.HashMap
|
||||||
import Pantograph
|
import Pantograph
|
||||||
|
|
||||||
namespace Pantograph
|
namespace Pantograph.Repl
|
||||||
|
|
||||||
structure Context where
|
structure Context where
|
||||||
imports: List String
|
imports: List String
|
||||||
|
@ -204,17 +204,20 @@ def execute (command: Protocol.Command): MainM Lean.Json := do
|
||||||
compile_unit (args: Protocol.CompileUnit): MainM (CR Protocol.CompileUnitResult) := do
|
compile_unit (args: Protocol.CompileUnit): MainM (CR Protocol.CompileUnitResult) := do
|
||||||
let module := args.module.toName
|
let module := args.module.toName
|
||||||
try
|
try
|
||||||
let steps ← Compile.processSource module
|
let li ← Frontend.runFrontendMInFile module {} <| Frontend.mapCompilationSteps λ step => do
|
||||||
let units? := if args.compilationUnits then
|
let unitBoundary := (step.src.startPos.byteIdx, step.src.stopPos.byteIdx)
|
||||||
.some $ steps.map λ step => (step.src.startPos.byteIdx, step.src.stopPos.byteIdx)
|
let tacticInvocations ← if args.invocations then
|
||||||
|
Frontend.collectTacticsFromCompilationStep step
|
||||||
|
else
|
||||||
|
pure []
|
||||||
|
return (unitBoundary, tacticInvocations)
|
||||||
|
let units := li.map λ (unit, _) => unit
|
||||||
|
let invocations? := if args.invocations then
|
||||||
|
.some $ li.bind λ (_, invocations) => invocations
|
||||||
else
|
else
|
||||||
.none
|
.none
|
||||||
let invocations? ← if args.invocations then
|
return .ok { units, invocations? }
|
||||||
pure $ .some (← Compile.collectTacticsFromCompilation steps)
|
|
||||||
else
|
|
||||||
pure .none
|
|
||||||
return .ok { units?, invocations? }
|
|
||||||
catch e =>
|
catch e =>
|
||||||
return .error $ errorI "compile" (← e.toMessageData.toString)
|
return .error $ errorI "compile" (← e.toMessageData.toString)
|
||||||
|
|
||||||
end Pantograph
|
end Pantograph.Repl
|
||||||
|
|
|
@ -6,13 +6,13 @@ import Repl
|
||||||
import Test.Common
|
import Test.Common
|
||||||
|
|
||||||
namespace Pantograph.Test.Integration
|
namespace Pantograph.Test.Integration
|
||||||
open Pantograph
|
open Pantograph.Repl
|
||||||
|
|
||||||
def step { α } [Lean.ToJson α] (cmd: String) (payload: List (String × Lean.Json))
|
def step { α } [Lean.ToJson α] (cmd: String) (payload: List (String × Lean.Json))
|
||||||
(expected: α) (name? : Option String := .none): MainM LSpec.TestSeq := do
|
(expected: α) (name? : Option String := .none): MainM LSpec.TestSeq := do
|
||||||
let payload := Lean.Json.mkObj payload
|
let payload := Lean.Json.mkObj payload
|
||||||
let name := name?.getD s!"{cmd} {payload.compress}"
|
let name := name?.getD s!"{cmd} {payload.compress}"
|
||||||
let result ← execute { cmd, payload }
|
let result ← Repl.execute { cmd, payload }
|
||||||
return LSpec.test name (toString result = toString (Lean.toJson expected))
|
return LSpec.test name (toString result = toString (Lean.toJson expected))
|
||||||
|
|
||||||
abbrev Test := List (MainM LSpec.TestSeq)
|
abbrev Test := List (MainM LSpec.TestSeq)
|
||||||
|
|
Loading…
Reference in New Issue