ToExpr instances for Mathlib #
def
Mathlib.instToExprULift_mathlib.toExpr
{α✝ : Type s}
[Lean.ToExpr α✝]
:
Lean.ToLevel → Lean.ToLevel → ULift α✝ → Lean.Expr
Equations
- Mathlib.instToExprULift_mathlib.toExpr inst✝¹ inst✝ { down := a } = ((Lean.Expr.const `ULift.up [Lean.toLevel, Lean.toLevel]).app (Lean.toTypeExpr α✝)).app (Lean.toExpr a)
Instances For
instance
Mathlib.instToExprULift_mathlib
{α✝ : Type s}
[Lean.ToExpr α✝]
[Lean.ToLevel]
[Lean.ToLevel]
:
Lean.ToExpr (ULift α✝)
Equations
- Mathlib.instToExprULift_mathlib = { toExpr := Mathlib.instToExprULift_mathlib.toExpr inst✝¹ inst✝, toTypeExpr := (Lean.Expr.const `ULift [Lean.toLevel, Lean.toLevel]).app (Lean.toTypeExpr α✝) }
Hand-written instance since PUnit is a Sort rather than a Type.
Equations
- Mathlib.instToExprPUnitOfToLevel_mathlib = { toExpr := fun (x : PUnit) => Lean.mkConst `PUnit.unit [Lean.toLevel], toTypeExpr := Lean.mkConst `PUnit [Lean.toLevel] }
Equations
- Mathlib.instToExprRaw_mathlib = { toExpr := Mathlib.instToExprRaw_mathlib.toExpr, toTypeExpr := Lean.Expr.const `String.Pos.Raw [] }
Equations
- Mathlib.instToExprRaw_mathlib.toExpr { byteIdx := a } = (Lean.Expr.const `String.Pos.Raw.mk []).app (Lean.toExpr a)
Instances For
Equations
- Mathlib.instToExprRaw_mathlib_1 = { toExpr := Mathlib.instToExprRaw_mathlib_1.toExpr, toTypeExpr := Lean.Expr.const `Substring.Raw [] }
Equations
- Mathlib.instToExprRaw_mathlib_1.toExpr { str := a, startPos := a_1, stopPos := a_2 } = (((Lean.Expr.const `Substring.Raw.mk []).app (Lean.toExpr a)).app (Lean.toExpr a_1)).app (Lean.toExpr a_2)
Instances For
Equations
- One or more equations did not get rendered due to their size.
- Mathlib.instToExprSourceInfo_mathlib.toExpr Lean.SourceInfo.none = Lean.Expr.const `Lean.SourceInfo.none []
Instances For
Equations
- Mathlib.instToExprSourceInfo_mathlib = { toExpr := Mathlib.instToExprSourceInfo_mathlib.toExpr, toTypeExpr := Lean.Expr.const `Lean.SourceInfo [] }
Equations
- Mathlib.instToExprSyntax_mathlib = { toExpr := Mathlib.instToExprSyntax_mathlib.toExpr, toTypeExpr := Lean.Expr.const `Lean.Syntax [] }
Equations
- Mathlib.instToExprMData_mathlib = { toExpr := Mathlib.toExprMData✝, toTypeExpr := Lean.mkConst `Lean.MData }
Equations
- Mathlib.instToExprMVarId_mathlib.toExpr { name := a } = (Lean.Expr.const `Lean.MVarId.mk []).app (Lean.toExpr a)
Instances For
Equations
- Mathlib.instToExprMVarId_mathlib = { toExpr := Mathlib.instToExprMVarId_mathlib.toExpr, toTypeExpr := Lean.Expr.const `Lean.MVarId [] }
Equations
- Mathlib.instToExprLevelMVarId_mathlib = { toExpr := Mathlib.instToExprLevelMVarId_mathlib.toExpr, toTypeExpr := Lean.Expr.const `Lean.LevelMVarId [] }
Equations
- Mathlib.instToExprLevelMVarId_mathlib.toExpr { name := a } = (Lean.Expr.const `Lean.LevelMVarId.mk []).app (Lean.toExpr a)
Instances For
Equations
- Mathlib.instToExprLevel_mathlib = { toExpr := Mathlib.instToExprLevel_mathlib.toExpr, toTypeExpr := Lean.Expr.const `Lean.Level [] }
Equations
- Mathlib.instToExprBinderInfo_mathlib = { toExpr := Mathlib.instToExprBinderInfo_mathlib.toExpr, toTypeExpr := Lean.Expr.const `Lean.BinderInfo [] }
Equations
- Mathlib.instToExprBinderInfo_mathlib.toExpr Lean.BinderInfo.default = Lean.Expr.const `Lean.BinderInfo.default []
- Mathlib.instToExprBinderInfo_mathlib.toExpr Lean.BinderInfo.implicit = Lean.Expr.const `Lean.BinderInfo.implicit []
- Mathlib.instToExprBinderInfo_mathlib.toExpr Lean.BinderInfo.strictImplicit = Lean.Expr.const `Lean.BinderInfo.strictImplicit []
- Mathlib.instToExprBinderInfo_mathlib.toExpr Lean.BinderInfo.instImplicit = Lean.Expr.const `Lean.BinderInfo.instImplicit []
Instances For
Equations
- Mathlib.instToExprExpr_mathlib = { toExpr := Mathlib.instToExprExpr_mathlib.toExpr, toTypeExpr := Lean.Expr.const `Lean.Expr [] }