The result of simplifying some expression e.
- expr : ExprThe simplified version of e
- A proof that - $e = $expr, where the simplified expression is on the RHS. If- none, the proof is assumed to be- refl.
- cache : BoolIf cache := truethe result is cached. Warning: we will remove this field in the future. It is currently used byarith := true, but we can now refactor the code to avoid the hack.
Instances For
Equations
Instances For
Equations
- r₁.mkEqTrans r₂ = Lean.Meta.Simp.mkEqTransOptProofResult r₁.proof? r₁.cache r₂
Instances For
Flip the proof in a Simp.Result.
Equations
- One or more equations did not get rendered due to their size.
Instances For
Instances For
Instances For
- config : Config
- zetaDeltaSet : FVarIdSetLocal declarations to propagate to Meta.Context
- initUsedZetaDelta : FVarIdSetWhen processing Simparguments,zetaDeltamay be performed ifzetaDeltaSetis not empty. We save the local free variable ids ininitUsedZetaDelta.initUsedZetaDeltais a subset ofzetaDeltaSet.
- metaConfig : ConfigWithKey
- indexConfig : ConfigWithKey
- maxDischargeDepth : UInt32maxDischargeDepthfromconfigas anUInt32.
- simpTheorems : SimpTheoremsArray
- congrTheorems : SimpCongrTheorems
- Stores the "parent" term for the term being simplified. If a simplification procedure result depends on this value, then it is its reponsability to set - Result.cache := false.- Motivation for this field: Suppose we have a simplification procedure for normalizing arithmetic terms. Then, given a term such as - t_1 + ... + t_n, we don't want to apply the procedure to every subterm- t_1 + ... + t_ifor- i < nfor performance issues. The procedure can accomplish this by checking whether the parent term is also an arithmetical expression and do nothing if it is. However, it should set- Result.cache := falseto ensure we don't miss simplification opportunities. For example, consider the following:- example (x y : Nat) (h : y = 0) : id ((x + x) + y) = id (x + x) := by simp +arith only ...- If we don't set - Result.cache := falsefor the first- x + x, then we get the resulting state:- ... |- id (2*x + y) = id (x + x)- instead of - ... |- id (2*x + y) = id (2*x)- as expected. - Remark: given an application - f a b cthe "parent" term for- f,- a,- b, and- cis- f a b c.
- dischargeDepth : UInt32
- lctxInitIndices : NatNumber of indices in the local context when starting simp. We use this information to decide which assumptions we can use without invalidating the cache.
- inDSimp : Bool
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.
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
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
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
Equations
- ctx.isDeclToUnfold declName = ctx.simpTheorems.isDeclToUnfold declName
Instances For
Equations
- One or more equations did not get rendered due to their size.
Instances For
- Number of times each simp theorem has been used/applied. 
- Number of times each simp theorem has been tried. 
- Number of times each congr theorem has been tried. 
- thmsWithBadKeys : PArray SimpTheoremWhen using Simp.Config.index := false, andset_option diagnostics true, for every theorem used bysimp, we check whether the theorem would be also applied ifindex := true, and we store it here if it would not have been tried.
Instances For
- cache : Cache
- congrCache : CongrCache
- usedTheorems : UsedSimps
- numSteps : Nat
- diag : Diagnostics
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
Equations
- One or more equations did not get rendered due to their size.
Instances For
Executes x using a MetaM configuration for indexing terms.
It is inferred from Simp.Config.
For example, if the user has set simp (config := { zeta := false }),
isDefEq and whnf in MetaM should not perform zeta reduction.
Equations
- Lean.Meta.Simp.withSimpIndexConfig x = do let __do_lift ← readThe Lean.Meta.Simp.Context Lean.Meta.withConfigWithKey __do_lift.indexConfig x
Instances For
Executes x using a MetaM configuration for inferred from Simp.Config.
Equations
- Lean.Meta.Simp.withSimpMetaConfig x = do let __do_lift ← readThe Lean.Meta.Simp.Context Lean.Meta.withConfigWithKey __do_lift.metaConfig x
Instances For
Equations
- One or more equations did not get rendered due to their size.
Instances For
Equations
Equations
Instances For
Similar to Simproc, but resulting expression should be definitionally equal to the input one.
Equations
Instances For
Equations
- (Lean.TransformStep.done e).toStep = Lean.Meta.Simp.Step.done { expr := e }
- (Lean.TransformStep.visit e).toStep = Lean.Meta.Simp.Step.visit { expr := e }
- (Lean.TransformStep.continue (some e)).toStep = Lean.Meta.Simp.Step.continue (some { expr := e })
- Lean.TransformStep.continue.toStep = Lean.Meta.Simp.Step.continue
Instances For
Equations
- One or more equations did not get rendered due to their size.
- Lean.Meta.Simp.mkEqTransResultStep r (Lean.Meta.Simp.Step.done r') = do let __do_lift ← Lean.Meta.Simp.mkEqTransOptProofResult r.proof? r.cache r' pure (Lean.Meta.Simp.Step.done __do_lift)
- Lean.Meta.Simp.mkEqTransResultStep r (Lean.Meta.Simp.Step.visit r') = do let __do_lift ← Lean.Meta.Simp.mkEqTransOptProofResult r.proof? r.cache r' pure (Lean.Meta.Simp.Step.visit __do_lift)
- Lean.Meta.Simp.mkEqTransResultStep r Lean.Meta.Simp.Step.continue = pure (Lean.Meta.Simp.Step.continue (some r))
Instances For
"Compose" the two given simplification procedures. We use the following semantics.
- If fproducesdoneorvisit, then returnf's result.
- If fproducescontinue, then applygto new expression returned byf.
See Simp.Step type.
Equations
- One or more equations did not get rendered due to their size.
Instances For
Equations
- Lean.Meta.Simp.instAndThenSimproc = { andThen := fun (s₁ : Lean.Meta.Simp.Simproc) (s₂ : Unit → Lean.Meta.Simp.Simproc) => Lean.Meta.Simp.andThen s₁ (s₂ ()) }
Equations
- One or more equations did not get rendered due to their size.
Instances For
Equations
- Lean.Meta.Simp.instAndThenDSimproc = { andThen := fun (s₁ : Lean.Meta.Simp.DSimproc) (s₂ : Unit → Lean.Meta.Simp.DSimproc) => Lean.Meta.Simp.dandThen s₁ (s₂ ()) }
Simproc entry. It is the .olean entry plus the actual function.
- Recall that we cannot store - Simprocinto .olean files because it is a closure. Given- SimprocOLeanEntry.declName, we convert it into a- Simprocby using the unsafe function- evalConstCheck.
Instances For
Instances For
- pre : SimprocTree
- post : SimprocTree
Instances For
- pre : Simproc
- post : Simproc
- dpre : DSimproc
- dpost : DSimproc
- wellBehavedDischarge : BoolwellBehavedDischargemust not be set totrueIFdischarge?access local declarations with index >=Context.lctxInitIndiceswhencontextual := false. Reason: it would prevent us from aggressively cachingsimpresults.
Instances For
Equations
Instances For
Equations
Instances For
Equations
- Lean.Meta.Simp.getMethods = do let __do_lift ← read pure (Lean.Meta.Simp.MethodsRef.toMethods __do_lift)
Instances For
Equations
- Lean.Meta.Simp.pre e = do let __do_lift ← Lean.Meta.Simp.getMethods __do_lift.pre e
Instances For
Equations
- Lean.Meta.Simp.post e = do let __do_lift ← Lean.Meta.Simp.getMethods __do_lift.post e
Instances For
Instances For
Equations
- Lean.Meta.Simp.getConfig = do let __do_lift ← Lean.Meta.Simp.getContext pure __do_lift.config
Instances For
Equations
- Lean.Meta.Simp.getSimpTheorems = do let __do_lift ← readThe Lean.Meta.Simp.Context pure __do_lift.simpTheorems
Instances For
Equations
- Lean.Meta.Simp.getSimpCongrTheorems = do let __do_lift ← readThe Lean.Meta.Simp.Context pure __do_lift.congrTheorems
Instances For
Returns true if simp is in dsimp mode.
That is, only transformations that preserve definitional equality should be applied.
Equations
- Lean.Meta.Simp.inDSimp = do let __do_lift ← readThe Lean.Meta.Simp.Context pure __do_lift.inDSimp
Instances For
Equations
- One or more equations did not get rendered due to their size.
Instances For
Save current cache, reset it, execute x, and then restore original cache.
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
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
Equations
- One or more equations did not get rendered due to their size.
Instances For
Similar to Result.getProof, but adds a mkExpectedTypeHint if proof? is none
(i.e., result is definitionally equal to input), but we cannot establish that
source and r.expr are definitionally when using TransparencyMode.reducible.
Equations
- One or more equations did not get rendered due to their size.
Instances For
Construct the Expr cast h e, from a Simp.Result with proof h.
Instances For
Equations
- One or more equations did not get rendered due to their size.
Instances For
Equations
- Lean.Meta.Simp.removeUnnecessaryCasts.isDummyEqRec e = ((e.isAppOfArity `Eq.rec 6 || e.isAppOfArity `Eq.ndrec 6) && e.appArg!.isAppOf `Eq.refl)
Instances For
Try to use automatically generated congruence theorems. See mkCongrSimp?.
Equations
- One or more equations did not get rendered due to their size.
Instances For
Equations
- (Lean.Meta.Simp.Step.visit r).addExtraArgs extraArgs = do let __do_lift ← r.addExtraArgs extraArgs pure (Lean.Meta.Simp.Step.visit __do_lift)
- (Lean.Meta.Simp.Step.done r).addExtraArgs extraArgs = do let __do_lift ← r.addExtraArgs extraArgs pure (Lean.Meta.Simp.Step.done __do_lift)
- Lean.Meta.Simp.Step.continue.addExtraArgs extraArgs = pure Lean.Meta.Simp.Step.continue
- (Lean.Meta.Simp.Step.continue (some r)).addExtraArgs extraArgs = do let __do_lift ← r.addExtraArgs extraArgs pure (Lean.Meta.Simp.Step.continue (some __do_lift))
Instances For
Equations
- Lean.Meta.Simp.DStep.addExtraArgs (Lean.TransformStep.visit eNew) extraArgs = Lean.TransformStep.visit (Lean.mkAppN eNew extraArgs)
- Lean.Meta.Simp.DStep.addExtraArgs (Lean.TransformStep.done eNew) extraArgs = Lean.TransformStep.done (Lean.mkAppN eNew extraArgs)
- Lean.Meta.Simp.DStep.addExtraArgs Lean.TransformStep.continue extraArgs = Lean.TransformStep.continue
- Lean.Meta.Simp.DStep.addExtraArgs (Lean.TransformStep.continue (some eNew)) extraArgs = Lean.TransformStep.continue (some (Lean.mkAppN eNew extraArgs))
Instances For
Auxiliary method.
Given the current target of mvarId, apply r which is a new target and proof that it is equal to the current one.
Equations
- One or more equations did not get rendered due to their size.