Documentation

Init.Prelude

@[inline]
def id {α : Sort u} (a : α) :
α
Equations
@[inline]
def Function.comp {α : Sort u} {β : Sort v} {δ : Sort w} (f : βδ) (g : αβ) :
αδ
Equations
@[inline]
def Function.const {α : Sort u} (β : Sort v) (a : α) :
βα
Equations
def inferInstance {α : Sort u} [i : α] :
α
Equations
  • inferInstance = i
def inferInstanceAs (α : Sort u) [i : α] :
α
Equations
inductive PUnit:
Sort u
@[inline]
abbrev Unit:
Type

An abbreviation for PUnit.{0}, its most common instantiation. This Type should be preferred over PUnit where possible to avoid unnecessary universe parameters.

Equations
@[matchPattern, inline]
Equations
unsafe axiom lcProof {α : Prop} :
α

Auxiliary unsafe constant used by the Compiler when erasing proofs from code.

unsafe axiom lcUnreachable {α : Sort u} :
α

Auxiliary unsafe constant used by the Compiler to mark unreachable code.

inductive True:
Prop
inductive False:
Prop
    inductive Empty:
    Type
      inductive PEmpty:
      Sort u
        def Not (a : Prop) :
        Prop
        Equations
        @[macroInline]
        def False.elim {C : Sort u} (h : False) :
        C
        Equations
        @[macroInline]
        def absurd {a : Prop} {b : Sort v} (h₁ : a) (h₂ : ¬a) :
        b
        Equations
        inductive Eq {α : Sort u_1} :
        ααProp
        • refl: ∀ {α : Sort u_1} (a : α), a = a
        @[matchPattern]
        def rfl {α : Sort u} {a : α} :
        a = a
        Equations
        @[simp]
        theorem id_eq {α : Sort u_1} (a : α) :
        id a = a
        theorem Eq.subst {α : Sort u} {motive : αProp} {a : α} {b : α} (h₁ : a = b) (h₂ : motive a) :
        motive b
        theorem Eq.symm {α : Sort u} {a : α} {b : α} (h : a = b) :
        b = a
        theorem Eq.trans {α : Sort u} {a : α} {b : α} {c : α} (h₁ : a = b) (h₂ : b = c) :
        a = c
        @[macroInline]
        def cast {α : Sort u} {β : Sort u} (h : α = β) (a : α) :
        β
        Equations
        theorem congrArg {α : Sort u} {β : Sort v} {a₁ : α} {a₂ : α} (f : αβ) (h : a₁ = a₂) :
        f a₁ = f a₂
        theorem congr {α : Sort u} {β : Sort v} {f₁ : αβ} {f₂ : αβ} {a₁ : α} {a₂ : α} (h₁ : f₁ = f₂) (h₂ : a₁ = a₂) :
        f₁ a₁ = f₂ a₂
        theorem congrFun {α : Sort u} {β : αSort v} {f : (x : α) → β x} {g : (x : α) → β x} (h : f = g) (a : α) :
        f a = g a
        inductive HEq {α : Sort u} :
        α{β : Sort u} → βProp
        • refl: ∀ {α : Sort u} (a : α), HEq a a
        @[matchPattern]
        def HEq.rfl {α : Sort u} {a : α} :
        HEq a a
        Equations
        theorem eq_of_heq {α : Sort u} {a : α} {a' : α} (h : HEq a a') :
        a = a'
        @[unbox]
        structure Prod (α : Type u) (β : Type v) :
        Type (max u v)
        • fst : α
        • snd : β
        structure PProd (α : Sort u) (β : Sort v) :
        Sort (max (max 1 u) v)
        • fst : α
        • snd : β

        Similar to Prod, but α and β can be propositions. We use this Type internally to automatically generate the brecOn recursor.

        structure MProd (α : Type u) (β : Type u) :
        Type u
        • fst : α
        • snd : β

        Similar to Prod, but α and β are in the same universe.

        structure And (a : Prop) (b : Prop) :
        Prop
        • intro :: (
          • left : a
          • right : b
        • )
        inductive Or (a : Prop) (b : Prop) :
        Prop
        • inl: ∀ {a b : Prop}, aa b
        • inr: ∀ {a b : Prop}, ba b
        theorem Or.intro_left {a : Prop} (b : Prop) (h : a) :
        a b
        theorem Or.intro_right {b : Prop} (a : Prop) (h : b) :
        a b
        theorem Or.elim {a : Prop} {b : Prop} {c : Prop} (h : a b) (left : ac) (right : bc) :
        c
        inductive Bool:
        Type
        structure Subtype {α : Sort u} (p : αProp) :
        Sort (max 1 u)
        • val : α
        • property : p val
        def optParam (α : Sort u) (default : α) :
        Sort u

        Gadget for optional parameter support.

        Equations
        def outParam (α : Sort u) :
        Sort u

        Gadget for marking output parameters in type classes.

        Equations
        def typedExpr (α : Sort u) (a : α) :
        α

        Auxiliary Declaration used to implement the notation (a : α)

        Equations
        def namedPattern {α : Sort u} (x : α) (a : α) (h : x = a) :
        α

        Auxiliary Declaration used to implement the named patterns x@h:p

        Equations
        @[neverExtract, extern lean_sorry]
        axiom sorryAx (α : Sort u) (synthetic : optParam Bool false) :
        α
        theorem eq_false_of_ne_true {b : Bool} :
        ¬b = trueb = false
        theorem eq_true_of_ne_false {b : Bool} :
        ¬b = falseb = true
        theorem ne_false_of_eq_true {b : Bool} :
        b = true¬b = false
        theorem ne_true_of_eq_false {b : Bool} :
        b = false¬b = true
        @[nospecialize]
        class Inhabited (α : Sort u) :
        Sort (max 1 u)
        • default : α
        Instances
        axiom Classical.choice {α : Sort u} :
        Nonempty αα
        def Nonempty.elim {α : Sort u} {p : Prop} (h₁ : Nonempty α) (h₂ : αp) :
        p
        Equations
        instance instNonempty {α : Sort u} [inst : Inhabited α] :
        Equations
        noncomputable def Classical.ofNonempty {α : Sort u} [inst : Nonempty α] :
        α
        Equations
        instance instNonemptyForAll (α : Sort u) {β : Sort v} [inst : Nonempty β] :
        Nonempty (αβ)
        Equations
        instance instNonemptyForAll_1 (α : Sort u) {β : αSort v} [inst : ∀ (a : α), Nonempty (β a)] :
        Nonempty ((a : α) → β a)
        Equations
        instance instInhabitedSort:
        Inhabited (Sort u)
        Equations
        instance instInhabitedForAll (α : Sort u) {β : Sort v} [inst : Inhabited β] :
        Inhabited (αβ)
        Equations
        instance instInhabitedForAll_1 (α : Sort u) {β : αSort v} [inst : (a : α) → Inhabited (β a)] :
        Inhabited ((a : α) → β a)
        Equations
        structure PLift (α : Sort u) :
        Type u
        • up :: (
          • down : α
        • )

        Universe lifting operation from Sort to Type

        theorem PLift.up_down {α : Sort u} (b : PLift α) :
        { down := b.down } = b
        theorem PLift.down_up {α : Sort u} (a : α) :
        { down := a }.down = a
        def NonemptyType:
        Type (u + 1)
        @[inline]
        abbrev NonemptyType.type (type : NonemptyType) :
        Type u
        Equations
        Equations
        structure ULift (α : Type s) :
        Type (max s r)
        • up :: (
          • down : α
        • )

        Universe lifting operation

        theorem ULift.up_down {α : Type u} (b : ULift α) :
        { down := b.down } = b
        theorem ULift.down_up {α : Type u} (a : α) :
        { down := a }.down = a
        class inductive Decidable (p : Prop) :
        Type
        Instances
        @[inlineIfReduce, nospecialize]
        def Decidable.decide (p : Prop) [h : Decidable p] :
        Equations
        @[inline]
        abbrev DecidablePred {α : Sort u} (r : αProp) :
        Sort (max 1 u)
        Equations
        @[inline]
        abbrev DecidableRel {α : Sort u} (r : ααProp) :
        Sort (max 1 u)
        Equations
        @[inline]
        abbrev DecidableEq (α : Sort u) :
        Sort (max 1 u)
        Equations
        def decEq {α : Sort u} [s : DecidableEq α] (a : α) (b : α) :
        Decidable (a = b)
        Equations
        theorem decide_eq_true {p : Prop} [s : Decidable p] :
        pdecide p = true
        theorem decide_eq_false {p : Prop} [inst : Decidable p] :
        ¬pdecide p = false
        theorem of_decide_eq_true {p : Prop} [s : Decidable p] :
        decide p = truep
        theorem of_decide_eq_false {p : Prop} [s : Decidable p] :
        decide p = false¬p
        theorem of_decide_eq_self_eq_true {α : Sort u_1} [s : DecidableEq α] (a : α) :
        decide (a = a) = true
        @[inline]
        Equations
        • One or more equations did not get rendered due to their size.
        class BEq (α : Type u) :
        Type u
        • beq : ααBool
        Instances
        instance instBEq {α : Type u_1} [inst : DecidableEq α] :
        BEq α
        Equations
        • instBEq = { beq := fun a b => decide (a = b) }
        @[macroInline]
        def dite {α : Sort u} (c : Prop) [h : Decidable c] (t : cα) (e : ¬cα) :
        α
        Equations
        @[macroInline]
        def ite {α : Sort u} (c : Prop) [h : Decidable c] (t : α) (e : α) :
        α
        Equations
        @[macroInline]
        instance instDecidableAnd {p : Prop} {q : Prop} [dp : Decidable p] [dq : Decidable q] :
        Equations
        • One or more equations did not get rendered due to their size.
        @[macroInline]
        instance instDecidableOr {p : Prop} {q : Prop} [dp : Decidable p] [dq : Decidable q] :
        Equations
        instance instDecidableNot {p : Prop} [dp : Decidable p] :
        Equations
        @[macroInline]
        def cond {α : Type u} (c : Bool) (x : α) (y : α) :
        α
        Equations
        • (bif c then x else y) = match c with | true => x | false => y
        @[macroInline]
        def or (x : Bool) (y : Bool) :
        Equations
        @[macroInline]
        def and (x : Bool) (y : Bool) :
        Equations
        @[inline]
        def not:
        Equations
        inductive Nat:
        Type
        @[defaultInstance 100]
        instance instOfNatNat (n : Nat) :
        Equations
        def GE.ge {α : Type u} [inst : LE α] (a : α) (b : α) :
        Prop
        Equations
        def GT.gt {α : Type u} [inst : LT α] (a : α) (b : α) :
        Prop
        Equations
        • (a > b) = (b < a)
        @[inline]
        def max {α : Type u_1} [inst : LT α] [inst : DecidableRel LT.lt] (a : α) (b : α) :
        α
        Equations
        • max a b = if b < a then a else b
        @[inline]
        def min {α : Type u_1} [inst : LE α] [inst : DecidableRel LE.le] (a : α) (b : α) :
        α
        Equations
        class Trans {α : Sort u_1} {β : Sort u_2} {γ : Sort u_3} (r : αβProp) (s : βγProp) (t : outParam (αγProp)) :
        Type
        • trans : {a : α} → {b : β} → {c : γ} → r a bs b ct a c

        Transitive chaining of proofs, used e.g. by calc.

        Instances
        instance instTransEq {α : Sort u_1} {γ : Sort u_2} (r : αγProp) :
        Trans Eq r r
        Equations
        instance instTransEq_1 {α : Sort u_1} {β : Sort u_2} (r : αβProp) :
        Trans r Eq r
        Equations
        class HAdd (α : Type u) (β : Type v) (γ : outParam (Type w)) :
        Type (max (max u v) w)
        • hAdd : αβγ
        Instances
        class HSub (α : Type u) (β : Type v) (γ : outParam (Type w)) :
        Type (max (max u v) w)
        • hSub : αβγ
        Instances
        class HMul (α : Type u) (β : Type v) (γ : outParam (Type w)) :
        Type (max (max u v) w)
        • hMul : αβγ
        Instances
        class HDiv (α : Type u) (β : Type v) (γ : outParam (Type w)) :
        Type (max (max u v) w)
        • hDiv : αβγ
        Instances
        class HMod (α : Type u) (β : Type v) (γ : outParam (Type w)) :
        Type (max (max u v) w)
        • hMod : αβγ
        Instances
        class HPow (α : Type u) (β : Type v) (γ : outParam (Type w)) :
        Type (max (max u v) w)
        • hPow : αβγ
        Instances
        class HAppend (α : Type u) (β : Type v) (γ : outParam (Type w)) :
        Type (max (max u v) w)
        • hAppend : αβγ
        Instances
        class HOrElse (α : Type u) (β : Type v) (γ : outParam (Type w)) :
        Type (max (max u v) w)
        • hOrElse : α(Unitβ) → γ
        Instances
        class HAndThen (α : Type u) (β : Type v) (γ : outParam (Type w)) :
        Type (max (max u v) w)
        • hAndThen : α(Unitβ) → γ
        Instances
        class HAnd (α : Type u) (β : Type v) (γ : outParam (Type w)) :
        Type (max (max u v) w)
        • hAnd : αβγ
        Instances
        class HXor (α : Type u) (β : Type v) (γ : outParam (Type w)) :
        Type (max (max u v) w)
        • hXor : αβγ
        Instances
        class HOr (α : Type u) (β : Type v) (γ : outParam (Type w)) :
        Type (max (max u v) w)
        • hOr : αβγ
        Instances
        class HShiftLeft (α : Type u) (β : Type v) (γ : outParam (Type w)) :
        Type (max (max u v) w)
        • hShiftLeft : αβγ
        Instances
        class HShiftRight (α : Type u) (β : Type v) (γ : outParam (Type w)) :
        Type (max (max u v) w)
        • hShiftRight : αβγ
        Instances
        class Neg (α : Type u) :
        Type u
        • neg : αα
        Instances
        class Pow (α : Type u) (β : Type v) :
        Type (max u v)
        • pow : αβα
        Instances
        @[defaultInstance 1000]
        instance instHAdd {α : Type u_1} [inst : Add α] :
        HAdd α α α
        Equations
        • instHAdd = { hAdd := fun a b => Add.add a b }
        @[defaultInstance 1000]
        instance instHSub {α : Type u_1} [inst : Sub α] :
        HSub α α α
        Equations
        • instHSub = { hSub := fun a b => Sub.sub a b }
        @[defaultInstance 1000]
        instance instHMul {α : Type u_1} [inst : Mul α] :
        HMul α α α
        Equations
        • instHMul = { hMul := fun a b => Mul.mul a b }
        @[defaultInstance 1000]
        instance instHDiv {α : Type u_1} [inst : Div α] :
        HDiv α α α
        Equations
        • instHDiv = { hDiv := fun a b => Div.div a b }
        @[defaultInstance 1000]
        instance instHMod {α : Type u_1} [inst : Mod α] :
        HMod α α α
        Equations
        • instHMod = { hMod := fun a b => Mod.mod a b }
        @[defaultInstance 1000]
        instance instHPow {α : Type u_1} {β : Type u_2} [inst : Pow α β] :
        HPow α β α
        Equations
        • instHPow = { hPow := fun a b => Pow.pow a b }
        @[defaultInstance 1000]
        instance instHAppend {α : Type u_1} [inst : Append α] :
        HAppend α α α
        Equations
        @[defaultInstance 1000]
        instance instHOrElse {α : Type u_1} [inst : OrElse α] :
        HOrElse α α α
        Equations
        @[defaultInstance 1000]
        instance instHAndThen {α : Type u_1} [inst : AndThen α] :
        HAndThen α α α
        Equations
        @[defaultInstance 1000]
        instance instHAnd {α : Type u_1} [inst : AndOp α] :
        HAnd α α α
        Equations
        @[defaultInstance 1000]
        instance instHXor {α : Type u_1} [inst : Xor α] :
        HXor α α α
        Equations
        • instHXor = { hXor := fun a b => Xor.xor a b }
        @[defaultInstance 1000]
        instance instHOr {α : Type u_1} [inst : OrOp α] :
        HOr α α α
        Equations
        • instHOr = { hOr := fun a b => OrOp.or a b }
        @[defaultInstance 1000]
        instance instHShiftLeft {α : Type u_1} [inst : ShiftLeft α] :
        HShiftLeft α α α
        Equations
        @[defaultInstance 1000]
        instance instHShiftRight {α : Type u_1} [inst : ShiftRight α] :
        HShiftRight α α α
        Equations
        class Membership (α : outParam (Type u)) (γ : Type v) :
        Type (max u v)
        • mem : αγProp
        Instances
        @[matchPattern, extern lean_nat_add]
        def Nat.add:
        NatNatNat
        Equations
        instance instAddNat:
        Equations
        @[extern lean_nat_mul]
        def Nat.mul:
        NatNatNat
        Equations
        instance instMulNat:
        Equations
        @[extern lean_nat_pow]
        def Nat.pow (m : Nat) :
        NatNat
        Equations
        Equations
        instance instBEqNat:
        Equations
        theorem Nat.eq_of_beq_eq_true {n : Nat} {m : Nat} :
        Nat.beq n m = truen = m
        theorem Nat.ne_of_beq_eq_false {n : Nat} {m : Nat} :
        Nat.beq n m = false¬n = m
        @[extern lean_nat_dec_eq]
        def Nat.decEq (n : Nat) (m : Nat) :
        Decidable (n = m)
        Equations
        inductive Nat.le (n : Nat) :
        NatProp
        instance instLENat:
        Equations
        def Nat.lt (n : Nat) (m : Nat) :
        Prop
        Equations
        instance instLTNat:
        Equations
        theorem Nat.not_lt_zero (n : Nat) :
        ¬n < 0
        @[simp]
        theorem Nat.zero_le (n : Nat) :
        0 n
        theorem Nat.succ_le_succ {n : Nat} {m : Nat} :
        n mNat.succ n Nat.succ m
        theorem Nat.le_step {n : Nat} {m : Nat} (h : n m) :
        theorem Nat.le_trans {n : Nat} {m : Nat} {k : Nat} :
        n mm kn k
        theorem Nat.lt_trans {n : Nat} {m : Nat} {k : Nat} (h₁ : n < m) :
        m < kn < k
        theorem Nat.le_succ (n : Nat) :
        theorem Nat.le_succ_of_le {n : Nat} {m : Nat} (h : n m) :
        @[simp]
        theorem Nat.le_refl (n : Nat) :
        n n
        theorem Nat.succ_pos (n : Nat) :
        @[extern c inline "lean_nat_sub(#1, lean_box(1))"]
        def Nat.pred:
        NatNat
        Equations
        theorem Nat.pred_le_pred {n : Nat} {m : Nat} :
        n mNat.pred n Nat.pred m
        theorem Nat.le_of_succ_le_succ {n : Nat} {m : Nat} :
        Nat.succ n Nat.succ mn m
        theorem Nat.le_of_lt_succ {m : Nat} {n : Nat} :
        m < Nat.succ nm n
        theorem Nat.eq_or_lt_of_le {n : Nat} {m : Nat} :
        n mn = m n < m
        theorem Nat.lt_or_ge (n : Nat) (m : Nat) :
        n < m n m
        @[simp]
        theorem Nat.lt_irrefl (n : Nat) :
        ¬n < n
        theorem Nat.lt_of_le_of_lt {n : Nat} {m : Nat} {k : Nat} (h₁ : n m) (h₂ : m < k) :
        n < k
        theorem Nat.le_antisymm {n : Nat} {m : Nat} (h₁ : n m) (h₂ : m n) :
        n = m
        theorem Nat.lt_of_le_of_ne {n : Nat} {m : Nat} (h₁ : n m) (h₂ : ¬n = m) :
        n < m
        theorem Nat.le_of_ble_eq_true {n : Nat} {m : Nat} (h : Nat.ble n m = true) :
        n m
        theorem Nat.ble_succ_eq_true {n : Nat} {m : Nat} :
        theorem Nat.ble_eq_true_of_le {n : Nat} {m : Nat} (h : n m) :
        theorem Nat.not_le_of_not_ble_eq_true {n : Nat} {m : Nat} (h : ¬Nat.ble n m = true) :
        ¬n m
        @[extern lean_nat_dec_le]
        instance Nat.decLe (n : Nat) (m : Nat) :
        Equations
        @[extern lean_nat_dec_lt]
        instance Nat.decLt (n : Nat) (m : Nat) :
        Decidable (n < m)
        Equations
        @[extern lean_nat_sub]
        def Nat.sub:
        NatNatNat
        Equations
        instance instSubNat:
        Equations
        structure Fin (n : Nat) :
        Type
        • val : Nat
        • isLt : val < n
        theorem Fin.eq_of_val_eq {n : Nat} {i : Fin n} {j : Fin n} :
        i.val = j.vali = j
        theorem Fin.val_eq_of_eq {n : Nat} {i : Fin n} {j : Fin n} (h : i = j) :
        i.val = j.val
        theorem Fin.ne_of_val_ne {n : Nat} {i : Fin n} {j : Fin n} (h : ¬i.val = j.val) :
        ¬i = j
        Equations
        instance instLTFin {n : Nat} :
        LT (Fin n)
        Equations
        • instLTFin = { lt := fun a b => a.val < b.val }
        instance instLEFin {n : Nat} :
        LE (Fin n)
        Equations
        • instLEFin = { le := fun a b => a.val b.val }
        instance Fin.decLt {n : Nat} (a : Fin n) (b : Fin n) :
        Decidable (a < b)
        Equations
        instance Fin.decLe {n : Nat} (a : Fin n) (b : Fin n) :
        Equations
        structure UInt8:
        Type
        @[extern lean_uint8_of_nat]
        def UInt8.ofNatCore (n : Nat) (h : n < UInt8.size) :
        Equations
        @[extern lean_uint8_dec_eq]
        def UInt8.decEq (a : UInt8) (b : UInt8) :
        Decidable (a = b)
        Equations
        • UInt8.decEq a b = match a, b with | { val := n }, { val := m } => if h : n = m then isTrue (_ : { val := n } = { val := m }) else isFalse (_ : { val := n } = { val := m }False)
        structure UInt16:
        Type
        @[extern lean_uint16_of_nat]
        Equations
        @[extern lean_uint16_dec_eq]
        def UInt16.decEq (a : UInt16) (b : UInt16) :
        Decidable (a = b)
        Equations
        • UInt16.decEq a b = match a, b with | { val := n }, { val := m } => if h : n = m then isTrue (_ : { val := n } = { val := m }) else isFalse (_ : { val := n } = { val := m }False)
        Equations
        structure UInt32:
        Type
        @[extern lean_uint32_of_nat]
        Equations
        @[extern lean_uint32_to_nat]
        Equations
        @[extern lean_uint32_dec_eq]
        def UInt32.decEq (a : UInt32) (b : UInt32) :
        Decidable (a = b)
        Equations
        • UInt32.decEq a b = match a, b with | { val := n }, { val := m } => if h : n = m then isTrue (_ : { val := n } = { val := m }) else isFalse (_ : { val := n } = { val := m }False)
        Equations
        Equations
        @[extern lean_uint32_dec_lt]
        def UInt32.decLt (a : UInt32) (b : UInt32) :
        Decidable (a < b)
        Equations
        @[extern lean_uint32_dec_le]
        def UInt32.decLe (a : UInt32) (b : UInt32) :
        Equations
        Equations
        structure UInt64:
        Type
        @[extern lean_uint64_of_nat]
        Equations
        @[extern lean_uint64_dec_eq]
        def UInt64.decEq (a : UInt64) (b : UInt64) :
        Decidable (a = b)
        Equations
        • UInt64.decEq a b = match a, b with | { val := n }, { val := m } => if h : n = m then isTrue (_ : { val := n } = { val := m }) else isFalse (_ : { val := n } = { val := m }False)
        theorem usize_size_eq:
        USize.size = 4294967296 USize.size = 18446744073709551616
        structure USize:
        Type
        @[extern lean_usize_of_nat]
        def USize.ofNatCore (n : Nat) (h : n < USize.size) :
        Equations
        @[extern lean_usize_dec_eq]
        def USize.decEq (a : USize) (b : USize) :
        Decidable (a = b)
        Equations
        • USize.decEq a b = match a, b with | { val := n }, { val := m } => if h : n = m then isTrue (_ : { val := n } = { val := m }) else isFalse (_ : { val := n } = { val := m }False)
        @[extern lean_usize_of_nat]
        def USize.ofNat32 (n : Nat) (h : n < 4294967296) :
        Equations
        @[inline]
        abbrev Nat.isValidChar (n : Nat) :
        Prop
        Equations
        structure Char:
        Type

        The Char Type represents an unicode scalar value. See http://www.unicode.org/glossary/#unicode_scalar_value).

        @[extern lean_uint32_of_nat]
        Equations
        @[matchPattern, noinline]
        def Char.ofNat (n : Nat) :
        Equations
        theorem Char.eq_of_val_eq {c : Char} {d : Char} :
        c.val = d.valc = d
        theorem Char.val_eq_of_eq {c : Char} {d : Char} :
        c = dc.val = d.val
        theorem Char.ne_of_val_ne {c : Char} {d : Char} (h : ¬c.val = d.val) :
        ¬c = d
        theorem Char.val_ne_of_ne {c : Char} {d : Char} (h : ¬c = d) :
        ¬c.val = d.val
        Equations
        Equations
        • One or more equations did not get rendered due to their size.
        @[unbox]
        inductive Option (α : Type u) :
        Type u
        • none: {α : Type u} → Option α
        • some: {α : Type u} → αOption α
        instance instInhabitedOption {α : Type u_1} :
        Equations
        • instInhabitedOption = { default := none }
        @[macroInline]
        def Option.getD {α : Type u_1} :
        Option ααα
        Equations
        • Option.getD _fun_discr _fun_discr = match _fun_discr, _fun_discr with | some x, x_1 => x | none, e => e
        @[inline]
        def Option.map {α : Type u_1} {β : Type u_2} (f : αβ) :
        Option αOption β
        Equations
        inductive List (α : Type u) :
        Type u
        • nil: {α : Type u} → List α
        • cons: {α : Type u} → αList αList α
        instance instInhabitedList {α : Type u_1} :
        Equations
        • instInhabitedList = { default := [] }
        def List.hasDecEq {α : Type u} [inst : DecidableEq α] (a : List α) (b : List α) :
        Decidable (a = b)
        Equations
        instance instDecidableEqList {α : Type u} [inst : DecidableEq α] :
        Equations
        • instDecidableEqList = List.hasDecEq
        @[specialize]
        def List.foldl {α : Sort u_1} {β : Type u_2} (f : αβα) (init : α) :
        List βα
        Equations
        def List.set {α : Type u_1} :
        List αNatαList α
        Equations
        def List.length {α : Type u_1} :
        List αNat
        Equations
        def List.lengthTRAux {α : Type u_1} :
        List αNatNat
        Equations
        def List.lengthTR {α : Type u_1} (as : List α) :
        Equations
        @[simp]
        theorem List.length_cons {α : Type u_1} (a : α) (as : List α) :
        def List.concat {α : Type u} :
        List ααList α
        Equations
        def List.get {α : Type u} (as : List α) :
        Fin (List.length as)α
        Equations
        structure String:
        Type
        @[extern lean_string_dec_eq]
        def String.decEq (s₁ : String) (s₂ : String) :
        Decidable (s₁ = s₂)
        Equations
        • One or more equations did not get rendered due to their size.
        structure String.Pos:
        Type

        A byte position in a String. Internally, Strings are UTF-8 encoded. Codepoint positions (counting the Unicode codepoints rather than bytes) are represented by plain Nats instead. Indexing a String by a byte position is constant-time, while codepoint positions need to be translated internally to byte positions in linear-time.

        Equations
        Equations
        • One or more equations did not get rendered due to their size.
        structure Substring:
        Type
        Equations
        @[inline]
        Equations
        • Substring.bsize _fun_discr = match _fun_discr with | { str := str, startPos := b, stopPos := e } => Nat.sub e.byteIdx b.byteIdx
        @[extern lean_string_utf8_byte_size]
        Equations
        Equations
        • instHAddPos = { hAdd := fun p₁ p₂ => { byteIdx := p₁.byteIdx + p₂.byteIdx } }
        Equations
        • instHSubPos = { hSub := fun p₁ p₂ => { byteIdx := p₁.byteIdx - p₂.byteIdx } }
        Equations
        Equations
        Equations
        • instLEPos = { le := fun p₁ p₂ => p₁.byteIdx p₂.byteIdx }
        Equations
        • instLTPos = { lt := fun p₁ p₂ => p₁.byteIdx < p₂.byteIdx }
        instance instDecidableLePosInstLEPos (p₁ : String.Pos) (p₂ : String.Pos) :
        Decidable (p₁ p₂)
        Equations
        instance instDecidableLtPosInstLTPos (p₁ : String.Pos) (p₂ : String.Pos) :
        Decidable (p₁ < p₂)
        Equations
        @[inline]
        Equations
        @[inline]
        Equations
        unsafe def unsafeCast {α : Type u} {β : Type v} (a : α) :
        β
        Equations
        @[neverExtract, extern lean_panic_fn]
        opaque panicCore {α : Type u} [inst : Inhabited α] (msg : String) :
        α
        @[neverExtract, noinline]
        def panic {α : Type u} [inst : Inhabited α] (msg : String) :
        α
        Equations
        structure Array (α : Type u) :
        Type u
        @[extern lean_mk_empty_array_with_capacity]
        def Array.mkEmpty {α : Type u} (c : Nat) :
        Equations
        def Array.empty {α : Type u} :
        Equations
        @[extern lean_array_get_size]
        def Array.size {α : Type u} (a : Array α) :
        Equations
        @[extern lean_array_fget]
        def Array.get {α : Type u} (a : Array α) (i : Fin (Array.size a)) :
        α
        Equations
        @[inline]
        def Array.getD {α : Type u_1} (a : Array α) (i : Nat) (v₀ : α) :
        α
        Equations
        @[extern lean_array_get]
        def Array.get! {α : Type u} [inst : Inhabited α] (a : Array α) (i : Nat) :
        α
        Equations
        def Array.getOp {α : Type u} [inst : Inhabited α] (self : Array α) (idx : Nat) :
        α
        Equations
        @[extern lean_array_push]
        def Array.push {α : Type u} (a : Array α) (v : α) :
        Equations
        @[extern lean_array_fset]
        def Array.set {α : Type u_1} (a : Array α) (i : Fin (Array.size a)) (v : α) :
        Equations
        @[inline]
        def Array.setD {α : Type u_1} (a : Array α) (i : Nat) (v : α) :
        Equations
        @[extern lean_array_set]
        def Array.set! {α : Type u_1} (a : Array α) (i : Nat) (v : α) :
        Equations
        def Array.appendCore {α : Type u} (as : Array α) (bs : Array α) :
        Equations
        def Array.appendCore.loop {α : Type u} (bs : Array α) (i : Nat) (j : Nat) (as : Array α) :
        Equations
        • One or more equations did not get rendered due to their size.
        @[inlineIfReduce]
        def List.toArrayAux {α : Type u_1} :
        List αArray αArray α
        Equations
        @[inlineIfReduce]
        def List.redLength {α : Type u_1} :
        List αNat
        Equations
        @[matchPattern, inline, export lean_list_to_array]
        def List.toArray {α : Type u_1} (as : List α) :
        Equations
        class Bind (m : Type uType v) :
        Type (max (u + 1) v)
        • bind : {α β : Type u} → m α(αm β) → m β
        Instances
        class Pure (f : Type uType v) :
        Type (max (u + 1) v)
        • pure : {α : Type u} → αf α
        Instances
        class Functor (f : Type uType v) :
        Type (max (u + 1) v)
        • map : {α β : Type u} → (αβ) → f αf β
        • mapConst : {α β : Type u} → αf βf α
        Instances
        class Seq (f : Type uType v) :
        Type (max (u + 1) v)
        • seq : {α β : Type u} → f (αβ)(Unitf α) → f β
        Instances
        class SeqLeft (f : Type uType v) :
        Type (max (u + 1) v)
        • seqLeft : {α β : Type u} → f α(Unitf β) → f α
        Instances
        class SeqRight (f : Type uType v) :
        Type (max (u + 1) v)
        • seqRight : {α β : Type u} → f α(Unitf β) → f β
        Instances
        class Applicative (f : Type uType v) extends Functor , Pure , Seq , SeqLeft , SeqRight :
        Type (max (u + 1) v)
          Instances
          instance instInhabitedForAll_2 {α : Type u} {m : Type uType v} [inst : Monad m] :
          Inhabited (αm α)
          Equations
          • instInhabitedForAll_2 = { default := pure }
          instance instInhabited {α : Type u} {m : Type uType v} [inst : Monad m] [inst : Inhabited α] :
          Inhabited (m α)
          Equations
          • instInhabited = { default := pure default }
          def Array.sequenceMap {α : Type u} {β : Type v} {m : Type vType w} [inst : Monad m] (as : Array α) (f : αm β) :
          m (Array β)
          Equations
          def Array.sequenceMap.loop {α : Type u} {β : Type v} {m : Type vType w} [inst : Monad m] (as : Array α) (f : αm β) (i : Nat) (j : Nat) (bs : Array β) :
          m (Array β)
          Equations
          • One or more equations did not get rendered due to their size.
          class MonadLiftT (m : Type uType v) (n : Type uType w) :
          Type (max (max (u + 1) v) w)
          • monadLift : {α : Type u} → m αn α

          The reflexive-transitive closure of MonadLift. monadLift is used to transitively lift monadic computations such as StateT.get or StateT.put s. Corresponds to MonadLift.

          Instances
          @[inline]
          abbrev liftM {m : Type u_1Type u_2} {n : Type u_1Type u_3} [self : MonadLiftT m n] {α : Type u_1} :
          m αn α
          Equations
          instance instMonadLiftT (m : Type u_1Type u_2) (n : Type u_1Type u_3) (o : Type u_1Type u_4) [inst : MonadLift n o] [inst : MonadLiftT m n] :
          Equations
          instance instMonadLiftT_1 (m : Type u_1Type u_2) :
          Equations
          class MonadFunctor (m : Type uType v) (n : Type uType w) :
          Type (max (max (u + 1) v) w)
          • monadMap : {α : Type u} → ({β : Type u} → m βm β) → n αn α

          A functor in the category of monads. Can be used to lift monad-transforming functions. Based on pipes' MFunctor, but not restricted to monad transformers. Alternatively, an implementation of MonadTransFunctor.

          Instances
          class MonadFunctorT (m : Type uType v) (n : Type uType w) :
          Type (max (max (u + 1) v) w)
          • monadMap : {α : Type u} → ({β : Type u} → m βm β) → n αn α

          The reflexive-transitive closure of MonadFunctor. monadMap is used to transitively lift Monad morphisms

          Instances
          instance instMonadFunctorT (m : Type u_1Type u_2) (n : Type u_1Type u_3) (o : Type u_1Type u_4) [inst : MonadFunctor n o] [inst : MonadFunctorT m n] :
          Equations
          instance monadFunctorRefl (m : Type u_1Type u_2) :
          Equations
          @[unbox]
          inductive Except (ε : Type u) (α : Type v) :
          Type (max u v)
          • error: {ε : Type u} → {α : Type v} → εExcept ε α
          • ok: {ε : Type u} → {α : Type v} → αExcept ε α
          instance instInhabitedExcept {ε : Type u} {α : Type v} [inst : Inhabited ε] :
          Equations
          @[inline]
          abbrev throwThe (ε : Type u) {m : Type vType w} [inst : MonadExceptOf ε m] {α : Type v} (e : ε) :
          m α
          Equations
          @[inline]
          abbrev tryCatchThe (ε : Type u) {m : Type vType w} [inst : MonadExceptOf ε m] {α : Type v} (x : m α) (handle : εm α) :
          m α
          Equations
          class MonadExcept (ε : outParam (Type u)) (m : Type vType w) :
          Type (max (max u (v + 1)) w)
          • throw : {α : Type v} → εm α
          • tryCatch : {α : Type v} → m α(εm α) → m α

          Similar to MonadExceptOf, but ε is an outParam for convenience

          Instances
          def MonadExcept.ofExcept {m : Type u_1Type u_2} {ε : Type u_3} {α : Type u_1} [inst : Monad m] [inst : MonadExcept ε m] :
          Except ε αm α
          Equations
          instance instMonadExcept (ε : outParam (Type u)) (m : Type vType w) [inst : MonadExceptOf ε m] :
          Equations
          @[inline]
          def MonadExcept.orElse {ε : Type u} {m : Type vType w} [inst : MonadExcept ε m] {α : Type v} (t₁ : m α) (t₂ : Unitm α) :
          m α
          Equations
          instance MonadExcept.instOrElse {ε : Type u} {m : Type vType w} [inst : MonadExcept ε m] {α : Type v} :
          OrElse (m α)
          Equations
          • MonadExcept.instOrElse = { orElse := MonadExcept.orElse }
          def ReaderT (ρ : Type u) (m : Type uType v) (α : Type u) :
          Type (max u v)

          An implementation of ReaderT

          Equations
          instance instInhabitedReaderT (ρ : Type u) (m : Type uType v) (α : Type u) [inst : Inhabited (m α)] :
          Inhabited (ReaderT ρ m α)
          Equations
          @[inline]
          def ReaderT.run {ρ : Type u} {m : Type uType v} {α : Type u} (x : ReaderT ρ m α) (r : ρ) :
          m α
          Equations
          instance ReaderT.instMonadLiftReaderT {ρ : Type u} {m : Type uType v} :
          Equations
          • ReaderT.instMonadLiftReaderT = { monadLift := fun {α} x x_1 => x }
          instance ReaderT.instMonadExceptOfReaderT {ρ : Type u} {m : Type uType v} (ε : Type u_1) [inst : MonadExceptOf ε m] :
          Equations
          @[inline]
          def ReaderT.read {ρ : Type u} {m : Type uType v} [inst : Monad m] :
          ReaderT ρ m ρ
          Equations
          • ReaderT.read = pure
          @[inline]
          def ReaderT.pure {ρ : Type u} {m : Type uType v} [inst : Monad m] {α : Type u} (a : α) :
          ReaderT ρ m α
          Equations
          @[inline]
          def ReaderT.bind {ρ : Type u} {m : Type uType v} [inst : Monad m] {α : Type u} {β : Type u} (x : ReaderT ρ m α) (f : αReaderT ρ m β) :
          ReaderT ρ m β
          Equations
          @[inline]
          def ReaderT.map {ρ : Type u} {m : Type uType v} [inst : Monad m] {α : Type u} {β : Type u} (f : αβ) (x : ReaderT ρ m α) :
          ReaderT ρ m β
          Equations
          instance ReaderT.instMonadReaderT {ρ : Type u} {m : Type uType v} [inst : Monad m] :
          Monad (ReaderT ρ m)
          Equations
          • ReaderT.instMonadReaderT = Monad.mk
          instance ReaderT.instMonadFunctorReaderT (ρ : Type u_1) (m : Type u_1Type u_2) [inst : Monad m] :
          Equations
          @[inline]
          def ReaderT.adapt {ρ : Type u} {m : Type uType v} {ρ' : Type u} [inst : Monad m] {α : Type u} (f : ρ'ρ) :
          ReaderT ρ m αReaderT ρ' m α
          Equations
          class MonadReaderOf (ρ : Type u) (m : Type uType v) :
          Type v
          • read : m ρ

          An implementation of MonadReader. It does not contain local because this Function cannot be lifted using monadLift. Instead, the MonadReaderAdapter class provides the more general adaptReader Function.

          Note: This class can be seen as a simplification of the more "principled" definition
          ```
          class MonadReader (ρ : outParam (Type u)) (n : Type u → Type u) where
            lift {α : Type u} : ({m : Type u → Type u} → [Monad m] → ReaderT ρ m α) → n α
          ```
          
          Instances
          @[inline]
          def readThe (ρ : Type u) {m : Type uType v} [inst : MonadReaderOf ρ m] :
          m ρ
          Equations
          class MonadReader (ρ : outParam (Type u)) (m : Type uType v) :
          Type v
          • read : m ρ

          Similar to MonadReaderOf, but ρ is an outParam for convenience

          Instances
          instance instMonadReader (ρ : Type u) (m : Type uType v) [inst : MonadReaderOf ρ m] :
          Equations
          instance instMonadReaderOf {ρ : Type u} {m : Type uType v} {n : Type uType w} [inst : MonadLift m n] [inst : MonadReaderOf ρ m] :
          Equations
          • instMonadReaderOf = { read := liftM read }
          instance instMonadReaderOfReaderT {ρ : Type u} {m : Type uType v} [inst : Monad m] :
          Equations
          • instMonadReaderOfReaderT = { read := ReaderT.read }
          @[inline]
          def withTheReader (ρ : Type u) {m : Type uType v} [inst : MonadWithReaderOf ρ m] {α : Type u} (f : ρρ) (x : m α) :
          m α
          Equations
          class MonadWithReader (ρ : outParam (Type u)) (m : Type uType v) :
          Type (max (u + 1) v)
          • withReader : {α : Type u} → (ρρ) → m αm α
          Instances
          instance instMonadWithReader (ρ : Type u) (m : Type uType v) [inst : MonadWithReaderOf ρ m] :
          Equations
          instance instMonadWithReaderOf {ρ : Type u} {m : Type uType v} {n : Type uType v} [inst : MonadFunctor m n] [inst : MonadWithReaderOf ρ m] :
          Equations
          instance instMonadWithReaderOfReaderT {ρ : Type u} {m : Type uType v} [inst : Monad m] :
          Equations
          • instMonadWithReaderOfReaderT = { withReader := fun {α} f x ctx => x (f ctx) }
          class MonadStateOf (σ : Type u) (m : Type uType v) :
          Type (max (u + 1) v)
          • get : m σ
          • set : σm PUnit
          • modifyGet : {α : Type u} → (σα × σ) → m α

          An implementation of MonadState. In contrast to the Haskell implementation, we use overlapping instances to derive instances automatically from monadLift.

          Instances
          @[inline]
          abbrev getThe (σ : Type u) {m : Type uType v} [inst : MonadStateOf σ m] :
          m σ
          Equations
          @[inline]
          abbrev modifyThe (σ : Type u) {m : Type uType v} [inst : MonadStateOf σ m] (f : σσ) :
          Equations
          @[inline]
          abbrev modifyGetThe {α : Type u} (σ : Type u) {m : Type uType v} [inst : MonadStateOf σ m] (f : σα × σ) :
          m α
          Equations
          class MonadState (σ : outParam (Type u)) (m : Type uType v) :
          Type (max (u + 1) v)
          • get : m σ
          • set : σm PUnit
          • modifyGet : {α : Type u} → (σα × σ) → m α

          Similar to MonadStateOf, but σ is an outParam for convenience

          Instances
          instance instMonadState (σ : Type u) (m : Type uType v) [inst : MonadStateOf σ m] :
          Equations
          @[inline]
          def modify {σ : Type u} {m : Type uType v} [inst : MonadState σ m] (f : σσ) :
          Equations
          @[inline]
          def getModify {σ : Type u} {m : Type uType v} [inst : MonadState σ m] [inst : Monad m] (f : σσ) :
          m σ
          Equations
          instance instMonadStateOf {σ : Type u} {m : Type uType v} {n : Type uType w} [inst : MonadLift m n] [inst : MonadStateOf σ m] :
          Equations
          inductive EStateM.Result (ε : Type u) (σ : Type u) (α : Type u) :
          Type u
          instance EStateM.instInhabitedResult {ε : Type u} {σ : Type u} {α : Type u} [inst : Inhabited ε] [inst : Inhabited σ] :
          Equations
          def EStateM (ε : Type u) (σ : Type u) (α : Type u) :
          Type u
          Equations
          instance EStateM.instInhabitedEStateM {ε : Type u} {σ : Type u} {α : Type u} [inst : Inhabited ε] :
          Inhabited (EStateM ε σ α)
          Equations
          @[inline]
          def EStateM.pure {ε : Type u} {σ : Type u} {α : Type u} (a : α) :
          EStateM ε σ α
          Equations
          @[inline]
          def EStateM.set {ε : Type u} {σ : Type u} (s : σ) :
          Equations
          @[inline]
          def EStateM.get {ε : Type u} {σ : Type u} :
          EStateM ε σ σ
          Equations
          @[inline]
          def EStateM.modifyGet {ε : Type u} {σ : Type u} {α : Type u} (f : σα × σ) :
          EStateM ε σ α
          Equations
          @[inline]
          def EStateM.throw {ε : Type u} {σ : Type u} {α : Type u} (e : ε) :
          EStateM ε σ α
          Equations
          class EStateM.Backtrackable (δ : outParam (Type u)) (σ : Type u) :
          Type u
          • save : σδ
          • restore : σδσ

          Auxiliary instance for saving/restoring the "backtrackable" part of the state.

          Instances
          @[inline]
          def EStateM.tryCatch {ε : Type u} {σ : Type u} {δ : Type u} [inst : EStateM.Backtrackable δ σ] {α : Type u} (x : EStateM ε σ α) (handle : εEStateM ε σ α) :
          EStateM ε σ α
          Equations
          @[inline]
          def EStateM.orElse {ε : Type u} {σ : Type u} {α : Type u} {δ : Type u} [inst : EStateM.Backtrackable δ σ] (x₁ : EStateM ε σ α) (x₂ : UnitEStateM ε σ α) :
          EStateM ε σ α
          Equations
          @[inline]
          def EStateM.adaptExcept {ε : Type u} {σ : Type u} {α : Type u} {ε' : Type u} (f : εε') (x : EStateM ε σ α) :
          EStateM ε' σ α
          Equations
          @[inline]
          def EStateM.bind {ε : Type u} {σ : Type u} {α : Type u} {β : Type u} (x : EStateM ε σ α) (f : αEStateM ε σ β) :
          EStateM ε σ β
          Equations
          @[inline]
          def EStateM.map {ε : Type u} {σ : Type u} {α : Type u} {β : Type u} (f : αβ) (x : EStateM ε σ α) :
          EStateM ε σ β
          Equations
          @[inline]
          def EStateM.seqRight {ε : Type u} {σ : Type u} {α : Type u} {β : Type u} (x : EStateM ε σ α) (y : UnitEStateM ε σ β) :
          EStateM ε σ β
          Equations
          instance EStateM.instMonadEStateM {ε : Type u} {σ : Type u} :
          Monad (EStateM ε σ)
          Equations
          • EStateM.instMonadEStateM = Monad.mk
          instance EStateM.instOrElseEStateM {ε : Type u} {σ : Type u} {α : Type u} {δ : Type u} [inst : EStateM.Backtrackable δ σ] :
          OrElse (EStateM ε σ α)
          Equations
          • EStateM.instOrElseEStateM = { orElse := EStateM.orElse }
          instance EStateM.instMonadStateOfEStateM {ε : Type u} {σ : Type u} :
          Equations
          • EStateM.instMonadStateOfEStateM = { get := EStateM.get, set := EStateM.set, modifyGet := fun {α} => EStateM.modifyGet }
          instance EStateM.instMonadExceptOfEStateM {ε : Type u} {σ : Type u} {δ : Type u} [inst : EStateM.Backtrackable δ σ] :
          Equations
          • EStateM.instMonadExceptOfEStateM = { throw := fun {α} => EStateM.throw, tryCatch := fun {α} => EStateM.tryCatch }
          @[inline]
          def EStateM.run {ε : Type u} {σ : Type u} {α : Type u} (x : EStateM ε σ α) (s : σ) :
          Equations
          @[inline]
          def EStateM.run' {ε : Type u} {σ : Type u} {α : Type u} (x : EStateM ε σ α) (s : σ) :
          Equations
          @[inline]
          def EStateM.dummySave {σ : Type u} :
          σPUnit
          Equations
          @[inline]
          def EStateM.dummyRestore {σ : Type u} :
          σPUnitσ
          Equations
          Equations
          • EStateM.nonBacktrackable = { save := EStateM.dummySave, restore := EStateM.dummyRestore }
          @[extern lean_uint64_to_usize]
          @[extern lean_usize_to_uint64]
          @[extern lean_uint64_mix_hash]
          opaque mixHash (u₁ : UInt64) (u₂ : UInt64) :
          @[extern lean_string_hash]
          opaque String.hash (s : String) :
          inductive Lean.Name:
          Type
          Equations
          @[export lean_name_mk_string]
          Equations
          @[export lean_name_mk_numeral]
          Equations
          @[extern lean_name_eq]
          Equations
          inductive Lean.SourceInfo:
          Type

          Source information of tokens.

          Equations
          • One or more equations did not get rendered due to their size.
          inductive Lean.Syntax:
          Type

          Syntax objects used by the parser, macro expander, delaborator, etc.

          structure Lean.TSyntax (ks : Lean.SyntaxNodeKinds) :
          Type

          A Syntax value of one of the given syntax kinds. Note that while syntax quotations produce/expect TSyntax values of the correct kinds, this is not otherwise enforced and can easily be circumvented by direct use of the constructor. The namespace TSyntax.Compat can be opened to expose a general coercion from Syntax to any TSyntax ks for porting older code.

          Equations
          • Lean.instInhabitedTSyntax = { default := { raw := default } }
          Equations
          • One or more equations did not get rendered due to their size.
          Equations
          Equations
          @[inline]
          Equations
          Equations
          Equations
          Equations
          • One or more equations did not get rendered due to their size.
          Equations
          Equations
          Equations
          Equations
          Equations

          Retrieve the left-most node or leaf's info in the Syntax tree.

          Retrieve the left-most leaf's info in the Syntax tree, or none if there is no token.

          Equations
          structure Lean.Syntax.SepArray (sep : String) :
          Type

          An array of syntax elements interspersed with separators. Can be coerced to/from Array Syntax to automatically remove/insert the separators.

          structure Lean.Syntax.TSepArray (ks : Lean.SyntaxNodeKinds) (sep : String) :
          Type

          A typed version of SepArray.

          Equations
          • Lean.TSyntaxArray.rawImpl = unsafeCast
          @[implementedBy Lean.TSyntaxArray.rawImpl]
          Equations
          • Lean.TSyntaxArray.mkImpl = unsafeCast
          @[implementedBy Lean.TSyntaxArray.mkImpl]
          Equations
          • One or more equations did not get rendered due to their size.
          @[inline]
          abbrev Lean.MacroScope:
          Type
          Equations

          Macro scope used internally. It is not available for our frontend.

          Equations

          First macro scope available for our frontend

          Equations
          instance Lean.instMonadRef (m : TypeType) (n : TypeType) [inst : MonadLift m n] [inst : MonadFunctor m n] [inst : Lean.MonadRef m] :
          Equations
          Equations
          @[inline]
          def Lean.withRef {m : TypeType} [inst : Monad m] [inst : Lean.MonadRef m] {α : Type} (ref : Lean.Syntax) (x : m α) :
          m α
          Equations
          class Lean.MonadQuotation (m : TypeType) extends Lean.MonadRef :
          Type 1

          A monad that supports syntax quotations. Syntax quotations (in term position) are monadic values that when executed retrieve the current "macro scope" from the monad and apply it to every identifier they introduce (independent of whether this identifier turns out to be a reference to an existing declaration, or an actually fresh binding during further elaboration). We also apply the position of the result of getRef to each introduced symbol, which results in better error positions than not applying any position.

          Instances
          def Lean.MonadRef.mkInfoFromRefPos {m : TypeType} [inst : Monad m] [inst : Lean.MonadRef m] :
          Equations
          instance Lean.instMonadQuotation {m : TypeType} {n : TypeType} [inst : MonadFunctor m n] [inst : MonadLift m n] [inst : Lean.MonadQuotation m] :
          Equations
          @[export lean_simp_macro_scopes]
          Equations
          structure Lean.MacroScopesView:
          Type
          Equations
          Equations
          • One or more equations did not get rendered due to their size.

          Revert all addMacroScope calls. v = extractMacroScopes n → n = v.review. This operation is useful for analyzing/transforming the original identifiers, then adding back the scopes (via MacroScopesView.review).

          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.
          @[inline]
          def Lean.MonadQuotation.addMacroScope {m : TypeType} [inst : Lean.MonadQuotation m] [inst : Monad m] (n : Lean.Name) :
          Equations
          Equations

          Function used for determining whether a syntax pattern `(id) is matched. There are various conceivable notions of when two syntactic identifiers should be regarded as identical, but semantic definitions like whether they refer to the same global name cannot be implemented without context information (i.e. MonadResolveName). Thus in patterns we default to the structural solution of comparing the identifiers' Name values, though we at least do so modulo macro scopes so that identifiers that "look" the same match. This is particularly useful when dealing with identifiers that do not actually refer to Lean bindings, e.g. in the stx pattern `(many($p)).

          Equations
          Equations
          • One or more equations did not get rendered due to their size.
          structure Lean.Macro.Context:
          Type
          structure Lean.Macro.State:
          Type
          Equations
          @[inline]
          abbrev Lean.Macro:
          Type
          Equations
          Equations
          • One or more equations did not get rendered due to their size.
          Equations
          Equations
          def Lean.Macro.throwError {α : Type} (msg : String) :
          Equations
          @[inline]
          Equations
          • One or more equations did not get rendered due to their size.
          @[inline]
          def Lean.Macro.withIncRecDepth {α : Type} (ref : Lean.Syntax) (x : Lean.MacroM α) :
          Equations
          • One or more equations did not get rendered due to their size.
          Equations
          structure Lean.Macro.Methods:
          Type
          Equations
          • One or more equations did not get rendered due to their size.
          @[implementedBy Lean.Macro.mkMethodsImp]
          @[implementedBy Lean.Macro.getMethodsImp]

          expandMacro? stx return some stxNew if stx is a macro, and stxNew is its expansion.

          Equations

          Return true if the environment contains a declaration with name declName

          Equations
          Equations
          @[inline]

          Function that tries to reverse macro expansions as a post-processing step of delaboration. While less general than an arbitrary delaborator, it can be declared without importing Lean. Used by the [appUnexpander] attribute.

          Equations