The top-down analyzer is an optional preprocessor to the delaborator that aims to determine the minimal annotations necessary to ensure that the delaborated expression can be re-elaborated correctly. Currently, the top-down analyzer is neither sound nor complete: there may be edge-cases in which the expression can still not be re-elaborated correctly, and it may also add many annotations that are not strictly necessary.
Equations
- Lean.getPPAnalyze o = Lean.KVMap.get o Lean.pp.analyze.name Lean.pp.analyze.defValue
Equations
Equations
Equations
Equations
Equations
Equations
- Lean.getPPAnalyzeTrustId o = Lean.KVMap.get o Lean.pp.analyze.trustId.name Lean.pp.analyze.trustId.defValue
Equations
- Lean.getPPAnalyzeTrustCoe o = Lean.KVMap.get o Lean.pp.analyze.trustCoe.name Lean.pp.analyze.trustCoe.defValue
Equations
Equations
- Lean.getPPAnalyzeOmitMax o = Lean.KVMap.get o Lean.pp.analyze.omitMax.name Lean.pp.analyze.omitMax.defValue
Equations
Equations
Equations
- Lean.getPPAnalysisSkip o = Lean.KVMap.get o `pp.analysis.skip false
Equations
- Lean.getPPAnalysisHole o = Lean.KVMap.get o `pp.analysis.hole false
Equations
- Lean.getPPAnalysisNamedArg o = Lean.KVMap.get o `pp.analysis.namedArg false
Equations
- Lean.getPPAnalysisLetVarType o = Lean.KVMap.get o `pp.analysis.letVarType false
Equations
- Lean.getPPAnalysisNeedsType o = Lean.KVMap.get o `pp.analysis.needsType false
Equations
- Lean.getPPAnalysisBlockImplicit o = Lean.KVMap.get o `pp.analysis.blockImplicit false
Equations
- Lean.PrettyPrinter.Delaborator.returnsPi motive = Lean.Meta.lambdaTelescope motive fun x b => pure (Lean.Expr.isForall b)
Equations
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.PrettyPrinter.Delaborator.isFOLike motive = let f := Lean.Expr.getAppFn motive; pure (Lean.Expr.isFVar f || Lean.Expr.isConst f)
Equations
- Lean.PrettyPrinter.Delaborator.isIdLike arg = match arg with | Lean.Expr.lam a a_1 (Lean.Expr.bvar a_2 a_3) a_4 => true | x => false
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
- 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
- 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.
def
Lean.PrettyPrinter.Delaborator.TopDownAnalyze.isDefEqAssigning
(t : Lean.Expr)
(s : Lean.Expr)
:
Equations
- One or more equations did not get rendered due to their size.
Equations
- Lean.PrettyPrinter.Delaborator.TopDownAnalyze.isHigherOrder type = Lean.Meta.forallTelescopeReducing type fun xs b => pure (decide (Array.size xs > 0) && Lean.Expr.isSort b)
Equations
- Lean.PrettyPrinter.Delaborator.TopDownAnalyze.isFunLike e = do let a ← Lean.Meta.inferType e Lean.Meta.forallTelescopeReducing a fun xs x => pure (decide (Array.size xs > 0))
Equations
- Lean.PrettyPrinter.Delaborator.TopDownAnalyze.isSubstLike e = (Lean.Expr.isAppOfArity e `Eq.ndrec 6 || Lean.Expr.isAppOfArity e `Eq.rec 6)
Equations
- One or more equations did not get rendered due to their size.
Equations
- Lean.PrettyPrinter.Delaborator.TopDownAnalyze.nameNotRoundtrippable.containsNum (Lean.Name.str p a a_1) = Lean.PrettyPrinter.Delaborator.TopDownAnalyze.nameNotRoundtrippable.containsNum p
- Lean.PrettyPrinter.Delaborator.TopDownAnalyze.nameNotRoundtrippable.containsNum (Lean.Name.num a a_1 a_2) = true
- Lean.PrettyPrinter.Delaborator.TopDownAnalyze.nameNotRoundtrippable.containsNum Lean.Name.anonymous = false
Equations
- Lean.PrettyPrinter.Delaborator.TopDownAnalyze.mvarName mvar = do let a ← Lean.Meta.getMVarDecl (Lean.Expr.mvarId! mvar) pure a.userName
Equations
- One or more equations did not get rendered due to their size.
- Lean.PrettyPrinter.Delaborator.TopDownAnalyze.containsBadMax (Lean.Level.succ u a) = Lean.PrettyPrinter.Delaborator.TopDownAnalyze.containsBadMax u
- Lean.PrettyPrinter.Delaborator.TopDownAnalyze.containsBadMax _fun_discr = false
- knowsType : Bool
- knowsLevel : Bool
- inBottomUp : Bool
- parentIsApp : Bool
- subExpr : Lean.SubExpr
Equations
- One or more equations did not get rendered due to their size.
- annotations : Lean.PrettyPrinter.Delaborator.OptionsPerPos
@[inline]
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
- One or more equations did not get rendered due to their size.
def
Lean.PrettyPrinter.Delaborator.TopDownAnalyze.inspectOutParams
(arg : Lean.Expr)
(mvar : Lean.Expr)
:
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.
def
Lean.PrettyPrinter.Delaborator.TopDownAnalyze.withKnowing
{α : Type}
(knowsType : Bool)
(knowsLevel : Bool)
(x : Lean.PrettyPrinter.Delaborator.TopDownAnalyze.AnalyzeM α)
:
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.
def
Lean.PrettyPrinter.Delaborator.TopDownAnalyze.annotateBoolAt
(n : Lean.Name)
(pos : Lean.SubExpr.Pos)
:
Equations
- One or more equations did not get rendered due to their size.
Equations
- Lean.PrettyPrinter.Delaborator.TopDownAnalyze.annotateBool n = do let a ← Lean.PrettyPrinter.Delaborator.SubExpr.getPos Lean.PrettyPrinter.Delaborator.TopDownAnalyze.annotateBoolAt n a
@[inline]
Equations
- One or more equations did not get rendered due to their size.
partial def
Lean.PrettyPrinter.Delaborator.TopDownAnalyze.analyze
(parentIsApp : optParam Bool false)
:
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
- 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
- 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
- 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.
partial def
Lean.PrettyPrinter.Delaborator.TopDownAnalyze.analyzeAppStagedCore.applyFunBinderHeuristic.core
(args : Array Lean.Expr)
(mvars : Array Lean.Expr)
(bInfos : Array Lean.BinderInfo)
(argIdx : Nat)
(mvarType : Lean.Expr)
:
def
Lean.PrettyPrinter.Delaborator.TopDownAnalyze.analyzeAppStagedCore.annotateNamedArg
(n : Lean.Name)
:
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.