Merge pull request 'feat(delate): Unfold matchers' (#176) from delate/instantiate into dev

Reviewed-on: #176
This commit is contained in:
Leni Aniva 2025-03-17 18:37:55 -07:00
commit bc37482212
6 changed files with 105 additions and 82 deletions

View File

@ -37,7 +37,7 @@ def analyzeProjection (env: Environment) (e: Expr): Projection :=
def anonymousLevel : Level := .mvar ⟨.anonymous⟩ def anonymousLevel : Level := .mvar ⟨.anonymous⟩
@[export pantograph_expr_proj_to_app] @[export pantograph_expr_proj_to_app]
def exprProjToApp (env: Environment) (e: Expr): Expr := def exprProjToApp (env : Environment) (e : Expr) : Expr :=
let anon : Expr := .mvar ⟨.anonymous⟩ let anon : Expr := .mvar ⟨.anonymous⟩
match analyzeProjection env e with match analyzeProjection env e with
| .field projector numParams => | .field projector numParams =>
@ -66,9 +66,18 @@ def exprProjToApp (env: Environment) (e: Expr): Expr :=
def _root_.Lean.Name.isAuxLemma (n : Lean.Name) : Bool := n matches .num (.str _ "_auxLemma") _ def _root_.Lean.Name.isAuxLemma (n : Lean.Name) : Bool := n matches .num (.str _ "_auxLemma") _
/-- Unfold all lemmas created by `Lean.Meta.mkAuxLemma`. These end in `_auxLemma.nn` where `nn` is a number. -/ /-- Unfold all lemmas created by `Lean.Meta.mkAuxLemma`. These end in `_auxLemma.nn` where `nn` is a number. -/
@[export pantograph_unfold_aux_lemmas] @[export pantograph_unfold_aux_lemmas_m]
def unfoldAuxLemmas (e : Expr) : CoreM Expr := do def unfoldAuxLemmas : Expr → CoreM Expr :=
Lean.Meta.deltaExpand e Lean.Name.isAuxLemma (Meta.deltaExpand · Lean.Name.isAuxLemma)
/-- Unfold all matcher applications -/
@[export pantograph_unfold_matchers_m]
def unfoldMatchers (expr : Expr) : CoreM Expr :=
Core.transform expr λ e => do
let .some mapp ← Meta.matchMatcherApp? e | return .continue e
let .some matcherInfo := (← getEnv).find? mapp.matcherName | panic! "Matcher must exist"
let f ← Meta.instantiateValueLevelParams matcherInfo mapp.matcherLevels.toList
let mdata := KVMap.empty.insert `matcher (DataValue.ofName mapp.matcherName)
return .visit $ .mdata mdata (f.betaRev e.getAppRevArgs (useZeta := true))
/-- /--
Force the instantiation of delayed metavariables even if they cannot be fully Force the instantiation of delayed metavariables even if they cannot be fully
@ -85,91 +94,92 @@ This function ensures any metavariable in the result is either
1. Delayed assigned with its pending mvar not assigned in any form 1. Delayed assigned with its pending mvar not assigned in any form
2. Not assigned (delay or not) 2. Not assigned (delay or not)
-/ -/
partial def instantiateDelayedMVars (eOrig: Expr) : MetaM Expr := do partial def instantiateDelayedMVars (expr : Expr) : MetaM Expr :=
--let padding := String.join $ List.replicate level "│ " withTraceNode `Pantograph.Delate (λ _ex => return m!":= {← Meta.ppExpr expr}") do
--IO.println s!"{padding}Starting {toString eOrig}" let mut result ← Meta.transform (← instantiateMVars expr)
let mut result ← Meta.transform (← instantiateMVars eOrig) λ e => e.withApp fun f args => do
(pre := fun e => e.withApp fun f args => do let .mvar mvarId := f | return .continue
let .mvar mvarId := f | return .continue trace[Pantograph.Delate] "V {e}"
--IO.println s!"{padding}├V {e}" let mvarDecl ← mvarId.getDecl
let mvarDecl ← mvarId.getDecl
-- This is critical to maintaining the interdependency of metavariables. -- This is critical to maintaining the interdependency of metavariables.
-- Without setting `.syntheticOpaque`, Lean's metavariable elimination -- Without setting `.syntheticOpaque`, Lean's metavariable elimination
-- system will not make the necessary delayed assigned mvars in case of -- system will not make the necessary delayed assigned mvars in case of
-- nested mvars. -- nested mvars.
mvarId.setKind .syntheticOpaque mvarId.setKind .syntheticOpaque
mvarId.withContext do mvarId.withContext do
let lctx ← MonadLCtx.getLCtx let lctx ← MonadLCtx.getLCtx
if mvarDecl.lctx.any (λ decl => !lctx.contains decl.fvarId) then if mvarDecl.lctx.any (λ decl => !lctx.contains decl.fvarId) then
let violations := mvarDecl.lctx.decls.foldl (λ acc decl? => match decl? with let violations := mvarDecl.lctx.decls.foldl (λ acc decl? => match decl? with
| .some decl => if lctx.contains decl.fvarId then acc else acc ++ [decl.fvarId.name] | .some decl => if lctx.contains decl.fvarId then acc else acc ++ [decl.fvarId.name]
| .none => acc) [] | .none => acc) []
panic! s!"In the context of {mvarId.name}, there are local context variable violations: {violations}" panic! s!"In the context of {mvarId.name}, there are local context variable violations: {violations}"
if let .some assign ← getExprMVarAssignment? mvarId then if let .some assign ← getExprMVarAssignment? mvarId then
--IO.println s!"{padding}├A ?{mvarId.name}" trace[Pantograph.Delate] "A ?{mvarId.name}"
assert! !(← mvarId.isDelayedAssigned) assert! !(← mvarId.isDelayedAssigned)
return .visit (mkAppN assign args) return .visit (mkAppN assign args)
else if let some { fvars, mvarIdPending } ← getDelayedMVarAssignment? mvarId then else if let some { fvars, mvarIdPending } ← getDelayedMVarAssignment? mvarId then
--let substTableStr := String.intercalate ", " $ Array.zipWith fvars args (λ fvar assign => s!"{fvar.fvarId!.name} := {assign}") |>.toList if ← isTracingEnabledFor `Pantograph.Delate then
--IO.println s!"{padding}├MD ?{mvarId.name} := ?{mvarIdPending.name} [{substTableStr}]" let substTableStr := String.intercalate ", " $ Array.zipWith fvars args (λ fvar assign => s!"{fvar.fvarId!.name} := {assign}") |>.toList
trace[Pantograph.Delate]"MD ?{mvarId.name} := ?{mvarIdPending.name} [{substTableStr}]"
if args.size < fvars.size then if args.size < fvars.size then
throwError "Not enough arguments to instantiate a delay assigned mvar. This is due to bad implementations of a tactic: {args.size} < {fvars.size}. Expr: {toString e}; Origin: {toString eOrig}" throwError "Not enough arguments to instantiate a delay assigned mvar. This is due to bad implementations of a tactic: {args.size} < {fvars.size}. Expr: {toString e}; Origin: {toString expr}"
--if !args.isEmpty then if !args.isEmpty then
--IO.println s!"{padding}├── Arguments Begin" trace[Pantograph.Delate] "─ Arguments Begin"
let args ← args.mapM self let args ← args.mapM self
--if !args.isEmpty then if !args.isEmpty then
--IO.println s!"{padding}├── Arguments End" trace[Pantograph.Delate] "─ Arguments End"
if !(← mvarIdPending.isAssignedOrDelayedAssigned) then if !(← mvarIdPending.isAssignedOrDelayedAssigned) then
--IO.println s!"{padding}├T1" trace[Pantograph.Delate] "T1"
let result := mkAppN f args let result := mkAppN f args
return .done result
let pending ← mvarIdPending.withContext do
let inner ← instantiateDelayedMVars (.mvar mvarIdPending) --(level := level + 1)
--IO.println s!"{padding}├Pre: {inner}"
pure <| (← inner.abstractM fvars).instantiateRev args
-- Tail arguments
let result := mkAppRange pending fvars.size args.size args
--IO.println s!"{padding}├MD {result}"
return .done result return .done result
else
assert! !(← mvarId.isAssigned)
assert! !(← mvarId.isDelayedAssigned)
--if !args.isEmpty then
-- IO.println s!"{padding}├── Arguments Begin"
let args ← args.mapM self
--if !args.isEmpty then
-- IO.println s!"{padding}├── Arguments End"
--IO.println s!"{padding}├M ?{mvarId.name}" let pending ← mvarIdPending.withContext do
return .done (mkAppN f args)) let inner ← instantiateDelayedMVars (.mvar mvarIdPending)
--IO.println s!"{padding}└Result {result}" trace[Pantograph.Delate] "Pre: {inner}"
pure <| (← inner.abstractM fvars).instantiateRev args
-- Tail arguments
let result := mkAppRange pending fvars.size args.size args
trace[Pantograph.Delate] "MD {result}"
return .done result
else
assert! !(← mvarId.isAssigned)
assert! !(← mvarId.isDelayedAssigned)
if !args.isEmpty then
trace[Pantograph.Delate] "─ Arguments Begin"
let args ← args.mapM self
if !args.isEmpty then
trace[Pantograph.Delate] "─ Arguments End"
trace[Pantograph.Delate] "M ?{mvarId.name}"
return .done (mkAppN f args)
trace[Pantoraph.Delate] "Result {result}"
return result return result
where where
self e := instantiateDelayedMVars e --(level := level + 1) self e := instantiateDelayedMVars e
/-- /--
Convert an expression to an equiavlent form with Convert an expression to an equiavlent form with
1. No nested delayed assigned mvars 1. No nested delayed assigned mvars
2. No aux lemmas 2. No aux lemmas or matchers
3. No assigned mvars 3. No assigned mvars
-/ -/
@[export pantograph_instantiate_all_m] @[export pantograph_instantiate_all_m]
def instantiateAll (e: Expr): MetaM Expr := do def instantiateAll (e : Expr) : MetaM Expr := do
let e ← instantiateDelayedMVars e let e ← instantiateDelayedMVars e
let e ← unfoldAuxLemmas e let e ← unfoldAuxLemmas e
let e ← unfoldMatchers e
return e return e
structure DelayedMVarInvocation where structure DelayedMVarInvocation where
mvarIdPending: MVarId mvarIdPending : MVarId
args: Array (FVarId × (Option Expr)) args : Array (FVarId × (Option Expr))
-- Extra arguments applied to the result of this substitution -- Extra arguments applied to the result of this substitution
tail: Array Expr tail : Array Expr
-- The pending mvar of any delayed assigned mvar must not be assigned in any way. -- The pending mvar of any delayed assigned mvar must not be assigned in any way.
@[export pantograph_to_delayed_mvar_invocation_m] @[export pantograph_to_delayed_mvar_invocation_m]
@ -594,4 +604,7 @@ protected def GoalState.diag (goalState: GoalState) (parent?: Option GoalState :
| other => s!"[{other}]" | other => s!"[{other}]"
parentHasMVar (mvarId: MVarId): Bool := parent?.map (λ state => state.mctx.decls.contains mvarId) |>.getD true parentHasMVar (mvarId: MVarId): Bool := parent?.map (λ state => state.mctx.decls.contains mvarId) |>.getD true
initialize
registerTraceClass `Pantograph.Delate
end Pantograph end Pantograph

View File

@ -62,7 +62,7 @@ private partial def translateExpr (srcExpr: Expr) : MetaTranslateM Expr := do
let sourceMCtx ← getSourceMCtx let sourceMCtx ← getSourceMCtx
-- We want to create as few mvars as possible -- We want to create as few mvars as possible
let (srcExpr, _) := instantiateMVarsCore (mctx := sourceMCtx) srcExpr let (srcExpr, _) := instantiateMVarsCore (mctx := sourceMCtx) srcExpr
--IO.println s!"Transform src: {srcExpr}" trace[Pantograph.Frontend.MetaTranslate] "Transform src: {srcExpr}"
let result ← Core.transform srcExpr λ e => do let result ← Core.transform srcExpr λ e => do
let state ← get let state ← get
match e with match e with
@ -100,10 +100,10 @@ partial def translateLocalDecl (srcLocalDecl: LocalDecl) : MetaTranslateM LocalD
addTranslatedFVar srcLocalDecl.fvarId fvarId addTranslatedFVar srcLocalDecl.fvarId fvarId
match srcLocalDecl with match srcLocalDecl with
| .cdecl index _ userName type bi kind => do | .cdecl index _ userName type bi kind => do
--IO.println s!"[CD] {userName} {toString type}" trace[Pantograph.Frontend.MetaTranslate] "[CD] {userName} {toString type}"
return .cdecl index fvarId userName (← translateExpr type) bi kind return .cdecl index fvarId userName (← translateExpr type) bi kind
| .ldecl index _ userName type value nonDep kind => do | .ldecl index _ userName type value nonDep kind => do
--IO.println s!"[LD] {toString type} := {toString value}" trace[Pantograph.Frontend.MetaTranslate] "[LD] {toString type} := {toString value}"
return .ldecl index fvarId userName (← translateExpr type) (← translateExpr value) nonDep kind return .ldecl index fvarId userName (← translateExpr type) (← translateExpr value) nonDep kind
partial def translateLCtx : MetaTranslateM LocalContext := do partial def translateLCtx : MetaTranslateM LocalContext := do
@ -162,4 +162,7 @@ end MetaTranslate
export MetaTranslate (MetaTranslateM) export MetaTranslate (MetaTranslateM)
initialize
registerTraceClass `Pantograph.Frontend.MetaTranslate
end Pantograph.Frontend end Pantograph.Frontend

View File

@ -346,10 +346,9 @@ structure FrontendProcess where
-- One of these two must be supplied: Either supply the file name or the content. -- One of these two must be supplied: Either supply the file name or the content.
fileName?: Option String := .none fileName?: Option String := .none
file?: Option String := .none file?: Option String := .none
-- If set to true, read the header (otherwise discards the header) -- Whether to read the header
readHeader : Bool := false readHeader : Bool := false
-- If set to true, Pantograph's environment will be altered to whatever is -- Alter the REPL environment after the compilation units.
-- after the compilation units.
inheritEnv : Bool := false inheritEnv : Bool := false
-- collect tactic invocations -- collect tactic invocations
invocations: Bool := false invocations: Bool := false

View File

@ -108,7 +108,7 @@ def strToTermSyntax (s: String): CoreM Syntax := do
(input := s) (input := s)
(fileName := ← getFileName) | panic! s!"Failed to parse {s}" (fileName := ← getFileName) | panic! s!"Failed to parse {s}"
return stx return stx
def parseSentence (s: String): Elab.TermElabM Expr := do def parseSentence (s : String) (expectedType? : Option Expr := .none) : Elab.TermElabM Expr := do
let stx ← match Parser.runParserCategory let stx ← match Parser.runParserCategory
(env := ← MonadEnv.getEnv) (env := ← MonadEnv.getEnv)
(catName := `term) (catName := `term)
@ -116,7 +116,7 @@ def parseSentence (s: String): Elab.TermElabM Expr := do
(fileName := ← getFileName) with (fileName := ← getFileName) with
| .ok syn => pure syn | .ok syn => pure syn
| .error error => throwError "Failed to parse: {error}" | .error error => throwError "Failed to parse: {error}"
Elab.Term.elabTerm (stx := stx) .none Elab.Term.elabTerm (stx := stx) expectedType?
def runTacticOnMVar (tacticM: Elab.Tactic.TacticM Unit) (goal: MVarId): Elab.TermElabM (List MVarId) := do def runTacticOnMVar (tacticM: Elab.Tactic.TacticM Unit) (goal: MVarId): Elab.TermElabM (List MVarId) := do
let (_, newGoals) ← tacticM { elaborator := .anonymous } |>.run { goals := [goal] } let (_, newGoals) ← tacticM { elaborator := .anonymous } |>.run { goals := [goal] }

View File

@ -3,10 +3,9 @@ import Pantograph.Delate
import Test.Common import Test.Common
import Lean import Lean
open Lean open Lean Pantograph
namespace Pantograph.Test.Delate
open Pantograph namespace Pantograph.Test.Delate
deriving instance Repr, DecidableEq for Protocol.BoundExpression deriving instance Repr, DecidableEq for Protocol.BoundExpression
@ -113,6 +112,13 @@ def test_projection_exists (env: Environment) : IO LSpec.TestSeq:= runTest do
checkEq "numParams" numParams 2 checkEq "numParams" numParams 2
checkEq "numFields" numFields 2 checkEq "numFields" numFields 2
def test_matcher : TestT Elab.TermElabM Unit := do
let t ← parseSentence "Nat → Nat"
let e ← parseSentence "fun (n : Nat) => match n with | 0 => 0 | k => k" (.some t)
let .some _ ← Meta.matchMatcherApp? e.bindingBody! | fail "Must be a matcher app"
let e' ← instantiateAll e
checkTrue "ok" <| ← Meta.isTypeCorrect e'
def suite (env: Environment): List (String × IO LSpec.TestSeq) := def suite (env: Environment): List (String × IO LSpec.TestSeq) :=
[ [
("serializeName", do pure test_serializeName), ("serializeName", do pure test_serializeName),
@ -123,6 +129,7 @@ def suite (env: Environment): List (String × IO LSpec.TestSeq) :=
("Instance", test_instance env), ("Instance", test_instance env),
("Projection Prod", test_projection_prod env), ("Projection Prod", test_projection_prod env),
("Projection Exists", test_projection_exists env), ("Projection Exists", test_projection_exists env),
("Matcher", runTestTermElabM env test_matcher),
] ]
end Pantograph.Test.Delate end Pantograph.Test.Delate

View File

@ -46,14 +46,15 @@ See `Pantograph/Protocol.lean` for a description of the parameters and return va
Save/Load a goal state to/from a file. The environment is not carried with the Save/Load a goal state to/from a file. The environment is not carried with the
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>], readHeader: <bool>, inheritEnv: <bool>, invocations:
<bool>, sorrys: <bool>, typeErrorsAsGoals: <bool>, newConstants: <bool> }`: <bool>, sorrys: <bool>, typeErrorsAsGoals: <bool>, newConstants: <bool> }`:
Executes the Lean frontend on a file, collecting the tactic invocations Executes the Lean frontend on a file, collecting the tactic invocations
(`"invocations": true`), the sorrys and type errors into goal states (`"invocations": true`), the sorrys and type errors into goal states
(`"sorrys": true`), and new constants (`"newConstants": true`). In the case of (`"sorrys": true`), and new constants (`"newConstants": true`). In the case of
`sorrys`, this command additionally outputs the position of each captured `sorrys`, this command additionally outputs the position of each captured
`sorry`. Warning: Behaviour is unstable in case of multiple `sorry`s. Use the `sorry`. Conditionally inherit the environment from executing the file.
draft tactic if possible. Warning: Behaviour is unstable in case of multiple `sorry`s. Use the draft
tactic if possible.
## Errors ## Errors