Pantograph/Pantograph/Commands.lean

152 lines
3.9 KiB
Plaintext
Raw Blame History

This file contains ambiguous Unicode characters

This file contains Unicode characters that might be confused with other characters. If you think that this is intentional, you can safely ignore this warning. Use the Escape button to reveal them.

/-
All the command input/output structures are stored here
Note that no command other than `InteractionError` may have `error` as one of
its field names to avoid confusion with error messages generated by the REPL.
-/
import Lean.Data.Json
namespace Pantograph.Commands
/-- Main Option structure, placed here to avoid name collision -/
structure Options where
-- When enabled, pretty print every expression
printExprPretty: Bool := true
-- When enabled, print the raw AST of expressions
printExprAST: Bool := false
-- When enabled, the types and values of persistent variables in a proof goal
-- are not shown unless they are new to the proof step. Reduces overhead
proofVariableDelta: Bool := false
deriving Lean.ToJson
--- Expression Objects ---
structure BoundExpression where
binders: Array (String × String)
target: String
deriving Lean.ToJson
structure Expression where
-- Pretty printed expression
pp?: Option String := .none
-- AST structure
sexp?: Option String := .none
deriving Lean.ToJson
structure Variable where
name: String
/-- Does the name contain a dagger -/
isInaccessible: Bool := false
type: Expression
value?: Option Expression := .none
deriving Lean.ToJson
structure Goal where
/-- String case id -/
caseName?: Option String := .none
/-- Is the goal in conversion mode -/
isConversion: Bool := false
/-- target expression type -/
target: Expression
/-- Variables -/
vars: Array Variable := #[]
deriving Lean.ToJson
--- Individual Commands and return types ---
structure Command where
cmd: String
payload: Lean.Json
deriving Lean.FromJson
structure InteractionError where
error: String
desc: String
deriving Lean.ToJson
--- Individual command and return types ---
/-- Set options; See `Options` struct above for meanings -/
structure OptionsSet where
printExprPretty?: Option Bool
printExprAST?: Option Bool
proofVariableDelta?: Option Bool
deriving Lean.FromJson
structure OptionsSetResult where
deriving Lean.ToJson
structure OptionsPrint where
deriving Lean.FromJson
abbrev OptionsPrintResult := Options
-- Print all symbols in environment
structure Catalog where
deriving Lean.FromJson
structure CatalogResult where
symbols: Array String
deriving Lean.ToJson
-- Print the type of a symbol
structure Inspect where
name: String
-- If true/false, show/hide the value expressions; By default definitions
-- values are shown and theorem values are hidden.
value?: Option Bool := .some false
deriving Lean.FromJson
structure InspectResult where
type: Expression
value?: Option Expression := .none
module?: Option String
deriving Lean.ToJson
structure ClearResult where
nTrees: Nat
deriving Lean.ToJson
-- Get the type of an expression
structure ExprType where
expr: String
deriving Lean.FromJson
structure ExprTypeResult where
type: String
roundTrip: String
deriving Lean.ToJson
structure ProofStart where
name: Option String -- Identifier of the proof
-- Only one of the fields below may be populated.
expr: Option String -- Proof expression
copyFrom: Option String -- Theorem name
deriving Lean.FromJson
structure ProofStartResult where
treeId: Nat := 0 -- Proof tree id
deriving Lean.ToJson
structure ProofTactic where
-- Identifiers for tree, state, and goal
treeId: Nat
stateId: Nat
goalId: Option Nat
tactic: String
deriving Lean.FromJson
structure ProofTacticResultSuccess where
goals: Array Goal
nextId?: Option Nat -- Next proof state id
deriving Lean.ToJson
structure ProofTacticResultFailure where
tacticErrors: Array String -- Error messages generated by tactic
deriving Lean.ToJson
structure ProofPrintTree where
treeId: Nat
deriving Lean.FromJson
structure ProofPrintTreeResult where
-- "" if no parents, otherwise "parentId.goalId"
parents: Array String
deriving Lean.ToJson
end Pantograph.Commands