You signed in with another tab or window. Reload to refresh your session.You signed out in another tab or window. Reload to refresh your session.You switched accounts on another tab or window. Reload to refresh your session.Dismiss alert
This seems to be the underlying cause of leanprover-community/mathlib4#15785 and leanprover-community/aesop#153 . Unfortunately, since it involves the metaprogramming API to simp, it is difficult to give a small reproduction. The example below uses a minimalistic simp wrapper which uses a pre method to replace occurrences of G α with Nat.zero. Importantly, while doing so it creates and instantiates a universe metavariable, and something in the main simp loop seems to forget that it is instantiated, presumably rolling back the metavariable context while still holding on to the expressions generated in that context, leading to panics and unsolved metavariables in the resulting proofs.
import Lean
axiomF (s : Nat) (f : Nat → Nat) : Nat
@[congr]axiomF_congr {s₁ s₂ : Nat} {f : Nat → Nat} (h : s₁ = s₂) : F s₁ f = F s₂ f
axiomG.{u} : Sort u → Nat
axiomG_zero (α) : G α = Nat.zero
theoremfoo : F (G True) (fun x => x) ≤ F (G True) id := byopen Lean Meta Elab Tactic in run_tac
let g ← getMainGoal
let ctx := {
simpTheorems := #[← Elab.Tactic.simpOnlyBuiltins.foldlM (·.addConst ·) {}]
congrTheorems := ← getSimpCongrTheorems }
let pre e := doif e.isAppOfArity ``G1thenlet u ← mkFreshLevelMVar
let α := e.appArg!
guard <| ← isDefEq (← inferType α) (.sort u)
let pf := Expr.app (.const ``G_zero [u]) α
-- let pf ← instantiateMVars pf
logInfo <| repr pf
logInfo pf
pure <| .done { expr := .const ``Nat.zero [], proof? := pf }
else
pure .continue
let tgt ← g.getType
let (res, _) ← Simp.main tgt ctx (methods := { pre })
logInfo res.proof?
replaceMainGoal [← applySimpResultToTarget g tgt res]
apply Nat.le_refl
-- invalid occurrence of universe level 'u_1' at 'foo', it does not occur at the declaration type,-- nor it is explicit universe level provided by the user, occurring at expression-- G.{u_1} True-- at declaration body-- Eq.mpr (id (congr (congrArg LE.le (F_congr (G_zero True))) (congrArg (fun (x : Nat) ↦ F x id) (G_zero True))))-- (Nat.le_refl (F Nat.zero fun (x : Nat) ↦ x))
The text was updated successfully, but these errors were encountered:
This seems to be the underlying cause of leanprover-community/mathlib4#15785 and leanprover-community/aesop#153 . Unfortunately, since it involves the metaprogramming API to simp, it is difficult to give a small reproduction. The example below uses a minimalistic simp wrapper which uses a
pre
method to replace occurrences ofG α
withNat.zero
. Importantly, while doing so it creates and instantiates a universe metavariable, and something in the main simp loop seems to forget that it is instantiated, presumably rolling back the metavariable context while still holding on to the expressions generated in that context, leading to panics and unsolved metavariables in the resulting proofs.The text was updated successfully, but these errors were encountered: