An Origin
is an identifier for simp theorems which indicates roughly
what action the user took which lead to this theorem existing in the simp set.
- decl: Lake.Name → optParam Bool true → optParam Bool false → Lean.Meta.Origin
A global declaration in the environment.
- fvar: Lean.FVarId → Lean.Meta.Origin
A local hypothesis. When
contextual := true
is enabled, this fvar may exist in an extension of the current local context; it will not be used for rewriting by simp once it is out of scope but it may end up in theusedSimps
trace. - stx: Lake.Name → Lean.Syntax → Lean.Meta.Origin
A proof term provided directly to a call to
simp [ref, ...]
whereref
is the provided simp argument (of kindParser.Tactic.simpLemma
). Theid
is a unique identifier for the call. - other: Lake.Name → Lean.Meta.Origin
Some other origin.
name
should not collide with the other types for erasure to work correctly, and simp trace will ignore this lemma. The other origins should be preferred if possible.
Instances For
Equations
- Lean.Meta.instInhabitedOrigin = { default := Lean.Meta.Origin.decl default default default }
Equations
- Lean.Meta.instReprOrigin = { reprPrec := Lean.Meta.reprOrigin✝ }
A unique identifier corresponding to the origin.
Equations
- (Lean.Meta.Origin.decl a a_1 a_2).key = a
- (Lean.Meta.Origin.fvar a).key = a.name
- (Lean.Meta.Origin.stx a a_1).key = a
- (Lean.Meta.Origin.other a).key = a
Instances For
Equations
- One or more equations did not get rendered due to their size.
Equations
- One or more equations did not get rendered due to their size.
Equations
- (Lean.Meta.Origin.decl declName₁ post inv₁).lt (Lean.Meta.Origin.decl declName₂ post_1 inv₂) = (declName₁.lt declName₂ || declName₁ == declName₂ && !inv₁ && inv₂)
- (Lean.Meta.Origin.decl declName).lt x = false
- x.lt (Lean.Meta.Origin.decl declName) = true
- x✝¹.lt x✝ = x✝¹.key.lt x✝.key
Instances For
Equations
- Lean.Meta.instLTOrigin = { lt := fun (a b : Lean.Meta.Origin) => a.lt b = true }
Equations
- Lean.Meta.instDecidableLtOrigin a b = inferInstanceAs (Decidable (a.lt b = true))
Instances For
The fields levelParams
and proof
are used to encode the proof of the simp theorem.
If the proof
is a global declaration c
, we store Expr.const c []
at proof
without the universe levels, and levelParams
is set to #[]
When using the lemma, we create fresh universe metavariables.
Motivation: most simp theorems are global declarations, and this approach is faster and saves memory.
The field levelParams
is not empty only when we elaborate an expression provided by the user, and it contains universe metavariables.
Then, we use abstractMVars
to abstract the universe metavariables and create new fresh universe parameters that are stored at the field levelParams
.
- keys : Array Lean.Meta.SimpTheoremKey
It stores universe parameter names for universe polymorphic proofs. Recall that it is non-empty only when we elaborate an expression provided by the user. When
proof
is just a constant, we can use the universe parameter names stored in the declaration.- proof : Lean.Expr
- priority : Nat
- post : Bool
- perm : Bool
perm
is true if lhs and rhs are identical modulo permutation of variables. - origin : Lean.Meta.Origin
origin
is mainly relevant for producing trace messages. It is also viewed anid
used to "erase"simp
theorems fromSimpTheorems
. - rfl : Bool
Instances For
Equations
- One or more equations did not get rendered due to their size.
Equations
- Lean.Meta.isRflProof (Lean.Expr.const declName us) = liftM (Lean.Meta.isRflTheorem declName)
- Lean.Meta.isRflProof proof = do let __do_lift ← Lean.Meta.inferType proof liftM (Lean.Meta.isRflProofCore __do_lift proof)
Instances For
Equations
- One or more equations did not get rendered due to their size.
Equations
- One or more equations did not get rendered due to their size.
- Lean.Meta.ppOrigin (Lean.Meta.Origin.fvar a) = pure (Lean.MessageData.ofExpr (Lean.mkFVar a))
- Lean.Meta.ppOrigin (Lean.Meta.Origin.stx a a_1) = pure (Lean.MessageData.ofSyntax a_1)
- Lean.Meta.ppOrigin (Lean.Meta.Origin.other a) = pure (Lean.MessageData.ofName a)
Instances For
Equations
- One or more equations did not get rendered due to their size.
Instances For
Equations
- Lean.Meta.instBEqSimpTheorem = { beq := fun (e₁ e₂ : Lean.Meta.SimpTheorem) => e₁.proof == e₂.proof }
Instances For
- post : Lean.Meta.SimpTheoremTree
- lemmaNames : Lean.PHashSet Lean.Meta.Origin
- toUnfold : Lean.PHashSet Lake.Name
Constants (and let-declaration
FVarId
) to unfold. WhenzetaDelta := false
, the simplifier will expand a let-declaration if it is in this set. - erased : Lean.PHashSet Lean.Meta.Origin
- toUnfoldThms : Lean.PHashMap Lake.Name (Array Lake.Name)
Instances For
Equations
- Lean.Meta.instInhabitedSimpTheorems = { default := { pre := default, post := default, lemmaNames := default, toUnfold := default, erased := default, toUnfoldThms := default } }
Configuration for the discrimination tree.
Equations
- Lean.Meta.simpDtConfig = { iota := false, beta := true, proj := Lean.Meta.ProjReductionKind.no, zeta := true, zetaDelta := false }
Instances For
Equations
- One or more equations did not get rendered due to their size.
Instances For
Equations
Instances For
Equations
- One or more equations did not get rendered due to their size.
Instances For
Equations
- One or more equations did not get rendered due to their size.
Instances For
Return true
if declName
is tagged to be unfolded using unfoldDefinition?
(i.e., without using equational theorems).
Equations
- d.isDeclToUnfold declName = Lean.PersistentHashSet.contains d.toUnfold declName
Instances For
Equations
- d.isLetDeclToUnfold fvarId = Lean.PersistentHashSet.contains d.toUnfold fvarId.name
Instances For
Equations
- d.isLemma thmId = Lean.PersistentHashSet.contains d.lemmaNames thmId
Instances For
Register the equational theorems for the given definition.
Equations
- One or more equations did not get rendered due to their size.
Instances For
Equations
- One or more equations did not get rendered due to their size.
Instances For
- thm: Lean.Meta.SimpTheorem → Lean.Meta.SimpEntry
- toUnfold: Lake.Name → Lean.Meta.SimpEntry
- toUnfoldThms: Lake.Name → Array Lake.Name → Lean.Meta.SimpEntry
Instances For
Equations
- Lean.Meta.instInhabitedSimpEntry = { default := Lean.Meta.SimpEntry.thm default }
Equations
Instances For
Equations
- ext.getTheorems = do let __do_lift ← Lean.getEnv pure (Lean.ScopedEnvExtension.getState ext __do_lift)
Instances For
Equations
- One or more equations did not get rendered due to their size.
Instances For
Equations
- One or more equations did not get rendered due to their size.
Instances For
Instances For
Equations
- Lean.Meta.getSimpExtension? attrName = do let __do_lift ← ST.Ref.get Lean.Meta.simpExtensionMapRef pure __do_lift[attrName]?
Instances For
Auxiliary method for adding a global declaration to a SimpTheorems
datastructure.
Equations
- One or more equations did not get rendered due to their size.
Instances For
Equations
- One or more equations did not get rendered due to their size.
Instances For
Auxiliary method for creating simp theorems from a proof term val
.
Equations
- One or more equations did not get rendered due to their size.
Instances For
Auxiliary method for adding a local simp theorem to a SimpTheorems
datastructure.
Equations
- One or more equations did not get rendered due to their size.
Instances For
Reducible functions and projection functions should always be put in toUnfold
, instead
of trying to use equational theorems.
The simplifiers has special support for structure and class projections, and gets confused when they suddenly rewrite, so ignore equations for them
Equations
- Lean.Meta.SimpTheorems.ignoreEquations declName = do let __do_lift ← Lean.isProjectionFn declName let __do_lift_1 ← Lean.isReducible declName pure (__do_lift || __do_lift_1)
Instances For
Even if a function has equation theorems,
we also store it in the toUnfold
set in the following two cases:
1- It was defined by structural recursion and has a smart-unfolding associated declaration.
2- It is non-recursive.
Reason: unfoldPartialApp := true
or conditional equations may not apply.
Remark: In the future, we are planning to disable this
behavior unless unfoldPartialApp := true
.
Moreover, users will have to use f.eq_def
if they want to force the definition to be
unfolded.
Equations
- One or more equations did not get rendered due to their size.
Instances For
Equations
- One or more equations did not get rendered due to their size.
Instances For
Instances For
Equations
- One or more equations did not get rendered due to their size.
Instances For
Equations
- thmsArray.eraseTheorem thmId = Array.map (fun (thms : Lean.Meta.SimpTheorems) => thms.eraseCore thmId) thmsArray
Instances For
Equations
- thmsArray.isErased thmId = Array.any thmsArray fun (thms : Lean.Meta.SimpTheorems) => Lean.PersistentHashSet.contains thms.erased thmId
Instances For
Equations
- thmsArray.isDeclToUnfold declName = Array.any thmsArray fun (thms : Lean.Meta.SimpTheorems) => thms.isDeclToUnfold declName
Instances For
Equations
- thmsArray.isLetDeclToUnfold fvarId = Array.any thmsArray fun (thms : Lean.Meta.SimpTheorems) => thms.isLetDeclToUnfold fvarId