Documentation

Init.Prelude

Init.Prelude #

This is the first file in the Lean import hierarchy. It is responsible for setting up basic definitions, most of which Lean already has "built in knowledge" about, so it is important that they be set up in exactly this way. (For example, Lean will use PUnit in the desugaring of do notation, or in the pattern match compiler.)

@[inline]
def id {α : Sort u} (a : α) :
α

The identity function. id takes an implicit argument α : Sort u (a type in any universe), and an argument a : α, and returns a.

Although this may look like a useless function, one application of the identity function is to explicitly put a type on an expression. If e has type T, and T' is definitionally equal to T, then @id T' e typechecks, and Lean knows that this expression has type T' rather than T. This can make a difference for typeclass inference, since T and T' may have different typeclass instances on them. show T' from e is sugar for an @id T' e expression.

Equations
Instances For
    @[inline]
    def Function.comp {α : Sort u} {β : Sort v} {δ : Sort w} (f : βδ) (g : αβ) :
    αδ

    Function composition is the act of pipelining the result of one function, to the input of another, creating an entirely new function. Example:

    #eval Function.comp List.reverse (List.drop 2) [3, 2, 4, 1]
    -- [1, 4]
    

    You can use the notation f ∘ g as shorthand for Function.comp f g.

    #eval (List.reverse ∘ List.drop 2) [3, 2, 4, 1]
    -- [1, 4]
    

    A simpler way of thinking about it, is that List.reverseList.drop 2 is equivalent to fun xs => List.reverse (List.drop 2 xs), the benefit is that the meaning of composition is obvious, and the representation is compact.

    Equations
    • (f g) x = f (g x)
    Instances For
      @[inline]
      def Function.const {α : Sort u} (β : Sort v) (a : α) :
      βα

      The constant function. If a : α, then Function.const β a : β → α is the "constant function with value a", that is, Function.const β a b = a.

      example (b : Bool) : Function.const Bool 10 b = 10 :=
        rfl
      
      #check Function.const Bool 10
      -- BoolNat
      
      Equations
      Instances For
        @[irreducible]
        def letFun {α : Sort u} {β : αSort v} (v : α) (f : (x : α) → β x) :
        β v

        The encoding of let_fun x := v; b is letFun v (fun x => b). This is equal to (fun x => b) v, so the value of x is not accessible to b. This is in contrast to let x := v; b, where the value of x is accessible to b.

        There is special support for letFun. Both WHNF and simp are aware of letFun and can reduce it when zeta reduction is enabled, despite the fact it is marked irreducible. For metaprogramming, the function Lean.Expr.letFun? can be used to recognize a let_fun expression to extract its parts as if it were a let expression.

        Equations
        Instances For
          @[inline, reducible]
          abbrev inferInstance {α : Sort u} [i : α] :
          α

          inferInstance synthesizes a value of any target type by typeclass inference. This function has the same type signature as the identity function, but the square brackets on the [i : α] argument means that it will attempt to construct this argument by typeclass inference. (This will fail if α is not a class.) Example:

          #check (inferInstance : Inhabited Nat) -- Inhabited Nat
          
          def foo : Inhabited (Nat × Nat) :=
            inferInstance
          
          example : foo.default = (default, default) :=
            rfl
          
          Equations
          • inferInstance = i
          Instances For
            @[inline, reducible]
            abbrev inferInstanceAs (α : Sort u) [i : α] :
            α

            inferInstanceAs α synthesizes a value of any target type by typeclass inference. This is just like inferInstance except that α is given explicitly instead of being inferred from the target type. It is especially useful when the target type is some α' which is definitionally equal to α, but the instance we are looking for is only registered for α (because typeclass search does not unfold most definitions, but definitional equality does.) Example:

            #check inferInstanceAs (Inhabited Nat) -- Inhabited Nat
            
            Equations
            Instances For
              inductive PUnit :

              The unit type, the canonical type with one element, named unit or (). This is the universe-polymorphic version of Unit; it is preferred to use Unit instead where applicable. For more information about universe levels: Types as objects

              Instances For
                @[inline, reducible]
                abbrev Unit :

                The unit type, the canonical type with one element, named unit or (). In other words, it describes only a single value, which consists of said constructor applied to no arguments whatsoever. The Unit type is similar to void in languages derived from C.

                Unit is actually defined as PUnit.{1} where PUnit is the universe polymorphic version. The Unit should be preferred over PUnit where possible to avoid unnecessary universe parameters.

                In functional programming, Unit is the return type of things that "return nothing", since a type with one element conveys no additional information. When programming with monads, the type m Unit represents an action that has some side effects but does not return a value, while m α would be an action that has side effects and returns a value of type α.

                Equations
                Instances For
                  @[match_pattern, inline, reducible]
                  abbrev Unit.unit :

                  Unit.unit : Unit is the canonical element of the unit type. It can also be written as ().

                  Equations
                  Instances For
                    unsafe axiom lcErased :

                    Marker for information that has been erased by the code generator.

                    unsafe axiom lcProof {α : Prop} :
                    α

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

                    It may look strange to have an axiom that says "every proposition is true", since this is obviously unsound, but the unsafe marker ensures that the kernel will not let this through into regular proofs. The lower levels of the code generator don't need proofs in terms, so this is used to stub the proofs out.

                    unsafe axiom lcCast {α : Sort u} {β : Sort v} (a : α) :
                    β

                    Auxiliary unsafe constant used by the Compiler when erasing casts.

                    unsafe axiom lcUnreachable {α : Sort u} :
                    α

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

                    Like lcProof, this is an unsafe axiom, which means that even though it is not sound, the kernel will not let us use it for regular proofs.

                    Executing this expression to actually synthesize a value of type α causes immediate undefined behavior, and the compiler does take advantage of this to optimize the code assuming that it is not called. If it is not optimized out, it is likely to appear as a print message saying "unreachable code", but this behavior is not guaranteed or stable in any way.

                    inductive True :

                    True is a proposition and has only an introduction rule, True.intro : True. In other words, True is simply true, and has a canonical proof, True.intro For more information: Propositional Logic

                    Instances For
                      inductive False :

                      False is the empty proposition. Thus, it has no introduction rules. It represents a contradiction. False elimination rule, False.rec, expresses the fact that anything follows from a contradiction. This rule is sometimes called ex falso (short for ex falso sequitur quodlibet), or the principle of explosion. For more information: Propositional Logic

                        Instances For
                          inductive Empty :

                          The empty type. It has no constructors. The Empty.rec eliminator expresses the fact that anything follows from the empty type.

                            Instances For
                              inductive PEmpty :

                              The universe-polymorphic empty type. Prefer Empty or False where possible.

                                Instances For
                                  def Not (a : Prop) :

                                  Not p, or ¬p, is the negation of p. It is defined to be p → False, so if your goal is ¬p you can use intro h to turn the goal into h : p ⊢ False, and if you have hn : ¬p and h : p then hn h : False and (hn h).elim will prove anything. For more information: Propositional Logic

                                  Equations
                                  Instances For
                                    @[macro_inline]
                                    def False.elim {C : Sort u} (h : False) :
                                    C

                                    False.elim : False → C says that from False, any desired proposition C holds. Also known as ex falso quodlibet (EFQ) or the principle of explosion.

                                    The target type is actually C : Sort u which means it works for both propositions and types. When executed, this acts like an "unreachable" instruction: it is undefined behavior to run, but it will probably print "unreachable code". (You would need to construct a proof of false to run it anyway, which you can only do using sorry or unsound axioms.)

                                    Equations
                                    Instances For
                                      @[macro_inline]
                                      def absurd {a : Prop} {b : Sort v} (h₁ : a) (h₂ : ¬a) :
                                      b

                                      Anything follows from two contradictory hypotheses. Example:

                                      example (hp : p) (hnp : ¬p) : q := absurd hp hnp
                                      

                                      For more information: Propositional Logic

                                      Equations
                                      Instances For
                                        inductive Eq {α : Sort u_1} :
                                        ααProp

                                        The equality relation. It has one introduction rule, Eq.refl. We use a = b as notation for Eq a b. A fundamental property of equality is that it is an equivalence relation.

                                        variable (α : Type) (a b c d : α)
                                        variable (hab : a = b) (hcb : c = b) (hcd : c = d)
                                        
                                        example : a = d :=
                                          Eq.trans (Eq.trans hab (Eq.symm hcb)) hcd
                                        

                                        Equality is much more than an equivalence relation, however. It has the important property that every assertion respects the equivalence, in the sense that we can substitute equal expressions without changing the truth value. That is, given h1 : a = b and h2 : p a, we can construct a proof for p b using substitution: Eq.subst h1 h2. Example:

                                        example (α : Type) (a b : α) (p : α → Prop)
                                                (h1 : a = b) (h2 : p a) : p b :=
                                          Eq.subst h1 h2
                                        
                                        example (α : Type) (a b : α) (p : α → Prop)
                                            (h1 : a = b) (h2 : p a) : p b :=
                                          h1 ▸ h2
                                        

                                        The triangle in the second presentation is a macro built on top of Eq.subst and Eq.symm, and you can enter it by typing \t. For more information: Equality

                                        • refl: ∀ {α : Sort u_1} (a : α), a = a

                                          Eq.refl a : a = a is reflexivity, the unique constructor of the equality type. See also rfl, which is usually used instead.

                                        Instances For
                                          @[match_pattern]
                                          def rfl {α : Sort u} {a : α} :
                                          a = a

                                          rfl : a = a is the unique constructor of the equality type. This is the same as Eq.refl except that it takes a implicitly instead of explicitly.

                                          This is a more powerful theorem than it may appear at first, because although the statement of the theorem is a = a, Lean will allow anything that is definitionally equal to that type. So, for instance, 2 + 2 = 4 is proven in Lean by rfl, because both sides are the same up to definitional equality.

                                          Equations
                                          • =
                                          Instances For
                                            @[simp]
                                            theorem id_eq {α : Sort u_1} (a : α) :
                                            id a = a

                                            id x = x, as a @[simp] lemma.

                                            theorem Eq.subst {α : Sort u} {motive : αProp} {a : α} {b : α} (h₁ : a = b) (h₂ : motive a) :
                                            motive b

                                            The substitution principle for equality. If a = b and P a holds, then P b also holds. We conventionally use the name motive for P here, so that you can specify it explicitly using e.g. Eq.subst (motive := fun x => x < 5) if it is not otherwise inferred correctly.

                                            This theorem is the underlying mechanism behind the rw tactic, which is essentially a fancy algorithm for finding good motive arguments to usefully apply this theorem to replace occurrences of a with b in the goal or hypotheses.

                                            For more information: Equality

                                            theorem Eq.symm {α : Sort u} {a : α} {b : α} (h : a = b) :
                                            b = a

                                            Equality is symmetric: if a = b then b = a.

                                            Because this is in the Eq namespace, if you have a variable h : a = b, h.symm can be used as shorthand for Eq.symm h as a proof of b = a.

                                            For more information: Equality

                                            theorem Eq.trans {α : Sort u} {a : α} {b : α} {c : α} (h₁ : a = b) (h₂ : b = c) :
                                            a = c

                                            Equality is transitive: if a = b and b = c then a = c.

                                            Because this is in the Eq namespace, if you have variables or expressions h₁ : a = b and h₂ : b = c, you can use h₁.trans h₂ : a = c as shorthand for Eq.trans h₁ h₂.

                                            For more information: Equality

                                            @[macro_inline]
                                            def cast {α : Sort u} {β : Sort u} (h : α = β) (a : α) :
                                            β

                                            Cast across a type equality. If h : α = β is an equality of types, and a : α, then a : β will usually not typecheck directly, but this function will allow you to work around this and embed a in type β as cast h a : β.

                                            It is best to avoid this function if you can, because it is more complicated to reason about terms containing casts, but if the types don't match up definitionally sometimes there isn't anything better you can do.

                                            For more information: Equality

                                            Equations
                                            Instances For
                                              theorem congrArg {α : Sort u} {β : Sort v} {a₁ : α} {a₂ : α} (f : αβ) (h : a₁ = a₂) :
                                              f a₁ = f a₂

                                              Congruence in the function argument: if a₁ = a₂ then f a₁ = f a₂ for any (nondependent) function f. This is more powerful than it might look at first, because you can also use a lambda expression for f to prove that <something containing a₁> = <something containing a₂>. This function is used internally by tactics like congr and simp to apply equalities inside subterms.

                                              For more information: Equality

                                              theorem congr {α : Sort u} {β : Sort v} {f₁ : αβ} {f₂ : αβ} {a₁ : α} {a₂ : α} (h₁ : f₁ = f₂) (h₂ : a₁ = a₂) :
                                              f₁ a₁ = f₂ a₂

                                              Congruence in both function and argument. If f₁ = f₂ and a₁ = a₂ then f₁ a₁ = f₂ a₂. This only works for nondependent functions; the theorem statement is more complex in the dependent case.

                                              For more information: Equality

                                              theorem congrFun {α : Sort u} {β : αSort v} {f : (x : α) → β x} {g : (x : α) → β x} (h : f = g) (a : α) :
                                              f a = g a

                                              Congruence in the function part of an application: If f = g then f a = g a.

                                              Initialize the Quotient Module, which effectively adds the following definitions:

                                              opaque Quot {α : Sort u} (r : α → α → Prop) : Sort u
                                              
                                              opaque Quot.mk {α : Sort u} (r : α → α → Prop) (a : α) : Quot r
                                              
                                              opaque Quot.lift {α : Sort u} {r : α → α → Prop} {β : Sort v} (f : α → β) :
                                                (∀ a b : α, r a b → Eq (f a) (f b)) → Quot r → β
                                              
                                              opaque Quot.ind {α : Sort u} {r : α → α → Prop} {β : Quot r → Prop} :
                                                (∀ a : α, β (Quot.mk r a)) → ∀ q : Quot r, β q
                                              
                                              unsafe axiom Quot.lcInv {α : Sort u} {r : ααProp} (q : Quot r) :
                                              α

                                              Unsafe auxiliary constant used by the compiler to erase Quot.lift.

                                              inductive HEq {α : Sort u} :
                                              α{β : Sort u} → βProp

                                              Heterogeneous equality. HEq a b asserts that a and b have the same type, and casting a across the equality yields b, and vice versa.

                                              You should avoid using this type if you can. Heterogeneous equality does not have all the same properties as Eq, because the assumption that the types of a and b are equal is often too weak to prove theorems of interest. One important non-theorem is the analogue of congr: If HEq f g and HEq x y and f x and g y are well typed it does not follow that HEq (f x) (g y). (This does follow if you have f = g instead.) However if a and b have the same type then a = b and HEq a b are equivalent.

                                              • refl: ∀ {α : Sort u} (a : α), HEq a a

                                                Reflexivity of heterogeneous equality.

                                              Instances For
                                                @[match_pattern]
                                                def HEq.rfl {α : Sort u} {a : α} :
                                                HEq a a

                                                A version of HEq.refl with an implicit argument.

                                                Equations
                                                • =
                                                Instances For
                                                  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)

                                                  Product type (aka pair). You can use α × β as notation for Prod α β. Given a : α and b : β, Prod.mk a b : Prod α β. You can use (a, b) as notation for Prod.mk a b. Moreover, (a, b, c) is notation for Prod.mk a (Prod.mk b c). Given p : Prod α β, p.1 : α and p.2 : β. They are short for Prod.fst p and Prod.snd p respectively. You can also write p.fst and p.snd. For more information: Constructors with Arguments

                                                  • fst : α

                                                    The first projection out of a pair. if p : α × β then p.1 : α.

                                                  • snd : β

                                                    The second projection out of a pair. if p : α × β then p.2 : β.

                                                  Instances For
                                                    structure PProd (α : Sort u) (β : Sort v) :
                                                    Sort (max (max 1 u) v)

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

                                                    • fst : α

                                                      The first projection out of a pair. if p : PProd α β then p.1 : α.

                                                    • snd : β

                                                      The second projection out of a pair. if p : PProd α β then p.2 : β.

                                                    Instances For
                                                      structure MProd (α : Type u) (β : Type u) :

                                                      Similar to Prod, but α and β are in the same universe. We say MProd is the universe monomorphic product type.

                                                      • fst : α

                                                        The first projection out of a pair. if p : MProd α β then p.1 : α.

                                                      • snd : β

                                                        The second projection out of a pair. if p : MProd α β then p.2 : β.

                                                      Instances For
                                                        structure And (a : Prop) (b : Prop) :

                                                        And a b, or a ∧ b, is the conjunction of propositions. It can be constructed and destructed like a pair: if ha : a and hb : b then ⟨ha, hb⟩ : a ∧ b, and if h : a ∧ b then h.left : a and h.right : b.

                                                        • intro :: (
                                                          • left : a

                                                            Extract the left conjunct from a conjunction. h : a ∧ b then h.left, also notated as h.1, is a proof of a.

                                                          • right : b

                                                            Extract the right conjunct from a conjunction. h : a ∧ b then h.right, also notated as h.2, is a proof of b.

                                                        • )
                                                        Instances For
                                                          inductive Or (a : Prop) (b : Prop) :

                                                          Or a b, or a ∨ b, is the disjunction of propositions. There are two constructors for Or, called Or.inl : a → a ∨ b and Or.inr : b → a ∨ b, and you can use match or cases to destruct an Or assumption into the two cases.

                                                          • inl: ∀ {a b : Prop}, aa b

                                                            Or.inl is "left injection" into an Or. If h : a then Or.inl h : a ∨ b.

                                                          • inr: ∀ {a b : Prop}, ba b

                                                            Or.inr is "right injection" into an Or. If h : b then Or.inr h : a ∨ b.

                                                          Instances For
                                                            theorem Or.intro_left {a : Prop} (b : Prop) (h : a) :
                                                            a b

                                                            Alias for Or.inl.

                                                            theorem Or.intro_right {b : Prop} (a : Prop) (h : b) :
                                                            a b

                                                            Alias for Or.inr.

                                                            theorem Or.elim {a : Prop} {b : Prop} {c : Prop} (h : a b) (left : ac) (right : bc) :
                                                            c

                                                            Proof by cases on an Or. If a ∨ b, and both a and b imply proposition c, then c is true.

                                                            theorem Or.resolve_left {a : Prop} {b : Prop} (h : a b) (na : ¬a) :
                                                            b
                                                            theorem Or.resolve_right {a : Prop} {b : Prop} (h : a b) (nb : ¬b) :
                                                            a
                                                            theorem Or.neg_resolve_left {a : Prop} {b : Prop} (h : ¬a b) (ha : a) :
                                                            b
                                                            theorem Or.neg_resolve_right {a : Prop} {b : Prop} (h : a ¬b) (nb : b) :
                                                            a
                                                            inductive Bool :

                                                            Bool is the type of boolean values, true and false. Classically, this is equivalent to Prop (the type of propositions), but the distinction is important for programming, because values of type Prop are erased in the code generator, while Bool corresponds to the type called bool or boolean in most programming languages.

                                                            • false: Bool

                                                              The boolean value false, not to be confused with the proposition False.

                                                            • true: Bool

                                                              The boolean value true, not to be confused with the proposition True.

                                                            Instances For
                                                              structure Subtype {α : Sort u} (p : αProp) :
                                                              Sort (max 1 u)

                                                              Subtype p, usually written as {x : α // p x}, is a type which represents all the elements x : α for which p x is true. It is structurally a pair-like type, so if you have x : α and h : p x then ⟨x, h⟩ : {x // p x}. An element s : {x // p x} will coerce to α but you can also make it explicit using s.1 or s.val.

                                                              • val : α

                                                                If s : {x // p x} then s.val : α is the underlying element in the base type. You can also write this as s.1, or simply as s when the type is known from context.

                                                              • property : p self.val

                                                                If s : {x // p x} then s.2 or s.property is the assertion that p s.1, that is, that s is in fact an element for which p holds.

                                                              Instances For
                                                                @[reducible]
                                                                def optParam (α : Sort u) (default : α) :

                                                                Gadget for optional parameter support.

                                                                A binder like (x : α := default) in a declaration is syntax sugar for x : optParam α default, and triggers the elaborator to attempt to use default to supply the argument if it is not supplied.

                                                                Equations
                                                                Instances For
                                                                  @[reducible]
                                                                  def outParam (α : Sort u) :

                                                                  Gadget for marking output parameters in type classes.

                                                                  For example, the Membership class is defined as:

                                                                  class Membership (α : outParam (Type u)) (γ : Type v)
                                                                  

                                                                  This means that whenever a typeclass goal of the form Membership ?α ?γ comes up, Lean will wait to solve it until is known, but then it will run typeclass inference, and take the first solution it finds, for any value of , which thereby determines what should be.

                                                                  This expresses that in a term like a ∈ s, s might be a Set α or List α or some other type with a membership operation, and in each case the "member" type α is determined by looking at the container type.

                                                                  Equations
                                                                  Instances For
                                                                    @[reducible]
                                                                    def semiOutParam (α : Sort u) :

                                                                    Gadget for marking semi output parameters in type classes.

                                                                    Semi-output parameters influence the order in which arguments to type class instances are processed. Lean determines an order where all non-(semi-)output parameters to the instance argument have to be figured out before attempting to synthesize an argument (that is, they do not contain assignable metavariables created during TC synthesis). This rules out instances such as [Mul β] : Add α (because β could be anything). Marking a parameter as semi-output is a promise that instances of the type class will always fill in a value for that parameter.

                                                                    For example, the Coe class is defined as:

                                                                    class Coe (α : semiOutParam (Sort u)) (β : Sort v)
                                                                    

                                                                    This means that all Coe instances should provide a concrete value for α (i.e., not an assignable metavariable). An instance like Coe Nat Int or Coe α (Option α) is fine, but Coe α Nat is not since it does not provide a value for α.

                                                                    Equations
                                                                    Instances For
                                                                      @[reducible]
                                                                      def namedPattern {α : Sort u} (x : α) (a : α) (h : x = a) :
                                                                      α

                                                                      Auxiliary declaration used to implement named patterns like x@h:p.

                                                                      Equations
                                                                      Instances For
                                                                        @[never_extract, extern lean_sorry]
                                                                        axiom sorryAx (α : Sort u) (synthetic : optParam Bool false) :
                                                                        α

                                                                        Auxiliary axiom used to implement sorry.

                                                                        The sorry term/tactic expands to sorryAx _ (synthetic := false). This is a proof of anything, which is intended for stubbing out incomplete parts of a proof while still having a syntactically correct proof skeleton. Lean will give a warning whenever a proof uses sorry, so you aren't likely to miss it, but you can double check if a theorem depends on sorry by using #print axioms my_thm and looking for sorryAx in the axiom list.

                                                                        The synthetic flag is false when written explicitly by the user, but it is set to true when a tactic fails to prove a goal, or if there is a type error in the expression. A synthetic sorry acts like a regular one, except that it suppresses follow-up errors in order to prevent one error from causing a cascade of other errors because the desired term was not constructed.

                                                                        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
                                                                        class Inhabited (α : Sort u) :
                                                                        Sort (max 1 u)

                                                                        Inhabited α is a typeclass that says that α has a designated element, called (default : α). This is sometimes referred to as a "pointed type".

                                                                        This class is used by functions that need to return a value of the type when called "out of domain". For example, Array.get! arr i : α returns a value of type α when arr : Array α, but if i is not in range of the array, it reports a panic message, but this does not halt the program, so it must still return a value of type α (and in fact this is required for logical consistency), so in this case it returns default.

                                                                        • default : α

                                                                          default is a function that produces a "default" element of any Inhabited type. This element does not have any particular specified properties, but it is often an all-zeroes value.

                                                                        Instances
                                                                          class inductive Nonempty (α : Sort u) :

                                                                          Nonempty α is a typeclass that says that α is not an empty type, that is, there exists an element in the type. It differs from Inhabited α in that Nonempty α is a Prop, which means that it does not actually carry an element of α, only a proof that there exists such an element. Given Nonempty α, you can construct an element of α nonconstructively using Classical.choice.

                                                                          • intro: ∀ {α : Sort u}, αNonempty α

                                                                            If val : α, then α is nonempty.

                                                                          Instances
                                                                            axiom Classical.choice {α : Sort u} :
                                                                            Nonempty αα

                                                                            The axiom of choice. Nonempty α is a proof that α has an element, but the element itself is erased. The axiom choice supplies a particular element of α given only this proof.

                                                                            The textbook axiom of choice normally makes a family of choices all at once, but that is implied from this formulation, because if α : ι → Type is a family of types and h : ∀ i, Nonempty (α i) is a proof that they are all nonempty, then fun i => Classical.choice (h i) : ∀ i, α i is a family of chosen elements. This is actually a bit stronger than the ZFC choice axiom; this is sometimes called "global choice".

                                                                            In Lean, we use the axiom of choice to derive the law of excluded middle (see Classical.em), so it will often show up in axiom listings where you may not expect. You can use #print axioms my_thm to find out if a given theorem depends on this or other axioms.

                                                                            This axiom can be used to construct "data", but obviously there is no algorithm to compute it, so Lean will require you to mark any definition that would involve executing Classical.choice or other axioms as noncomputable, and will not produce any executable code for such definitions.

                                                                            def Nonempty.elim {α : Sort u} {p : Prop} (h₁ : Nonempty α) (h₂ : αp) :
                                                                            p

                                                                            The elimination principle for Nonempty α. If Nonempty α, and we can prove p given any element x : α, then p holds. Note that it is essential that p is a Prop here; the version with p being a Sort u is equivalent to Classical.choice.

                                                                            Equations
                                                                            • =
                                                                            Instances For
                                                                              instance instNonempty {α : Sort u} [Inhabited α] :
                                                                              Equations
                                                                              • =
                                                                              noncomputable def Classical.ofNonempty {α : Sort u} [Nonempty α] :
                                                                              α

                                                                              A variation on Classical.choice that uses typeclass inference to infer the proof of Nonempty α.

                                                                              Equations
                                                                              Instances For
                                                                                instance instNonemptyForAll (α : Sort u) {β : Sort v} [Nonempty β] :
                                                                                Nonempty (αβ)
                                                                                Equations
                                                                                • =
                                                                                instance instNonemptyForAll_1 (α : Sort u) {β : αSort v} [∀ (a : α), Nonempty (β a)] :
                                                                                Nonempty ((a : α) → β a)
                                                                                Equations
                                                                                • =
                                                                                Equations
                                                                                instance instInhabitedForAll (α : Sort u) {β : Sort v} [Inhabited β] :
                                                                                Inhabited (αβ)
                                                                                Equations
                                                                                instance instInhabitedForAll_1 (α : Sort u) {β : αSort v} [(a : α) → Inhabited (β a)] :
                                                                                Inhabited ((a : α) → β a)
                                                                                Equations
                                                                                structure PLift (α : Sort u) :

                                                                                Universe lifting operation from Sort u to Type u.

                                                                                • up :: (
                                                                                  • down : α

                                                                                    Extract a value from PLift α

                                                                                • )
                                                                                Instances For
                                                                                  theorem PLift.up_down {α : Sort u} (b : PLift α) :
                                                                                  { down := b.down } = b

                                                                                  Bijection between α and PLift α

                                                                                  theorem PLift.down_up {α : Sort u} (a : α) :
                                                                                  { down := a }.down = a

                                                                                  Bijection between α and PLift α

                                                                                  def NonemptyType :
                                                                                  Type (u + 1)

                                                                                  NonemptyType.{u} is the type of nonempty types in universe u. It is mainly used in constant declarations where we wish to introduce a type and simultaneously assert that it is nonempty, but otherwise make the type opaque.

                                                                                  Equations
                                                                                  Instances For
                                                                                    @[inline, reducible]

                                                                                    The underlying type of a NonemptyType.

                                                                                    Equations
                                                                                    Instances For

                                                                                      NonemptyType is inhabited, because PUnit is a nonempty type.

                                                                                      Equations
                                                                                      structure ULift (α : Type s) :
                                                                                      Type (max s r)

                                                                                      Universe lifting operation from a lower Type universe to a higher one. To express this using level variables, the input is Type s and the output is Type (max s r), so if s ≤ r then the latter is (definitionally) Type r.

                                                                                      The universe variable r is written first so that ULift.{r} α can be used when s can be inferred from the type of α.

                                                                                      • up :: (
                                                                                        • down : α

                                                                                          Extract a value from ULift α

                                                                                      • )
                                                                                      Instances For
                                                                                        theorem ULift.up_down {α : Type u} (b : ULift α) :
                                                                                        { down := b.down } = b

                                                                                        Bijection between α and ULift.{v} α

                                                                                        theorem ULift.down_up {α : Type u} (a : α) :
                                                                                        { down := a }.down = a

                                                                                        Bijection between α and ULift.{v} α

                                                                                        class inductive Decidable (p : Prop) :

                                                                                        Decidable p is a data-carrying class that supplies a proof that p is either true or false. It is equivalent to Bool (and in fact it has the same code generation as Bool) together with a proof that the Bool is true iff p is.

                                                                                        Decidable instances are used to infer "computation strategies" for propositions, so that you can have the convenience of writing propositions inside if statements and executing them (which actually executes the inferred decidability instance instead of the proposition, which has no code).

                                                                                        If a proposition p is Decidable, then (by decide : p) will prove it by evaluating the decidability instance to isTrue h and returning h.

                                                                                        Because Decidable carries data, when writing @[simp] lemmas which include a Decidable instance on the LHS, it is best to use {_ : Decidable p} rather than [Decidable p] so that non-canonical instances can be found via unification rather than typeclass search.

                                                                                        • isFalse: {p : Prop} → ¬pDecidable p

                                                                                          Prove that p is decidable by supplying a proof of ¬p

                                                                                        • isTrue: {p : Prop} → pDecidable p

                                                                                          Prove that p is decidable by supplying a proof of p

                                                                                        Instances
                                                                                          @[inline_if_reduce]
                                                                                          def Decidable.decide (p : Prop) [h : Decidable p] :

                                                                                          Convert a decidable proposition into a boolean value.

                                                                                          If p : Prop is decidable, then decide p : Bool is the boolean value which is true if p is true and false if p is false.

                                                                                          Equations
                                                                                          Instances For
                                                                                            @[inline, reducible]
                                                                                            abbrev DecidablePred {α : Sort u} (r : αProp) :
                                                                                            Sort (max 1 u)

                                                                                            A decidable predicate. See Decidable.

                                                                                            Equations
                                                                                            Instances For
                                                                                              @[inline, reducible]
                                                                                              abbrev DecidableRel {α : Sort u} (r : ααProp) :
                                                                                              Sort (max 1 u)

                                                                                              A decidable relation. See Decidable.

                                                                                              Equations
                                                                                              Instances For
                                                                                                @[inline, reducible]
                                                                                                abbrev DecidableEq (α : Sort u) :
                                                                                                Sort (max 1 u)

                                                                                                Asserts that α has decidable equality, that is, a = b is decidable for all a b : α. See Decidable.

                                                                                                Equations
                                                                                                Instances For
                                                                                                  def decEq {α : Sort u} [inst : DecidableEq α] (a : α) (b : α) :
                                                                                                  Decidable (a = b)

                                                                                                  Proves that a = b is decidable given DecidableEq α.

                                                                                                  Equations
                                                                                                  Instances For
                                                                                                    theorem decide_eq_true {p : Prop} [inst : Decidable p] :
                                                                                                    pdecide p = true
                                                                                                    theorem decide_eq_false {p : Prop} [Decidable p] :
                                                                                                    ¬pdecide p = false
                                                                                                    theorem of_decide_eq_true {p : Prop} [inst : Decidable p] :
                                                                                                    decide p = truep
                                                                                                    theorem of_decide_eq_false {p : Prop} [inst : Decidable p] :
                                                                                                    decide p = false¬p
                                                                                                    theorem of_decide_eq_self_eq_true {α : Sort u_1} [inst : DecidableEq α] (a : α) :
                                                                                                    decide (a = a) = true
                                                                                                    @[inline]
                                                                                                    def Bool.decEq (a : Bool) (b : Bool) :
                                                                                                    Decidable (a = b)

                                                                                                    Decidable equality for Bool

                                                                                                    Equations
                                                                                                    Instances For
                                                                                                      class BEq (α : Type u) :

                                                                                                      BEq α is a typeclass for supplying a boolean-valued equality relation on α, notated as a == b. Unlike DecidableEq α (which uses a = b), this is Bool valued instead of Prop valued, and it also does not have any axioms like being reflexive or agreeing with =. It is mainly intended for programming applications. See LawfulBEq for a version that requires that == and = coincide.

                                                                                                      • beq : ααBool

                                                                                                        Boolean equality, notated as a == b.

                                                                                                      Instances
                                                                                                        instance instBEq {α : Type u_1} [DecidableEq α] :
                                                                                                        BEq α
                                                                                                        Equations
                                                                                                        • instBEq = { beq := fun (a b : α) => decide (a = b) }
                                                                                                        @[macro_inline]
                                                                                                        def dite {α : Sort u} (c : Prop) [h : Decidable c] (t : cα) (e : ¬cα) :
                                                                                                        α

                                                                                                        "Dependent" if-then-else, normally written via the notation if h : c then t(h) else e(h), is sugar for dite c (fun h => t(h)) (fun h => e(h)), and it is the same as if c then t else e except that t is allowed to depend on a proof h : c, and e can depend on h : ¬c. (Both branches use the same name for the hypothesis, even though it has different types in the two cases.)

                                                                                                        We use this to be able to communicate the if-then-else condition to the branches. For example, Array.get arr ⟨i, h⟩ expects a proof h : i < arr.size in order to avoid a bounds check, so you can write if h : i < arr.size then arr.get ⟨i, h⟩ else ... to avoid the bounds check inside the if branch. (Of course in this case we have only lifted the check into an explicit if, but we could also use this proof multiple times or derive i < arr.size from some other proposition that we are checking in the if.)

                                                                                                        Equations
                                                                                                        Instances For

                                                                                                          if-then-else #

                                                                                                          @[macro_inline]
                                                                                                          def ite {α : Sort u} (c : Prop) [h : Decidable c] (t : α) (e : α) :
                                                                                                          α

                                                                                                          if c then t else e is notation for ite c t e, "if-then-else", which decides to return t or e depending on whether c is true or false. The explicit argument c : Prop does not have any actual computational content, but there is an additional [Decidable c] argument synthesized by typeclass inference which actually determines how to evaluate c to true or false. Write if h : c then t else e instead for a "dependent if-then-else" dite, which allows t/e to use the fact that c is true/false.

                                                                                                          Equations
                                                                                                          Instances For
                                                                                                            @[macro_inline]
                                                                                                            instance instDecidableAnd {p : Prop} {q : Prop} [dp : Decidable p] [dq : Decidable q] :
                                                                                                            Equations
                                                                                                            @[macro_inline]
                                                                                                            instance instDecidableOr {p : Prop} {q : Prop} [dp : Decidable p] [dq : Decidable q] :
                                                                                                            Equations
                                                                                                            instance instDecidableNot {p : Prop} [dp : Decidable p] :
                                                                                                            Equations

                                                                                                            Boolean operators #

                                                                                                            @[macro_inline]
                                                                                                            def cond {α : Type u} (c : Bool) (x : α) (y : α) :
                                                                                                            α

                                                                                                            cond b x y is the same as if b then x else y, but optimized for a boolean condition. It can also be written as bif b then x else y. This is @[macro_inline] because x and y should not be eagerly evaluated (see ite).

                                                                                                            Equations
                                                                                                            • (bif c then x else y) = match c with | true => x | false => y
                                                                                                            Instances For
                                                                                                              @[macro_inline]
                                                                                                              def or (x : Bool) (y : Bool) :

                                                                                                              or x y, or x || y, is the boolean "or" operation (not to be confused with Or : Prop → Prop → Prop, which is the propositional connective). It is @[macro_inline] because it has C-like short-circuiting behavior: if x is true then y is not evaluated.

                                                                                                              Equations
                                                                                                              Instances For
                                                                                                                @[macro_inline]
                                                                                                                def and (x : Bool) (y : Bool) :

                                                                                                                and x y, or x && y, is the boolean "and" operation (not to be confused with And : Prop → Prop → Prop, which is the propositional connective). It is @[macro_inline] because it has C-like short-circuiting behavior: if x is false then y is not evaluated.

                                                                                                                Equations
                                                                                                                Instances For
                                                                                                                  @[inline]
                                                                                                                  def not :

                                                                                                                  not x, or !x, is the boolean "not" operation (not to be confused with Not : Prop → Prop, which is the propositional connective).

                                                                                                                  Equations
                                                                                                                  Instances For
                                                                                                                    inductive Nat :

                                                                                                                    The type of natural numbers, starting at zero. It is defined as an inductive type freely generated by "zero is a natural number" and "the successor of a natural number is a natural number".

                                                                                                                    You can prove a theorem P n about n : Nat by induction n, which will expect a proof of the theorem for P 0, and a proof of P (succ i) assuming a proof of P i. The same method also works to define functions by recursion on natural numbers: induction and recursion are two expressions of the same operation from Lean's point of view.

                                                                                                                    open Nat
                                                                                                                    example (n : Nat) : n < succ n := by
                                                                                                                      induction n with
                                                                                                                      | zero =>
                                                                                                                        show 0 < 1
                                                                                                                        decide
                                                                                                                      | succ i ih => -- ih : i < succ i
                                                                                                                        show succ i < succ (succ i)
                                                                                                                        exact Nat.succ_lt_succ ih
                                                                                                                    

                                                                                                                    This type is special-cased by both the kernel and the compiler:

                                                                                                                    • The type of expressions contains "Nat literals" as a primitive constructor, and the kernel knows how to reduce zero/succ expressions to nat literals.
                                                                                                                    • If implemented naively, this type would represent a numeral n in unary as a linked list with n links, which is horribly inefficient. Instead, the runtime itself has a special representation for Nat which stores numbers up to 2^63 directly and larger numbers use an arbitrary precision "bignum" library (usually GMP).
                                                                                                                    • zero: Nat

                                                                                                                      Nat.zero, normally written 0 : Nat, is the smallest natural number. This is one of the two constructors of Nat.

                                                                                                                    • succ: NatNat

                                                                                                                      The successor function on natural numbers, succ n = n + 1. This is one of the two constructors of Nat.

                                                                                                                    Instances For
                                                                                                                      class OfNat (α : Type u) :
                                                                                                                      NatType u

                                                                                                                      The class OfNat α n powers the numeric literal parser. If you write 37 : α, Lean will attempt to synthesize OfNat α 37, and will generate the term (OfNat.ofNat 37 : α).

                                                                                                                      There is a bit of infinite regress here since the desugaring apparently still contains a literal 37 in it. The type of expressions contains a primitive constructor for "raw natural number literals", which you can directly access using the macro nat_lit 37. Raw number literals are always of type Nat. So it would be more correct to say that Lean looks for an instance of OfNat α (nat_lit 37), and it generates the term (OfNat.ofNat (nat_lit 37) : α).

                                                                                                                      • ofNat : α

                                                                                                                        The OfNat.ofNat function is automatically inserted by the parser when the user writes a numeric literal like 1 : α. Implementations of this typeclass can therefore customize the behavior of n : α based on n and α.

                                                                                                                      Instances
                                                                                                                        @[defaultInstance 100]
                                                                                                                        instance instOfNatNat (n : Nat) :
                                                                                                                        Equations
                                                                                                                        class LE (α : Type u) :

                                                                                                                        LE α is the typeclass which supports the notation x ≤ y where x y : α.

                                                                                                                        • le : ααProp

                                                                                                                          The less-equal relation: x ≤ y

                                                                                                                        Instances
                                                                                                                          class LT (α : Type u) :

                                                                                                                          LT α is the typeclass which supports the notation x < y where x y : α.

                                                                                                                          • lt : ααProp

                                                                                                                            The less-than relation: x < y

                                                                                                                          Instances
                                                                                                                            @[reducible]
                                                                                                                            def GE.ge {α : Type u} [LE α] (a : α) (b : α) :

                                                                                                                            a ≥ b is an abbreviation for b ≤ a.

                                                                                                                            Equations
                                                                                                                            Instances For
                                                                                                                              @[reducible]
                                                                                                                              def GT.gt {α : Type u} [LT α] (a : α) (b : α) :

                                                                                                                              a > b is an abbreviation for b < a.

                                                                                                                              Equations
                                                                                                                              • (a > b) = (b < a)
                                                                                                                              Instances For
                                                                                                                                class Max (α : Type u) :

                                                                                                                                Max α is the typeclass which supports the operation max x y where x y : α.

                                                                                                                                • max : ααα

                                                                                                                                  The maximum operation: max x y.

                                                                                                                                Instances
                                                                                                                                  @[inline]
                                                                                                                                  def maxOfLe {α : Type u_1} [LE α] [DecidableRel LE.le] :
                                                                                                                                  Max α

                                                                                                                                  Implementation of the max operation using .

                                                                                                                                  Equations
                                                                                                                                  • maxOfLe = { max := fun (x y : α) => if x y then y else x }
                                                                                                                                  Instances For
                                                                                                                                    class Min (α : Type u) :

                                                                                                                                    Min α is the typeclass which supports the operation min x y where x y : α.

                                                                                                                                    • min : ααα

                                                                                                                                      The minimum operation: min x y.

                                                                                                                                    Instances
                                                                                                                                      @[inline]
                                                                                                                                      def minOfLe {α : Type u_1} [LE α] [DecidableRel LE.le] :
                                                                                                                                      Min α

                                                                                                                                      Implementation of the min operation using .

                                                                                                                                      Equations
                                                                                                                                      • minOfLe = { min := fun (x y : α) => if x y then x else y }
                                                                                                                                      Instances For
                                                                                                                                        class Trans {α : Sort u_1} {β : Sort u_2} {γ : Sort u_3} (r : αβSort u) (s : βγSort v) (t : outParam (αγSort w)) :
                                                                                                                                        Sort (max (max (max (max (max (max 1 u) u_1) u_2) u_3) v) w)

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

                                                                                                                                        It takes two relations r and s as "input", and produces an "output" relation t, with the property that r a b and s b c implies t a c. The calc tactic uses this so that when it sees a chain with a ≤ b and b < c it knows that this should be a proof of a < c because there is an instance Trans (·≤·) (·<·) (·<·).

                                                                                                                                        • trans : {a : α} → {b : β} → {c : γ} → r a bs b ct a c

                                                                                                                                          Compose two proofs by transitivity, generalized over the relations involved.

                                                                                                                                        Instances
                                                                                                                                          instance instTransEq {α : Sort u_1} {γ : Sort u_2} (r : αγSort u) :
                                                                                                                                          Trans Eq r r
                                                                                                                                          Equations
                                                                                                                                          • instTransEq r = { trans := fun {a b : α} {c : γ} (heq : a = b) (h' : r b c) => h' }
                                                                                                                                          instance instTransEq_1 {α : Sort u_1} {β : Sort u_2} (r : αβSort u) :
                                                                                                                                          Trans r Eq r
                                                                                                                                          Equations
                                                                                                                                          • instTransEq_1 r = { trans := fun {a : α} {b c : β} (h' : r a b) (heq : b = c) => heqh' }
                                                                                                                                          class HAdd (α : Type u) (β : Type v) (γ : outParam (Type w)) :
                                                                                                                                          Type (max (max u v) w)

                                                                                                                                          The notation typeclass for heterogeneous addition. This enables the notation a + b : γ where a : α, b : β.

                                                                                                                                          • hAdd : αβγ

                                                                                                                                            a + b computes the sum of a and b. The meaning of this notation is type-dependent.

                                                                                                                                          Instances
                                                                                                                                            class HSub (α : Type u) (β : Type v) (γ : outParam (Type w)) :
                                                                                                                                            Type (max (max u v) w)

                                                                                                                                            The notation typeclass for heterogeneous subtraction. This enables the notation a - b : γ where a : α, b : β.

                                                                                                                                            • hSub : αβγ

                                                                                                                                              a - b computes the difference of a and b. The meaning of this notation is type-dependent.

                                                                                                                                              • For natural numbers, this operator saturates at 0: a - b = 0 when a ≤ b.
                                                                                                                                            Instances
                                                                                                                                              class HMul (α : Type u) (β : Type v) (γ : outParam (Type w)) :
                                                                                                                                              Type (max (max u v) w)

                                                                                                                                              The notation typeclass for heterogeneous multiplication. This enables the notation a * b : γ where a : α, b : β.

                                                                                                                                              • hMul : αβγ

                                                                                                                                                a * b computes the product of a and b. The meaning of this notation is type-dependent.

                                                                                                                                              Instances
                                                                                                                                                class HDiv (α : Type u) (β : Type v) (γ : outParam (Type w)) :
                                                                                                                                                Type (max (max u v) w)

                                                                                                                                                The notation typeclass for heterogeneous division. This enables the notation a / b : γ where a : α, b : β.

                                                                                                                                                • hDiv : αβγ

                                                                                                                                                  a / b computes the result of dividing a by b. The meaning of this notation is type-dependent.

                                                                                                                                                  • For most types like Nat, Int, Rat, Real, a / 0 is defined to be 0.
                                                                                                                                                  • For Nat and Int, a / b rounds toward 0.
                                                                                                                                                  • For Float, a / 0 follows the IEEE 754 semantics for division, usually resulting in inf or nan.
                                                                                                                                                Instances
                                                                                                                                                  class HMod (α : Type u) (β : Type v) (γ : outParam (Type w)) :
                                                                                                                                                  Type (max (max u v) w)

                                                                                                                                                  The notation typeclass for heterogeneous modulo / remainder. This enables the notation a % b : γ where a : α, b : β.

                                                                                                                                                  • hMod : αβγ

                                                                                                                                                    a % b computes the remainder upon dividing a by b. The meaning of this notation is type-dependent.

                                                                                                                                                    • For Nat and Int, a % 0 is defined to be a.
                                                                                                                                                  Instances
                                                                                                                                                    class HPow (α : Type u) (β : Type v) (γ : outParam (Type w)) :
                                                                                                                                                    Type (max (max u v) w)

                                                                                                                                                    The notation typeclass for heterogeneous exponentiation. This enables the notation a ^ b : γ where a : α, b : β.

                                                                                                                                                    • hPow : αβγ

                                                                                                                                                      a ^ b computes a to the power of b. The meaning of this notation is type-dependent.

                                                                                                                                                    Instances
                                                                                                                                                      class HAppend (α : Type u) (β : Type v) (γ : outParam (Type w)) :
                                                                                                                                                      Type (max (max u v) w)

                                                                                                                                                      The notation typeclass for heterogeneous append. This enables the notation a ++ b : γ where a : α, b : β.

                                                                                                                                                      • hAppend : αβγ

                                                                                                                                                        a ++ b is the result of concatenation of a and b, usually read "append". The meaning of this notation is type-dependent.

                                                                                                                                                      Instances
                                                                                                                                                        class HOrElse (α : Type u) (β : Type v) (γ : outParam (Type w)) :
                                                                                                                                                        Type (max (max u v) w)

                                                                                                                                                        The typeclass behind the notation a <|> b : γ where a : α, b : β. Because b is "lazy" in this notation, it is passed as Unit → β to the implementation so it can decide when to evaluate it.

                                                                                                                                                        • hOrElse : α(Unitβ)γ

                                                                                                                                                          a <|> b executes a and returns the result, unless it fails in which case it executes and returns b. Because b is not always executed, it is passed as a thunk so it can be forced only when needed. The meaning of this notation is type-dependent.

                                                                                                                                                        Instances
                                                                                                                                                          class HAndThen (α : Type u) (β : Type v) (γ : outParam (Type w)) :
                                                                                                                                                          Type (max (max u v) w)

                                                                                                                                                          The typeclass behind the notation a >> b : γ where a : α, b : β. Because b is "lazy" in this notation, it is passed as Unit → β to the implementation so it can decide when to evaluate it.

                                                                                                                                                          • hAndThen : α(Unitβ)γ

                                                                                                                                                            a >> b executes a, ignores the result, and then executes b. If a fails then b is not executed. Because b is not always executed, it is passed as a thunk so it can be forced only when needed. The meaning of this notation is type-dependent.

                                                                                                                                                          Instances
                                                                                                                                                            class HAnd (α : Type u) (β : Type v) (γ : outParam (Type w)) :
                                                                                                                                                            Type (max (max u v) w)

                                                                                                                                                            The typeclass behind the notation a &&& b : γ where a : α, b : β.

                                                                                                                                                            • hAnd : αβγ

                                                                                                                                                              a &&& b computes the bitwise AND of a and b. The meaning of this notation is type-dependent.

                                                                                                                                                            Instances
                                                                                                                                                              class HXor (α : Type u) (β : Type v) (γ : outParam (Type w)) :
                                                                                                                                                              Type (max (max u v) w)

                                                                                                                                                              The typeclass behind the notation a ^^^ b : γ where a : α, b : β.

                                                                                                                                                              • hXor : αβγ

                                                                                                                                                                a ^^^ b computes the bitwise XOR of a and b. The meaning of this notation is type-dependent.

                                                                                                                                                              Instances
                                                                                                                                                                class HOr (α : Type u) (β : Type v) (γ : outParam (Type w)) :
                                                                                                                                                                Type (max (max u v) w)

                                                                                                                                                                The typeclass behind the notation a ||| b : γ where a : α, b : β.

                                                                                                                                                                • hOr : αβγ

                                                                                                                                                                  a ||| b computes the bitwise OR of a and b. The meaning of this notation is type-dependent.

                                                                                                                                                                Instances
                                                                                                                                                                  class HShiftLeft (α : Type u) (β : Type v) (γ : outParam (Type w)) :
                                                                                                                                                                  Type (max (max u v) w)

                                                                                                                                                                  The typeclass behind the notation a <<< b : γ where a : α, b : β.

                                                                                                                                                                  • hShiftLeft : αβγ

                                                                                                                                                                    a <<< b computes a shifted to the left by b places. The meaning of this notation is type-dependent.

                                                                                                                                                                    • On Nat, this is equivalent to a * 2 ^ b.
                                                                                                                                                                    • On UInt8 and other fixed width unsigned types, this is the same but truncated to the bit width.
                                                                                                                                                                  Instances
                                                                                                                                                                    class HShiftRight (α : Type u) (β : Type v) (γ : outParam (Type w)) :
                                                                                                                                                                    Type (max (max u v) w)

                                                                                                                                                                    The typeclass behind the notation a >>> b : γ where a : α, b : β.

                                                                                                                                                                    • hShiftRight : αβγ

                                                                                                                                                                      a >>> b computes a shifted to the right by b places. The meaning of this notation is type-dependent.

                                                                                                                                                                      • On Nat and fixed width unsigned types like UInt8, this is equivalent to a / 2 ^ b.
                                                                                                                                                                    Instances
                                                                                                                                                                      class Add (α : Type u) :

                                                                                                                                                                      The homogeneous version of HAdd: a + b : α where a b : α.

                                                                                                                                                                      • add : ααα

                                                                                                                                                                        a + b computes the sum of a and b. See HAdd.

                                                                                                                                                                      Instances
                                                                                                                                                                        class Sub (α : Type u) :

                                                                                                                                                                        The homogeneous version of HSub: a - b : α where a b : α.

                                                                                                                                                                        • sub : ααα

                                                                                                                                                                          a - b computes the difference of a and b. See HSub.

                                                                                                                                                                        Instances
                                                                                                                                                                          class Mul (α : Type u) :

                                                                                                                                                                          The homogeneous version of HMul: a * b : α where a b : α.

                                                                                                                                                                          • mul : ααα

                                                                                                                                                                            a * b computes the product of a and b. See HMul.

                                                                                                                                                                          Instances
                                                                                                                                                                            class Neg (α : Type u) :

                                                                                                                                                                            The notation typeclass for negation. This enables the notation -a : α where a : α.

                                                                                                                                                                            • neg : αα

                                                                                                                                                                              -a computes the negative or opposite of a. The meaning of this notation is type-dependent.

                                                                                                                                                                            Instances
                                                                                                                                                                              class Div (α : Type u) :

                                                                                                                                                                              The homogeneous version of HDiv: a / b : α where a b : α.

                                                                                                                                                                              • div : ααα

                                                                                                                                                                                a / b computes the result of dividing a by b. See HDiv.

                                                                                                                                                                              Instances
                                                                                                                                                                                class Mod (α : Type u) :

                                                                                                                                                                                The homogeneous version of HMod: a % b : α where a b : α.

                                                                                                                                                                                • mod : ααα

                                                                                                                                                                                  a % b computes the remainder upon dividing a by b. See HMod.

                                                                                                                                                                                Instances
                                                                                                                                                                                  class Dvd (α : Type u_1) :
                                                                                                                                                                                  Type u_1

                                                                                                                                                                                  Notation typeclass for the operation (typed as \|), which represents divisibility.

                                                                                                                                                                                  • dvd : ααProp

                                                                                                                                                                                    Divisibility. a ∣ b (typed as \|) means that there is some c such that b = a * c.

                                                                                                                                                                                  Instances
                                                                                                                                                                                    class Pow (α : Type u) (β : Type v) :
                                                                                                                                                                                    Type (max u v)

                                                                                                                                                                                    The homogeneous version of HPow: a ^ b : α where a : α, b : β. (The right argument is not the same as the left since we often want this even in the homogeneous case.)

                                                                                                                                                                                    Types can choose to subscribe to particular defaulting behavior by providing an instance to either NatPow or HomogeneousPow:

                                                                                                                                                                                    • NatPow is for types whose exponents is preferentially a Nat.
                                                                                                                                                                                    • HomogeneousPow is for types whose base and exponent are preferentially the same.
                                                                                                                                                                                    • pow : αβα

                                                                                                                                                                                      a ^ b computes a to the power of b. See HPow.

                                                                                                                                                                                    Instances
                                                                                                                                                                                      class NatPow (α : Type u) :

                                                                                                                                                                                      The homogenous version of Pow where the exponent is a Nat. The purpose of this class is that it provides a default Pow instance, which can be used to specialize the exponent to Nat during elaboration.

                                                                                                                                                                                      For example, if x ^ 2 should preferentially elaborate with 2 : Nat then x's type should provide an instance for this class.

                                                                                                                                                                                      • pow : αNatα

                                                                                                                                                                                        a ^ n computes a to the power of n where n : Nat. See Pow.

                                                                                                                                                                                      Instances
                                                                                                                                                                                        class HomogeneousPow (α : Type u) :

                                                                                                                                                                                        The completely homogeneous version of Pow where the exponent has the same type as the base. The purpose of this class is that it provides a default Pow instance, which can be used to specialize the exponent to have the same type as the base's type during elaboration. This is to say, a type should provide an instance for this class in case x ^ y should be elaborated with both x and y having the same type.

                                                                                                                                                                                        For example, the Float type provides an instance of this class, which causes expressions such as (2.2 ^ 2.2 : Float) to elaborate.

                                                                                                                                                                                        • pow : ααα

                                                                                                                                                                                          a ^ b computes a to the power of b where a and b both have the same type.

                                                                                                                                                                                        Instances
                                                                                                                                                                                          class Append (α : Type u) :

                                                                                                                                                                                          The homogeneous version of HAppend: a ++ b : α where a b : α.

                                                                                                                                                                                          • append : ααα

                                                                                                                                                                                            a ++ b is the result of concatenation of a and b. See HAppend.

                                                                                                                                                                                          Instances
                                                                                                                                                                                            class OrElse (α : Type u) :

                                                                                                                                                                                            The homogeneous version of HOrElse: a <|> b : α where a b : α. Because b is "lazy" in this notation, it is passed as Unit → α to the implementation so it can decide when to evaluate it.

                                                                                                                                                                                            • orElse : α(Unitα)α

                                                                                                                                                                                              The implementation of a <|> b : α. See HOrElse.

                                                                                                                                                                                            Instances
                                                                                                                                                                                              class AndThen (α : Type u) :

                                                                                                                                                                                              The homogeneous version of HAndThen: a >> b : α where a b : α. Because b is "lazy" in this notation, it is passed as Unit → α to the implementation so it can decide when to evaluate it.

                                                                                                                                                                                              • andThen : α(Unitα)α

                                                                                                                                                                                                The implementation of a >> b : α. See HAndThen.

                                                                                                                                                                                              Instances
                                                                                                                                                                                                class AndOp (α : Type u) :

                                                                                                                                                                                                The homogeneous version of HAnd: a &&& b : α where a b : α. (It is called AndOp because And is taken for the propositional connective.)

                                                                                                                                                                                                • and : ααα

                                                                                                                                                                                                  The implementation of a &&& b : α. See HAnd.

                                                                                                                                                                                                Instances
                                                                                                                                                                                                  class Xor (α : Type u) :

                                                                                                                                                                                                  The homogeneous version of HXor: a ^^^ b : α where a b : α.

                                                                                                                                                                                                  • xor : ααα

                                                                                                                                                                                                    The implementation of a ^^^ b : α. See HXor.

                                                                                                                                                                                                  Instances
                                                                                                                                                                                                    class OrOp (α : Type u) :

                                                                                                                                                                                                    The homogeneous version of HOr: a ||| b : α where a b : α. (It is called OrOp because Or is taken for the propositional connective.)

                                                                                                                                                                                                    • or : ααα

                                                                                                                                                                                                      The implementation of a ||| b : α. See HOr.

                                                                                                                                                                                                    Instances
                                                                                                                                                                                                      class Complement (α : Type u) :

                                                                                                                                                                                                      The typeclass behind the notation ~~~a : α where a : α.

                                                                                                                                                                                                      • complement : αα

                                                                                                                                                                                                        The implementation of ~~~a : α.

                                                                                                                                                                                                      Instances
                                                                                                                                                                                                        class ShiftLeft (α : Type u) :

                                                                                                                                                                                                        The homogeneous version of HShiftLeft: a <<< b : α where a b : α.

                                                                                                                                                                                                        • shiftLeft : ααα

                                                                                                                                                                                                          The implementation of a <<< b : α. See HShiftLeft.

                                                                                                                                                                                                        Instances
                                                                                                                                                                                                          class ShiftRight (α : Type u) :

                                                                                                                                                                                                          The homogeneous version of HShiftRight: a >>> b : α where a b : α.

                                                                                                                                                                                                          • shiftRight : ααα

                                                                                                                                                                                                            The implementation of a >>> b : α. See HShiftRight.

                                                                                                                                                                                                          Instances
                                                                                                                                                                                                            @[defaultInstance 1000]
                                                                                                                                                                                                            instance instHAdd {α : Type u_1} [Add α] :
                                                                                                                                                                                                            HAdd α α α
                                                                                                                                                                                                            Equations
                                                                                                                                                                                                            • instHAdd = { hAdd := fun (a b : α) => Add.add a b }
                                                                                                                                                                                                            @[defaultInstance 1000]
                                                                                                                                                                                                            instance instHSub {α : Type u_1} [Sub α] :
                                                                                                                                                                                                            HSub α α α
                                                                                                                                                                                                            Equations
                                                                                                                                                                                                            • instHSub = { hSub := fun (a b : α) => Sub.sub a b }
                                                                                                                                                                                                            @[defaultInstance 1000]
                                                                                                                                                                                                            instance instHMul {α : Type u_1} [Mul α] :
                                                                                                                                                                                                            HMul α α α
                                                                                                                                                                                                            Equations
                                                                                                                                                                                                            • instHMul = { hMul := fun (a b : α) => Mul.mul a b }
                                                                                                                                                                                                            @[defaultInstance 1000]
                                                                                                                                                                                                            instance instHDiv {α : Type u_1} [Div α] :
                                                                                                                                                                                                            HDiv α α α
                                                                                                                                                                                                            Equations
                                                                                                                                                                                                            • instHDiv = { hDiv := fun (a b : α) => Div.div a b }
                                                                                                                                                                                                            @[defaultInstance 1000]
                                                                                                                                                                                                            instance instHMod {α : Type u_1} [Mod α] :
                                                                                                                                                                                                            HMod α α α
                                                                                                                                                                                                            Equations
                                                                                                                                                                                                            • instHMod = { hMod := fun (a b : α) => Mod.mod a b }
                                                                                                                                                                                                            @[defaultInstance 1000]
                                                                                                                                                                                                            instance instHPow {α : Type u_1} {β : Type u_2} [Pow α β] :
                                                                                                                                                                                                            HPow α β α
                                                                                                                                                                                                            Equations
                                                                                                                                                                                                            • instHPow = { hPow := fun (a : α) (b : β) => Pow.pow a b }
                                                                                                                                                                                                            @[defaultInstance 1000]
                                                                                                                                                                                                            instance instPowNat {α : Type u_1} [NatPow α] :
                                                                                                                                                                                                            Pow α Nat
                                                                                                                                                                                                            Equations
                                                                                                                                                                                                            @[defaultInstance 1000]
                                                                                                                                                                                                            instance instPow {α : Type u_1} [HomogeneousPow α] :
                                                                                                                                                                                                            Pow α α
                                                                                                                                                                                                            Equations
                                                                                                                                                                                                            @[defaultInstance 1000]
                                                                                                                                                                                                            instance instHAppend {α : Type u_1} [Append α] :
                                                                                                                                                                                                            HAppend α α α
                                                                                                                                                                                                            Equations
                                                                                                                                                                                                            @[defaultInstance 1000]
                                                                                                                                                                                                            instance instHOrElse {α : Type u_1} [OrElse α] :
                                                                                                                                                                                                            HOrElse α α α
                                                                                                                                                                                                            Equations
                                                                                                                                                                                                            @[defaultInstance 1000]
                                                                                                                                                                                                            instance instHAndThen {α : Type u_1} [AndThen α] :
                                                                                                                                                                                                            HAndThen α α α
                                                                                                                                                                                                            Equations
                                                                                                                                                                                                            @[defaultInstance 1000]
                                                                                                                                                                                                            instance instHAnd {α : Type u_1} [AndOp α] :
                                                                                                                                                                                                            HAnd α α α
                                                                                                                                                                                                            Equations
                                                                                                                                                                                                            • instHAnd = { hAnd := fun (a b : α) => AndOp.and a b }
                                                                                                                                                                                                            @[defaultInstance 1000]
                                                                                                                                                                                                            instance instHXor {α : Type u_1} [Xor α] :
                                                                                                                                                                                                            HXor α α α
                                                                                                                                                                                                            Equations
                                                                                                                                                                                                            • instHXor = { hXor := fun (a b : α) => Xor.xor a b }
                                                                                                                                                                                                            @[defaultInstance 1000]
                                                                                                                                                                                                            instance instHOr {α : Type u_1} [OrOp α] :
                                                                                                                                                                                                            HOr α α α
                                                                                                                                                                                                            Equations
                                                                                                                                                                                                            • instHOr = { hOr := fun (a b : α) => OrOp.or a b }
                                                                                                                                                                                                            @[defaultInstance 1000]
                                                                                                                                                                                                            instance instHShiftLeft {α : Type u_1} [ShiftLeft α] :
                                                                                                                                                                                                            HShiftLeft α α α
                                                                                                                                                                                                            Equations
                                                                                                                                                                                                            @[defaultInstance 1000]
                                                                                                                                                                                                            instance instHShiftRight {α : Type u_1} [ShiftRight α] :
                                                                                                                                                                                                            HShiftRight α α α
                                                                                                                                                                                                            Equations
                                                                                                                                                                                                            class Membership (α : outParam (Type u)) (γ : Type v) :
                                                                                                                                                                                                            Type (max u v)

                                                                                                                                                                                                            The typeclass behind the notation a ∈ s : Prop where a : α, s : γ. Because α is an outParam, the "container type" γ determines the type of the elements of the container.

                                                                                                                                                                                                            • mem : αγProp

                                                                                                                                                                                                              The membership relation a ∈ s : Prop where a : α, s : γ.

                                                                                                                                                                                                            Instances
                                                                                                                                                                                                              @[match_pattern, extern lean_nat_add]
                                                                                                                                                                                                              def Nat.add :
                                                                                                                                                                                                              NatNatNat

                                                                                                                                                                                                              Addition of natural numbers.

                                                                                                                                                                                                              This definition is overridden in both the kernel and the compiler to efficiently evaluate using the "bignum" representation (see Nat). The definition provided here is the logical model (and it is soundness-critical that they coincide).

                                                                                                                                                                                                              Equations
                                                                                                                                                                                                              Instances For
                                                                                                                                                                                                                instance instAddNat :
                                                                                                                                                                                                                Equations
                                                                                                                                                                                                                @[extern lean_nat_mul]
                                                                                                                                                                                                                def Nat.mul :
                                                                                                                                                                                                                NatNatNat

                                                                                                                                                                                                                Multiplication of natural numbers.

                                                                                                                                                                                                                This definition is overridden in both the kernel and the compiler to efficiently evaluate using the "bignum" representation (see Nat). The definition provided here is the logical model (and it is soundness-critical that they coincide).

                                                                                                                                                                                                                Equations
                                                                                                                                                                                                                Instances For
                                                                                                                                                                                                                  instance instMulNat :
                                                                                                                                                                                                                  Equations
                                                                                                                                                                                                                  @[extern lean_nat_pow]
                                                                                                                                                                                                                  def Nat.pow (m : Nat) :
                                                                                                                                                                                                                  NatNat

                                                                                                                                                                                                                  The power operation on natural numbers.

                                                                                                                                                                                                                  This definition is overridden in both the kernel and the compiler to efficiently evaluate using the "bignum" representation (see Nat). The definition provided here is the logical model.

                                                                                                                                                                                                                  Equations
                                                                                                                                                                                                                  Instances For
                                                                                                                                                                                                                    Equations
                                                                                                                                                                                                                    @[extern lean_nat_dec_eq]
                                                                                                                                                                                                                    def Nat.beq :
                                                                                                                                                                                                                    NatNatBool

                                                                                                                                                                                                                    (Boolean) equality of natural numbers.

                                                                                                                                                                                                                    This definition is overridden in both the kernel and the compiler to efficiently evaluate using the "bignum" representation (see Nat). The definition provided here is the logical model (and it is soundness-critical that they coincide).

                                                                                                                                                                                                                    Equations
                                                                                                                                                                                                                    Instances For
                                                                                                                                                                                                                      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
                                                                                                                                                                                                                      @[reducible, extern lean_nat_dec_eq]
                                                                                                                                                                                                                      def Nat.decEq (n : Nat) (m : Nat) :
                                                                                                                                                                                                                      Decidable (n = m)

                                                                                                                                                                                                                      A decision procedure for equality of natural numbers.

                                                                                                                                                                                                                      This definition is overridden in the compiler to efficiently evaluate using the "bignum" representation (see Nat). The definition provided here is the logical model.

                                                                                                                                                                                                                      Equations
                                                                                                                                                                                                                      Instances For
                                                                                                                                                                                                                        @[extern lean_nat_dec_le]
                                                                                                                                                                                                                        def Nat.ble :
                                                                                                                                                                                                                        NatNatBool

                                                                                                                                                                                                                        The (Boolean) less-equal relation on natural numbers.

                                                                                                                                                                                                                        This definition is overridden in both the kernel and the compiler to efficiently evaluate using the "bignum" representation (see Nat). The definition provided here is the logical model (and it is soundness-critical that they coincide).

                                                                                                                                                                                                                        Equations
                                                                                                                                                                                                                        Instances For
                                                                                                                                                                                                                          inductive Nat.le (n : Nat) :
                                                                                                                                                                                                                          NatProp

                                                                                                                                                                                                                          An inductive definition of the less-equal relation on natural numbers, characterized as the least relation such that n ≤ n and n ≤ m → n ≤ m + 1.

                                                                                                                                                                                                                          Instances For
                                                                                                                                                                                                                            instance instLENat :
                                                                                                                                                                                                                            Equations
                                                                                                                                                                                                                            def Nat.lt (n : Nat) (m : Nat) :

                                                                                                                                                                                                                            The strict less than relation on natural numbers is defined as n < m := n + 1 ≤ m.

                                                                                                                                                                                                                            Equations
                                                                                                                                                                                                                            Instances For
                                                                                                                                                                                                                              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
                                                                                                                                                                                                                              @[simp]
                                                                                                                                                                                                                              theorem Nat.zero_lt_succ (n : Nat) :
                                                                                                                                                                                                                              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 lean_nat_pred]
                                                                                                                                                                                                                              def Nat.pred :
                                                                                                                                                                                                                              NatNat

                                                                                                                                                                                                                              The predecessor function on natural numbers.

                                                                                                                                                                                                                              This definition is overridden in the compiler to use n - 1 instead. The definition provided here is the logical model.

                                                                                                                                                                                                                              Equations
                                                                                                                                                                                                                              Instances For
                                                                                                                                                                                                                                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
                                                                                                                                                                                                                                instance instMinNat :
                                                                                                                                                                                                                                Equations
                                                                                                                                                                                                                                @[extern lean_nat_sub]
                                                                                                                                                                                                                                def Nat.sub :
                                                                                                                                                                                                                                NatNatNat

                                                                                                                                                                                                                                (Truncated) subtraction of natural numbers. Because natural numbers are not closed under subtraction, we define n - m to be 0 when n < m.

                                                                                                                                                                                                                                This definition is overridden in both the kernel and the compiler to efficiently evaluate using the "bignum" representation (see Nat). The definition provided here is the logical model (and it is soundness-critical that they coincide).

                                                                                                                                                                                                                                Equations
                                                                                                                                                                                                                                Instances For
                                                                                                                                                                                                                                  instance instSubNat :
                                                                                                                                                                                                                                  Equations
                                                                                                                                                                                                                                  @[extern lean_system_platform_nbits]
                                                                                                                                                                                                                                  opaque System.Platform.getNumBits :
                                                                                                                                                                                                                                  Unit{ n : Nat // n = 32 n = 64 }

                                                                                                                                                                                                                                  Gets the word size of the platform. That is, whether the platform is 64 or 32 bits.

                                                                                                                                                                                                                                  This function is opaque because we cannot guarantee at compile time that the target will have the same size as the host, and also because we would like to avoid typechecking being architecture-dependent. Nevertheless, Lean only works on 64 and 32 bit systems so we can encode this as a fact available for proof purposes.

                                                                                                                                                                                                                                  Gets the word size of the platform. That is, whether the platform is 64 or 32 bits.

                                                                                                                                                                                                                                  Equations
                                                                                                                                                                                                                                  Instances For
                                                                                                                                                                                                                                    structure Fin (n : Nat) :

                                                                                                                                                                                                                                    Fin n is a natural number i with the constraint that 0 ≤ i < n. It is the "canonical type with n elements".

                                                                                                                                                                                                                                    • val : Nat

                                                                                                                                                                                                                                      If i : Fin n, then i.val : ℕ is the described number. It can also be written as i.1 or just i when the target type is known.

                                                                                                                                                                                                                                    • isLt : self < n

                                                                                                                                                                                                                                      If i : Fin n, then i.2 is a proof that i.1 < n.

                                                                                                                                                                                                                                    Instances For
                                                                                                                                                                                                                                      theorem Fin.eq_of_val_eq {n : Nat} {i : Fin n} {j : Fin n} :
                                                                                                                                                                                                                                      i = ji = j
                                                                                                                                                                                                                                      theorem Fin.val_eq_of_eq {n : Nat} {i : Fin n} {j : Fin n} (h : i = j) :
                                                                                                                                                                                                                                      i = j
                                                                                                                                                                                                                                      theorem Fin.ne_of_val_ne {n : Nat} {i : Fin n} {j : Fin n} (h : ¬i = j) :
                                                                                                                                                                                                                                      ¬i = j
                                                                                                                                                                                                                                      Equations
                                                                                                                                                                                                                                      instance instLTFin {n : Nat} :
                                                                                                                                                                                                                                      LT (Fin n)
                                                                                                                                                                                                                                      Equations
                                                                                                                                                                                                                                      • instLTFin = { lt := fun (a b : Fin n) => a < b }
                                                                                                                                                                                                                                      instance instLEFin {n : Nat} :
                                                                                                                                                                                                                                      LE (Fin n)
                                                                                                                                                                                                                                      Equations
                                                                                                                                                                                                                                      • instLEFin = { le := fun (a b : Fin n) => a b }
                                                                                                                                                                                                                                      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
                                                                                                                                                                                                                                      @[inline, reducible]
                                                                                                                                                                                                                                      abbrev UInt8.size :

                                                                                                                                                                                                                                      The size of type UInt8, that is, 2^8 = 256.

                                                                                                                                                                                                                                      Equations
                                                                                                                                                                                                                                      Instances For
                                                                                                                                                                                                                                        structure UInt8 :

                                                                                                                                                                                                                                        The type of unsigned 8-bit integers. This type has special support in the compiler to make it actually 8 bits rather than wrapping a Nat.

                                                                                                                                                                                                                                        • Unpack a UInt8 as a Nat less than 2^8. This function is overridden with a native implementation.

                                                                                                                                                                                                                                        Instances For
                                                                                                                                                                                                                                          @[extern lean_uint8_of_nat]
                                                                                                                                                                                                                                          def UInt8.ofNatCore (n : Nat) (h : n < UInt8.size) :

                                                                                                                                                                                                                                          Pack a Nat less than 2^8 into a UInt8. This function is overridden with a native implementation.

                                                                                                                                                                                                                                          Equations
                                                                                                                                                                                                                                          Instances For
                                                                                                                                                                                                                                            @[extern lean_uint8_dec_eq]
                                                                                                                                                                                                                                            def UInt8.decEq (a : UInt8) (b : UInt8) :
                                                                                                                                                                                                                                            Decidable (a = b)

                                                                                                                                                                                                                                            Decides equality on UInt8. This function is overridden with a native implementation.

                                                                                                                                                                                                                                            Equations
                                                                                                                                                                                                                                            Instances For
                                                                                                                                                                                                                                              @[inline, reducible]

                                                                                                                                                                                                                                              The size of type UInt16, that is, 2^16 = 65536.

                                                                                                                                                                                                                                              Equations
                                                                                                                                                                                                                                              Instances For
                                                                                                                                                                                                                                                structure UInt16 :

                                                                                                                                                                                                                                                The type of unsigned 16-bit integers. This type has special support in the compiler to make it actually 16 bits rather than wrapping a Nat.

                                                                                                                                                                                                                                                • Unpack a UInt16 as a Nat less than 2^16. This function is overridden with a native implementation.

                                                                                                                                                                                                                                                Instances For
                                                                                                                                                                                                                                                  @[extern lean_uint16_of_nat]

                                                                                                                                                                                                                                                  Pack a Nat less than 2^16 into a UInt16. This function is overridden with a native implementation.

                                                                                                                                                                                                                                                  Equations
                                                                                                                                                                                                                                                  Instances For
                                                                                                                                                                                                                                                    @[extern lean_uint16_dec_eq]
                                                                                                                                                                                                                                                    def UInt16.decEq (a : UInt16) (b : UInt16) :
                                                                                                                                                                                                                                                    Decidable (a = b)

                                                                                                                                                                                                                                                    Decides equality on UInt16. This function is overridden with a native implementation.

                                                                                                                                                                                                                                                    Equations
                                                                                                                                                                                                                                                    Instances For
                                                                                                                                                                                                                                                      @[inline, reducible]

                                                                                                                                                                                                                                                      The size of type UInt32, that is, 2^32 = 4294967296.

                                                                                                                                                                                                                                                      Equations
                                                                                                                                                                                                                                                      Instances For
                                                                                                                                                                                                                                                        structure UInt32 :

                                                                                                                                                                                                                                                        The type of unsigned 32-bit integers. This type has special support in the compiler to make it actually 32 bits rather than wrapping a Nat.

                                                                                                                                                                                                                                                        • Unpack a UInt32 as a Nat less than 2^32. This function is overridden with a native implementation.

                                                                                                                                                                                                                                                        Instances For
                                                                                                                                                                                                                                                          @[extern lean_uint32_of_nat]

                                                                                                                                                                                                                                                          Pack a Nat less than 2^32 into a UInt32. This function is overridden with a native implementation.

                                                                                                                                                                                                                                                          Equations
                                                                                                                                                                                                                                                          Instances For
                                                                                                                                                                                                                                                            @[extern lean_uint32_to_nat]

                                                                                                                                                                                                                                                            Unpack a UInt32 as a Nat. This function is overridden with a native implementation.

                                                                                                                                                                                                                                                            Equations
                                                                                                                                                                                                                                                            Instances For
                                                                                                                                                                                                                                                              @[extern lean_uint32_dec_eq]
                                                                                                                                                                                                                                                              def UInt32.decEq (a : UInt32) (b : UInt32) :
                                                                                                                                                                                                                                                              Decidable (a = b)

                                                                                                                                                                                                                                                              Decides equality on UInt32. This function is overridden with a native implementation.

                                                                                                                                                                                                                                                              Equations
                                                                                                                                                                                                                                                              Instances For
                                                                                                                                                                                                                                                                Equations
                                                                                                                                                                                                                                                                Equations
                                                                                                                                                                                                                                                                @[extern lean_uint32_dec_lt]
                                                                                                                                                                                                                                                                def UInt32.decLt (a : UInt32) (b : UInt32) :
                                                                                                                                                                                                                                                                Decidable (a < b)

                                                                                                                                                                                                                                                                Decides less-equal on UInt32. This function is overridden with a native implementation.

                                                                                                                                                                                                                                                                Equations
                                                                                                                                                                                                                                                                Instances For
                                                                                                                                                                                                                                                                  @[extern lean_uint32_dec_le]
                                                                                                                                                                                                                                                                  def UInt32.decLe (a : UInt32) (b : UInt32) :

                                                                                                                                                                                                                                                                  Decides less-than on UInt32. This function is overridden with a native implementation.

                                                                                                                                                                                                                                                                  Equations
                                                                                                                                                                                                                                                                  Instances For
                                                                                                                                                                                                                                                                    Equations
                                                                                                                                                                                                                                                                    Equations
                                                                                                                                                                                                                                                                    @[inline, reducible]

                                                                                                                                                                                                                                                                    The size of type UInt64, that is, 2^64 = 18446744073709551616.

                                                                                                                                                                                                                                                                    Equations
                                                                                                                                                                                                                                                                    Instances For
                                                                                                                                                                                                                                                                      structure UInt64 :

                                                                                                                                                                                                                                                                      The type of unsigned 64-bit integers. This type has special support in the compiler to make it actually 64 bits rather than wrapping a Nat.

                                                                                                                                                                                                                                                                      • Unpack a UInt64 as a Nat less than 2^64. This function is overridden with a native implementation.

                                                                                                                                                                                                                                                                      Instances For
                                                                                                                                                                                                                                                                        @[extern lean_uint64_of_nat]

                                                                                                                                                                                                                                                                        Pack a Nat less than 2^64 into a UInt64. This function is overridden with a native implementation.

                                                                                                                                                                                                                                                                        Equations
                                                                                                                                                                                                                                                                        Instances For
                                                                                                                                                                                                                                                                          @[extern lean_uint64_dec_eq]
                                                                                                                                                                                                                                                                          def UInt64.decEq (a : UInt64) (b : UInt64) :
                                                                                                                                                                                                                                                                          Decidable (a = b)

                                                                                                                                                                                                                                                                          Decides equality on UInt64. This function is overridden with a native implementation.

                                                                                                                                                                                                                                                                          Equations
                                                                                                                                                                                                                                                                          Instances For
                                                                                                                                                                                                                                                                            @[inline, reducible]
                                                                                                                                                                                                                                                                            abbrev USize.size :

                                                                                                                                                                                                                                                                            The size of type UInt16, that is, 2^System.Platform.numBits, which may be either 2^32 or 2^64 depending on the platform's architecture.

                                                                                                                                                                                                                                                                            Remark: we define USize.size using (2^numBits - 1) + 1 to ensure the Lean unifier can solve contraints such as ?m + 1 = USize.size. Recall that numBits does not reduce to a numeral in the Lean kernel since it is platform specific. Without this trick, the following definition would be rejected by the Lean type checker.

                                                                                                                                                                                                                                                                            def one: Fin USize.size := 1
                                                                                                                                                                                                                                                                            

                                                                                                                                                                                                                                                                            Because Lean would fail to synthesize instance OfNat (Fin USize.size) 1. Recall that the OfNat instance for Fin is

                                                                                                                                                                                                                                                                            instance : OfNat (Fin (n+1)) i where
                                                                                                                                                                                                                                                                              ofNat := Fin.ofNat i
                                                                                                                                                                                                                                                                            
                                                                                                                                                                                                                                                                            Equations
                                                                                                                                                                                                                                                                            Instances For
                                                                                                                                                                                                                                                                              theorem usize_size_eq :
                                                                                                                                                                                                                                                                              USize.size = 4294967296 USize.size = 18446744073709551616
                                                                                                                                                                                                                                                                              structure USize :

                                                                                                                                                                                                                                                                              A USize is an unsigned integer with the size of a word for the platform's architecture.

                                                                                                                                                                                                                                                                              For example, if running on a 32-bit machine, USize is equivalent to UInt32. Or on a 64-bit machine, UInt64.

                                                                                                                                                                                                                                                                              Instances For
                                                                                                                                                                                                                                                                                @[extern lean_usize_of_nat]
                                                                                                                                                                                                                                                                                def USize.ofNatCore (n : Nat) (h : n < USize.size) :

                                                                                                                                                                                                                                                                                Pack a Nat less than USize.size into a USize. This function is overridden with a native implementation.

                                                                                                                                                                                                                                                                                Equations
                                                                                                                                                                                                                                                                                Instances For
                                                                                                                                                                                                                                                                                  @[extern lean_usize_dec_eq]
                                                                                                                                                                                                                                                                                  def USize.decEq (a : USize) (b : USize) :
                                                                                                                                                                                                                                                                                  Decidable (a = b)

                                                                                                                                                                                                                                                                                  Decides equality on USize. This function is overridden with a native implementation.

                                                                                                                                                                                                                                                                                  Equations
                                                                                                                                                                                                                                                                                  Instances For
                                                                                                                                                                                                                                                                                    @[extern lean_usize_of_nat]
                                                                                                                                                                                                                                                                                    def USize.ofNat32 (n : Nat) (h : n < 4294967296) :

                                                                                                                                                                                                                                                                                    Upcast a Nat less than 2^32 to a USize. This is lossless because USize.size is either 2^32 or 2^64. This function is overridden with a native implementation.

                                                                                                                                                                                                                                                                                    Equations
                                                                                                                                                                                                                                                                                    Instances For
                                                                                                                                                                                                                                                                                      @[inline, reducible]
                                                                                                                                                                                                                                                                                      abbrev Nat.isValidChar (n : Nat) :

                                                                                                                                                                                                                                                                                      A Nat denotes a valid unicode codepoint if it is less than 0x110000, and it is also not a "surrogate" character (the range 0xd800 to 0xdfff inclusive).

                                                                                                                                                                                                                                                                                      Equations
                                                                                                                                                                                                                                                                                      Instances For
                                                                                                                                                                                                                                                                                        @[inline, reducible]

                                                                                                                                                                                                                                                                                        A UInt32 denotes a valid unicode codepoint if it is less than 0x110000, and it is also not a "surrogate" character (the range 0xd800 to 0xdfff inclusive).

                                                                                                                                                                                                                                                                                        Equations
                                                                                                                                                                                                                                                                                        Instances For
                                                                                                                                                                                                                                                                                          structure Char :

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

                                                                                                                                                                                                                                                                                          Instances For
                                                                                                                                                                                                                                                                                            @[extern lean_uint32_of_nat]

                                                                                                                                                                                                                                                                                            Pack a Nat encoding a valid codepoint into a Char. This function is overridden with a native implementation.

                                                                                                                                                                                                                                                                                            Equations
                                                                                                                                                                                                                                                                                            • Char.ofNatAux n h = { val := { val := { val := n, isLt := } }, valid := h }
                                                                                                                                                                                                                                                                                            Instances For
                                                                                                                                                                                                                                                                                              @[match_pattern, noinline]
                                                                                                                                                                                                                                                                                              def Char.ofNat (n : Nat) :

                                                                                                                                                                                                                                                                                              Convert a Nat into a Char. If the Nat does not encode a valid unicode scalar value, '\0' is returned instead.

                                                                                                                                                                                                                                                                                              Equations
                                                                                                                                                                                                                                                                                              Instances For
                                                                                                                                                                                                                                                                                                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

                                                                                                                                                                                                                                                                                                Returns the number of bytes required to encode this Char in UTF-8.

                                                                                                                                                                                                                                                                                                Equations
                                                                                                                                                                                                                                                                                                • One or more equations did not get rendered due to their size.
                                                                                                                                                                                                                                                                                                Instances For
                                                                                                                                                                                                                                                                                                  @[unbox]
                                                                                                                                                                                                                                                                                                  inductive Option (α : Type u) :

                                                                                                                                                                                                                                                                                                  Option α is the type of values which are either some a for some a : α, or none. In functional programming languages, this type is used to represent the possibility of failure, or sometimes nullability.

                                                                                                                                                                                                                                                                                                  For example, the function HashMap.find? : HashMap α β → α → Option β looks up a specified key a : α inside the map. Because we do not know in advance whether the key is actually in the map, the return type is Option β, where none means the value was not in the map, and some b means that the value was found and b is the value retrieved.

                                                                                                                                                                                                                                                                                                  To extract a value from an Option α, we use pattern matching:

                                                                                                                                                                                                                                                                                                  def map (f : α → β) (x : Option α) : Option β :=
                                                                                                                                                                                                                                                                                                    match x with
                                                                                                                                                                                                                                                                                                    | some a => some (f a)
                                                                                                                                                                                                                                                                                                    | none => none
                                                                                                                                                                                                                                                                                                  

                                                                                                                                                                                                                                                                                                  We can also use if let to pattern match on Option and get the value in the branch:

                                                                                                                                                                                                                                                                                                  def map (f : α → β) (x : Option α) : Option β :=
                                                                                                                                                                                                                                                                                                    if let some a := x then
                                                                                                                                                                                                                                                                                                      some (f a)
                                                                                                                                                                                                                                                                                                    else
                                                                                                                                                                                                                                                                                                      none
                                                                                                                                                                                                                                                                                                  
                                                                                                                                                                                                                                                                                                  • none: {α : Type u} → Option α

                                                                                                                                                                                                                                                                                                    No value.

                                                                                                                                                                                                                                                                                                  • some: {α : Type u} → αOption α

                                                                                                                                                                                                                                                                                                    Some value of type α.

                                                                                                                                                                                                                                                                                                  Instances For
                                                                                                                                                                                                                                                                                                    instance instInhabitedOption {α : Type u_1} :
                                                                                                                                                                                                                                                                                                    Equations
                                                                                                                                                                                                                                                                                                    • instInhabitedOption = { default := none }
                                                                                                                                                                                                                                                                                                    @[macro_inline]
                                                                                                                                                                                                                                                                                                    def Option.getD {α : Type u_1} (opt : Option α) (dflt : α) :
                                                                                                                                                                                                                                                                                                    α

                                                                                                                                                                                                                                                                                                    Get with default. If opt : Option α and dflt : α, then opt.getD dflt returns a if opt = some a and dflt otherwise.

                                                                                                                                                                                                                                                                                                    This function is @[macro_inline], so dflt will not be evaluated unless opt turns out to be none.

                                                                                                                                                                                                                                                                                                    Equations
                                                                                                                                                                                                                                                                                                    Instances For
                                                                                                                                                                                                                                                                                                      @[inline]
                                                                                                                                                                                                                                                                                                      def Option.map {α : Type u_1} {β : Type u_2} (f : αβ) :
                                                                                                                                                                                                                                                                                                      Option αOption β

                                                                                                                                                                                                                                                                                                      Map a function over an Option by applying the function to the contained value if present.

                                                                                                                                                                                                                                                                                                      Equations
                                                                                                                                                                                                                                                                                                      Instances For
                                                                                                                                                                                                                                                                                                        inductive List (α : Type u) :

                                                                                                                                                                                                                                                                                                        List α is the type of ordered lists with elements of type α. It is implemented as a linked list.

                                                                                                                                                                                                                                                                                                        List α is isomorphic to Array α, but they are useful for different things:

                                                                                                                                                                                                                                                                                                        • List α is easier for reasoning, and Array α is modeled as a wrapper around List α
                                                                                                                                                                                                                                                                                                        • List α works well as a persistent data structure, when many copies of the tail are shared. When the value is not shared, Array α will have better performance because it can do destructive updates.
                                                                                                                                                                                                                                                                                                        • nil: {α : Type u} → List α

                                                                                                                                                                                                                                                                                                          [] is the empty list.

                                                                                                                                                                                                                                                                                                        • cons: {α : Type u} → αList αList α

                                                                                                                                                                                                                                                                                                          If a : α and l : List α, then cons a l, or a :: l, is the list whose first element is a and with l as the rest of the list.

                                                                                                                                                                                                                                                                                                        Instances For
                                                                                                                                                                                                                                                                                                          instance instInhabitedList {α : Type u_1} :
                                                                                                                                                                                                                                                                                                          Equations
                                                                                                                                                                                                                                                                                                          • instInhabitedList = { default := [] }
                                                                                                                                                                                                                                                                                                          def List.hasDecEq {α : Type u} [DecidableEq α] (a : List α) (b : List α) :
                                                                                                                                                                                                                                                                                                          Decidable (a = b)

                                                                                                                                                                                                                                                                                                          Implements decidable equality for List α, assuming α has decidable equality.

                                                                                                                                                                                                                                                                                                          Equations
                                                                                                                                                                                                                                                                                                          Instances For
                                                                                                                                                                                                                                                                                                            instance instDecidableEqList {α : Type u} [DecidableEq α] :
                                                                                                                                                                                                                                                                                                            Equations
                                                                                                                                                                                                                                                                                                            • instDecidableEqList = List.hasDecEq
                                                                                                                                                                                                                                                                                                            @[specialize #[]]
                                                                                                                                                                                                                                                                                                            def List.foldl {α : Type u} {β : Type v} (f : αβα) (init : α) :
                                                                                                                                                                                                                                                                                                            List βα

                                                                                                                                                                                                                                                                                                            Folds a function over a list from the left: foldl f z [a, b, c] = f (f (f z a) b) c

                                                                                                                                                                                                                                                                                                            Equations
                                                                                                                                                                                                                                                                                                            Instances For
                                                                                                                                                                                                                                                                                                              def List.set {α : Type u_1} :
                                                                                                                                                                                                                                                                                                              List αNatαList α

                                                                                                                                                                                                                                                                                                              l.set n a sets the value of list l at (zero-based) index n to a: [a, b, c, d].set 1 b' = [a, b', c, d]

                                                                                                                                                                                                                                                                                                              Equations
                                                                                                                                                                                                                                                                                                              Instances For
                                                                                                                                                                                                                                                                                                                def List.length {α : Type u_1} :
                                                                                                                                                                                                                                                                                                                List αNat

                                                                                                                                                                                                                                                                                                                The length of a list: [].length = 0 and (a :: l).length = l.length + 1.

                                                                                                                                                                                                                                                                                                                This function is overridden in the compiler to lengthTR, which uses constant stack space, while leaving this function to use the "naive" recursion which is easier for reasoning.

                                                                                                                                                                                                                                                                                                                Equations
                                                                                                                                                                                                                                                                                                                Instances For
                                                                                                                                                                                                                                                                                                                  def List.lengthTRAux {α : Type u_1} :
                                                                                                                                                                                                                                                                                                                  List αNatNat

                                                                                                                                                                                                                                                                                                                  Auxiliary function for List.lengthTR.

                                                                                                                                                                                                                                                                                                                  Equations
                                                                                                                                                                                                                                                                                                                  Instances For
                                                                                                                                                                                                                                                                                                                    def List.lengthTR {α : Type u_1} (as : List α) :

                                                                                                                                                                                                                                                                                                                    A tail-recursive version of List.length, used to implement List.length without running out of stack space.

                                                                                                                                                                                                                                                                                                                    Equations
                                                                                                                                                                                                                                                                                                                    Instances For
                                                                                                                                                                                                                                                                                                                      @[simp]
                                                                                                                                                                                                                                                                                                                      theorem List.length_cons {α : Type u_1} (a : α) (as : List α) :
                                                                                                                                                                                                                                                                                                                      def List.concat {α : Type u} :
                                                                                                                                                                                                                                                                                                                      List ααList α

                                                                                                                                                                                                                                                                                                                      l.concat a appends a at the end of l, that is, l ++ [a].

                                                                                                                                                                                                                                                                                                                      Equations
                                                                                                                                                                                                                                                                                                                      Instances For
                                                                                                                                                                                                                                                                                                                        def List.get {α : Type u} (as : List α) :
                                                                                                                                                                                                                                                                                                                        Fin (List.length as)α

                                                                                                                                                                                                                                                                                                                        as.get i returns the i'th element of the list as. This version of the function uses i : Fin as.length to ensure that it will not index out of bounds.

                                                                                                                                                                                                                                                                                                                        Equations
                                                                                                                                                                                                                                                                                                                        Instances For
                                                                                                                                                                                                                                                                                                                          structure String :

                                                                                                                                                                                                                                                                                                                          String is the type of (UTF-8 encoded) strings.

                                                                                                                                                                                                                                                                                                                          The compiler overrides the data representation of this type to a byte sequence, and both String.utf8ByteSize and String.length are cached and O(1).

                                                                                                                                                                                                                                                                                                                          • data : List Char

                                                                                                                                                                                                                                                                                                                            Unpack String into a List Char. This function is overridden by the compiler and is O(n) in the length of the list.

                                                                                                                                                                                                                                                                                                                          Instances For
                                                                                                                                                                                                                                                                                                                            @[extern lean_string_dec_eq]
                                                                                                                                                                                                                                                                                                                            def String.decEq (s₁ : String) (s₂ : String) :
                                                                                                                                                                                                                                                                                                                            Decidable (s₁ = s₂)

                                                                                                                                                                                                                                                                                                                            Decides equality on String. This function is overridden with a native implementation.

                                                                                                                                                                                                                                                                                                                            Equations
                                                                                                                                                                                                                                                                                                                            • String.decEq s₁ s₂ = match s₁, s₂ with | { data := s₁ }, { data := s₂ } => if h : s₁ = s₂ then isTrue else isFalse
                                                                                                                                                                                                                                                                                                                            Instances For
                                                                                                                                                                                                                                                                                                                              structure String.Pos :

                                                                                                                                                                                                                                                                                                                              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.

                                                                                                                                                                                                                                                                                                                              A byte position p is valid for a string s if 0 ≤ p ≤ s.endPos and p lies on a UTF8 byte boundary.

                                                                                                                                                                                                                                                                                                                              Instances For
                                                                                                                                                                                                                                                                                                                                Equations
                                                                                                                                                                                                                                                                                                                                Equations
                                                                                                                                                                                                                                                                                                                                • One or more equations did not get rendered due to their size.
                                                                                                                                                                                                                                                                                                                                structure Substring :

                                                                                                                                                                                                                                                                                                                                A Substring is a view into some subslice of a String. The actual string slicing is deferred because this would require copying the string; here we only store a reference to the original string for garbage collection purposes.

                                                                                                                                                                                                                                                                                                                                • str : String

                                                                                                                                                                                                                                                                                                                                  The underlying string to slice.

                                                                                                                                                                                                                                                                                                                                • startPos : String.Pos

                                                                                                                                                                                                                                                                                                                                  The byte position of the start of the string slice.

                                                                                                                                                                                                                                                                                                                                • stopPos : String.Pos

                                                                                                                                                                                                                                                                                                                                  The byte position of the end of the string slice.

                                                                                                                                                                                                                                                                                                                                Instances For
                                                                                                                                                                                                                                                                                                                                  Equations
                                                                                                                                                                                                                                                                                                                                  @[inline]

                                                                                                                                                                                                                                                                                                                                  The byte length of the substring.

                                                                                                                                                                                                                                                                                                                                  Equations
                                                                                                                                                                                                                                                                                                                                  Instances For
                                                                                                                                                                                                                                                                                                                                    def String.csize (c : Char) :

                                                                                                                                                                                                                                                                                                                                    Returns the number of bytes required to encode this Char in UTF-8.

                                                                                                                                                                                                                                                                                                                                    Equations
                                                                                                                                                                                                                                                                                                                                    Instances For
                                                                                                                                                                                                                                                                                                                                      @[extern lean_string_utf8_byte_size]

                                                                                                                                                                                                                                                                                                                                      The UTF-8 byte length of this string. This is overridden by the compiler to be cached and O(1).

                                                                                                                                                                                                                                                                                                                                      Equations
                                                                                                                                                                                                                                                                                                                                      Instances For
                                                                                                                                                                                                                                                                                                                                        Equations
                                                                                                                                                                                                                                                                                                                                        Equations
                                                                                                                                                                                                                                                                                                                                        Equations
                                                                                                                                                                                                                                                                                                                                        Equations
                                                                                                                                                                                                                                                                                                                                        Equations
                                                                                                                                                                                                                                                                                                                                        Equations
                                                                                                                                                                                                                                                                                                                                        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]

                                                                                                                                                                                                                                                                                                                                        A String.Pos pointing at the end of this string.

                                                                                                                                                                                                                                                                                                                                        Equations
                                                                                                                                                                                                                                                                                                                                        Instances For
                                                                                                                                                                                                                                                                                                                                          @[inline]

                                                                                                                                                                                                                                                                                                                                          Convert a String into a Substring denoting the entire string.

                                                                                                                                                                                                                                                                                                                                          Equations
                                                                                                                                                                                                                                                                                                                                          Instances For

                                                                                                                                                                                                                                                                                                                                            String.toSubstring without [inline] annotation.

                                                                                                                                                                                                                                                                                                                                            Equations
                                                                                                                                                                                                                                                                                                                                            Instances For
                                                                                                                                                                                                                                                                                                                                              unsafe def unsafeCast {α : Sort u} {β : Sort v} (a : α) :
                                                                                                                                                                                                                                                                                                                                              β

                                                                                                                                                                                                                                                                                                                                              This function will cast a value of type α to type β, and is a no-op in the compiler. This function is extremely dangerous because there is no guarantee that types α and β have the same data representation, and this can lead to memory unsafety. It is also logically unsound, since you could just cast True to False. For all those reasons this function is marked as unsafe.

                                                                                                                                                                                                                                                                                                                                              It is implemented by lifting both α and β into a common universe, and then using cast (lcProof : ULift (PLift α) = ULift (PLift β)) to actually perform the cast. All these operations are no-ops in the compiler.

                                                                                                                                                                                                                                                                                                                                              Using this function correctly requires some knowledge of the data representation of the source and target types. Some general classes of casts which are safe in the current runtime:

                                                                                                                                                                                                                                                                                                                                              • Array α to Array β where α and β have compatible representations, or more generally for other inductive types.
                                                                                                                                                                                                                                                                                                                                              • Quot α r and α.
                                                                                                                                                                                                                                                                                                                                              • @Subtype α p and α, or generally any structure containing only one non-Prop field of type α.
                                                                                                                                                                                                                                                                                                                                              • Casting α to/from NonScalar when α is a boxed generic type (i.e. a function that accepts an arbitrary type α and is not specialized to a scalar type like UInt8).
                                                                                                                                                                                                                                                                                                                                              Equations
                                                                                                                                                                                                                                                                                                                                              Instances For
                                                                                                                                                                                                                                                                                                                                                @[never_extract, extern lean_panic_fn]
                                                                                                                                                                                                                                                                                                                                                def panicCore {α : Type u} [Inhabited α] (msg : String) :
                                                                                                                                                                                                                                                                                                                                                α

                                                                                                                                                                                                                                                                                                                                                Auxiliary definition for panic.

                                                                                                                                                                                                                                                                                                                                                Equations
                                                                                                                                                                                                                                                                                                                                                Instances For
                                                                                                                                                                                                                                                                                                                                                  @[never_extract, noinline]
                                                                                                                                                                                                                                                                                                                                                  def panic {α : Type u} [Inhabited α] (msg : String) :
                                                                                                                                                                                                                                                                                                                                                  α

                                                                                                                                                                                                                                                                                                                                                  (panic "msg" : α) has a built-in implementation which prints msg to the error buffer. It does not terminate execution, and because it is a safe function, it still has to return an element of α, so it takes [Inhabited α] and returns default. It is primarily intended for debugging in pure contexts, and assertion failures.

                                                                                                                                                                                                                                                                                                                                                  Because this is a pure function with side effects, it is marked as @[never_extract] so that the compiler will not perform common sub-expression elimination and other optimizations that assume that the expression is pure.

                                                                                                                                                                                                                                                                                                                                                  Equations
                                                                                                                                                                                                                                                                                                                                                  Instances For
                                                                                                                                                                                                                                                                                                                                                    class GetElem (cont : Type u) (idx : Type v) (elem : outParam (Type w)) (dom : outParam (contidxProp)) :
                                                                                                                                                                                                                                                                                                                                                    Type (max (max u v) w)

                                                                                                                                                                                                                                                                                                                                                    The class GetElem cont idx elem dom implements the xs[i] notation. When you write this, given xs : cont and i : idx, Lean looks for an instance of GetElem cont idx elem dom. Here elem is the type of xs[i], while dom is whatever proof side conditions are required to make this applicable. For example, the instance for arrays looks like GetElem (Array α) Nat α (fun xs i => i < xs.size).

                                                                                                                                                                                                                                                                                                                                                    The proof side-condition dom xs i is automatically dispatched by the get_elem_tactic tactic, which can be extended by adding more clauses to get_elem_tactic_trivial.

                                                                                                                                                                                                                                                                                                                                                    • getElem : (xs : cont) → (i : idx) → dom xs ielem

                                                                                                                                                                                                                                                                                                                                                      The syntax arr[i] gets the i'th element of the collection arr. If there are proof side conditions to the application, they will be automatically inferred by the get_elem_tactic tactic.

                                                                                                                                                                                                                                                                                                                                                      The actual behavior of this class is type-dependent, but here are some important implementations:

                                                                                                                                                                                                                                                                                                                                                      • arr[i] : α where arr : Array α and i : Nat or i : USize: does array indexing with no bounds check and a proof side goal i < arr.size.
                                                                                                                                                                                                                                                                                                                                                      • l[i] : α where l : List α and i : Nat: index into a list, with proof side goal i < l.length.
                                                                                                                                                                                                                                                                                                                                                      • stx[i] : Syntax where stx : Syntax and i : Nat: get a syntax argument, no side goal (returns .missing out of range)

                                                                                                                                                                                                                                                                                                                                                      There are other variations on this syntax:

                                                                                                                                                                                                                                                                                                                                                      • arr[i]: proves the proof side goal by get_elem_tactic
                                                                                                                                                                                                                                                                                                                                                      • arr[i]!: panics if the side goal is false
                                                                                                                                                                                                                                                                                                                                                      • arr[i]?: returns none if the side goal is false
                                                                                                                                                                                                                                                                                                                                                      • arr[i]'h: uses h to prove the side goal
                                                                                                                                                                                                                                                                                                                                                    Instances
                                                                                                                                                                                                                                                                                                                                                      structure Array (α : Type u) :

                                                                                                                                                                                                                                                                                                                                                      Array α is the type of dynamic arrays with elements from α. This type has special support in the runtime.

                                                                                                                                                                                                                                                                                                                                                      An array has a size and a capacity; the size is Array.size but the capacity is not observable from Lean code. Arrays perform best when unshared; as long as they are used "linearly" all updates will be performed destructively on the array, so it has comparable performance to mutable arrays in imperative programming languages.

                                                                                                                                                                                                                                                                                                                                                      From the point of view of proofs Array α is just a wrapper around List α.

                                                                                                                                                                                                                                                                                                                                                      • data : List α

                                                                                                                                                                                                                                                                                                                                                        Converts a Array α into an List α.

                                                                                                                                                                                                                                                                                                                                                        At runtime, this projection is implemented by Array.toList and is O(n) in the length of the array.

                                                                                                                                                                                                                                                                                                                                                      Instances For
                                                                                                                                                                                                                                                                                                                                                        @[extern lean_mk_empty_array_with_capacity]
                                                                                                                                                                                                                                                                                                                                                        def Array.mkEmpty {α : Type u} (c : Nat) :

                                                                                                                                                                                                                                                                                                                                                        Construct a new empty array with initial capacity c.

                                                                                                                                                                                                                                                                                                                                                        Equations
                                                                                                                                                                                                                                                                                                                                                        Instances For
                                                                                                                                                                                                                                                                                                                                                          def Array.empty {α : Type u} :

                                                                                                                                                                                                                                                                                                                                                          Construct a new empty array.

                                                                                                                                                                                                                                                                                                                                                          Equations
                                                                                                                                                                                                                                                                                                                                                          Instances For
                                                                                                                                                                                                                                                                                                                                                            @[reducible, extern lean_array_get_size]
                                                                                                                                                                                                                                                                                                                                                            def Array.size {α : Type u} (a : Array α) :

                                                                                                                                                                                                                                                                                                                                                            Get the size of an array. This is a cached value, so it is O(1) to access.

                                                                                                                                                                                                                                                                                                                                                            Equations
                                                                                                                                                                                                                                                                                                                                                            Instances For
                                                                                                                                                                                                                                                                                                                                                              @[extern lean_array_fget]
                                                                                                                                                                                                                                                                                                                                                              def Array.get {α : Type u} (a : Array α) (i : Fin (Array.size a)) :
                                                                                                                                                                                                                                                                                                                                                              α

                                                                                                                                                                                                                                                                                                                                                              Access an element from an array without bounds checks, using a Fin index.

                                                                                                                                                                                                                                                                                                                                                              Equations
                                                                                                                                                                                                                                                                                                                                                              Instances For
                                                                                                                                                                                                                                                                                                                                                                @[inline, reducible]
                                                                                                                                                                                                                                                                                                                                                                abbrev Array.getD {α : Type u_1} (a : Array α) (i : Nat) (v₀ : α) :
                                                                                                                                                                                                                                                                                                                                                                α

                                                                                                                                                                                                                                                                                                                                                                Access an element from an array, or return v₀ if the index is out of bounds.

                                                                                                                                                                                                                                                                                                                                                                Equations
                                                                                                                                                                                                                                                                                                                                                                Instances For
                                                                                                                                                                                                                                                                                                                                                                  @[extern lean_array_get]
                                                                                                                                                                                                                                                                                                                                                                  def Array.get! {α : Type u} [Inhabited α] (a : Array α) (i : Nat) :
                                                                                                                                                                                                                                                                                                                                                                  α

                                                                                                                                                                                                                                                                                                                                                                  Access an element from an array, or panic if the index is out of bounds.

                                                                                                                                                                                                                                                                                                                                                                  Equations
                                                                                                                                                                                                                                                                                                                                                                  Instances For
                                                                                                                                                                                                                                                                                                                                                                    instance instGetElemArrayNatLtInstLTNatSize {α : Type u_1} :
                                                                                                                                                                                                                                                                                                                                                                    GetElem (Array α) Nat α fun (xs : Array α) (i : Nat) => i < Array.size xs
                                                                                                                                                                                                                                                                                                                                                                    Equations
                                                                                                                                                                                                                                                                                                                                                                    • instGetElemArrayNatLtInstLTNatSize = { getElem := fun (xs : Array α) (i : Nat) (h : i < Array.size xs) => Array.get xs { val := i, isLt := h } }
                                                                                                                                                                                                                                                                                                                                                                    @[extern lean_array_push]
                                                                                                                                                                                                                                                                                                                                                                    def Array.push {α : Type u} (a : Array α) (v : α) :

                                                                                                                                                                                                                                                                                                                                                                    Push an element onto the end of an array. This is amortized O(1) because Array α is internally a dynamic array.

                                                                                                                                                                                                                                                                                                                                                                    Equations
                                                                                                                                                                                                                                                                                                                                                                    Instances For
                                                                                                                                                                                                                                                                                                                                                                      def Array.mkArray0 {α : Type u} :

                                                                                                                                                                                                                                                                                                                                                                      Create array #[]

                                                                                                                                                                                                                                                                                                                                                                      Equations
                                                                                                                                                                                                                                                                                                                                                                      Instances For
                                                                                                                                                                                                                                                                                                                                                                        def Array.mkArray1 {α : Type u} (a₁ : α) :

                                                                                                                                                                                                                                                                                                                                                                        Create array #[a₁]

                                                                                                                                                                                                                                                                                                                                                                        Equations
                                                                                                                                                                                                                                                                                                                                                                        Instances For
                                                                                                                                                                                                                                                                                                                                                                          def Array.mkArray2 {α : Type u} (a₁ : α) (a₂ : α) :

                                                                                                                                                                                                                                                                                                                                                                          Create array #[a₁, a₂]

                                                                                                                                                                                                                                                                                                                                                                          Equations
                                                                                                                                                                                                                                                                                                                                                                          Instances For
                                                                                                                                                                                                                                                                                                                                                                            def Array.mkArray3 {α : Type u} (a₁ : α) (a₂ : α) (a₃ : α) :

                                                                                                                                                                                                                                                                                                                                                                            Create array #[a₁, a₂, a₃]

                                                                                                                                                                                                                                                                                                                                                                            Equations
                                                                                                                                                                                                                                                                                                                                                                            Instances For
                                                                                                                                                                                                                                                                                                                                                                              def Array.mkArray4 {α : Type u} (a₁ : α) (a₂ : α) (a₃ : α) (a₄ : α) :

                                                                                                                                                                                                                                                                                                                                                                              Create array #[a₁, a₂, a₃, a₄]

                                                                                                                                                                                                                                                                                                                                                                              Equations
                                                                                                                                                                                                                                                                                                                                                                              Instances For
                                                                                                                                                                                                                                                                                                                                                                                def Array.mkArray5 {α : Type u} (a₁ : α) (a₂ : α) (a₃ : α) (a₄ : α) (a₅ : α) :

                                                                                                                                                                                                                                                                                                                                                                                Create array #[a₁, a₂, a₃, a₄, a₅]

                                                                                                                                                                                                                                                                                                                                                                                Equations
                                                                                                                                                                                                                                                                                                                                                                                Instances For
                                                                                                                                                                                                                                                                                                                                                                                  def Array.mkArray6 {α : Type u} (a₁ : α) (a₂ : α) (a₃ : α) (a₄ : α) (a₅ : α) (a₆ : α) :

                                                                                                                                                                                                                                                                                                                                                                                  Create array #[a₁, a₂, a₃, a₄, a₅, a₆]

                                                                                                                                                                                                                                                                                                                                                                                  Equations
                                                                                                                                                                                                                                                                                                                                                                                  Instances For
                                                                                                                                                                                                                                                                                                                                                                                    def Array.mkArray7 {α : Type u} (a₁ : α) (a₂ : α) (a₃ : α) (a₄ : α) (a₅ : α) (a₆ : α) (a₇ : α) :

                                                                                                                                                                                                                                                                                                                                                                                    Create array #[a₁, a₂, a₃, a₄, a₅, a₆, a₇]

                                                                                                                                                                                                                                                                                                                                                                                    Equations
                                                                                                                                                                                                                                                                                                                                                                                    Instances For
                                                                                                                                                                                                                                                                                                                                                                                      def Array.mkArray8 {α : Type u} (a₁ : α) (a₂ : α) (a₃ : α) (a₄ : α) (a₅ : α) (a₆ : α) (a₇ : α) (a₈ : α) :

                                                                                                                                                                                                                                                                                                                                                                                      Create array #[a₁, a₂, a₃, a₄, a₅, a₆, a₇, a₈]

                                                                                                                                                                                                                                                                                                                                                                                      Equations
                                                                                                                                                                                                                                                                                                                                                                                      • One or more equations did not get rendered due to their size.
                                                                                                                                                                                                                                                                                                                                                                                      Instances For
                                                                                                                                                                                                                                                                                                                                                                                        @[extern lean_array_fset]
                                                                                                                                                                                                                                                                                                                                                                                        def Array.set {α : Type u_1} (a : Array α) (i : Fin (Array.size a)) (v : α) :

                                                                                                                                                                                                                                                                                                                                                                                        Set an element in an array without bounds checks, using a Fin index.

                                                                                                                                                                                                                                                                                                                                                                                        This will perform the update destructively provided that a has a reference count of 1 when called.

                                                                                                                                                                                                                                                                                                                                                                                        Equations
                                                                                                                                                                                                                                                                                                                                                                                        Instances For
                                                                                                                                                                                                                                                                                                                                                                                          @[inline]
                                                                                                                                                                                                                                                                                                                                                                                          def Array.setD {α : Type u_1} (a : Array α) (i : Nat) (v : α) :

                                                                                                                                                                                                                                                                                                                                                                                          Set an element in an array, or do nothing if the index is out of bounds.

                                                                                                                                                                                                                                                                                                                                                                                          This will perform the update destructively provided that a has a reference count of 1 when called.

                                                                                                                                                                                                                                                                                                                                                                                          Equations
                                                                                                                                                                                                                                                                                                                                                                                          Instances For
                                                                                                                                                                                                                                                                                                                                                                                            @[extern lean_array_set]
                                                                                                                                                                                                                                                                                                                                                                                            def Array.set! {α : Type u_1} (a : Array α) (i : Nat) (v : α) :

                                                                                                                                                                                                                                                                                                                                                                                            Set an element in an array, or panic if the index is out of bounds.

                                                                                                                                                                                                                                                                                                                                                                                            This will perform the update destructively provided that a has a reference count of 1 when called.

                                                                                                                                                                                                                                                                                                                                                                                            Equations
                                                                                                                                                                                                                                                                                                                                                                                            Instances For
                                                                                                                                                                                                                                                                                                                                                                                              def Array.appendCore {α : Type u} (as : Array α) (bs : Array α) :

                                                                                                                                                                                                                                                                                                                                                                                              Slower Array.append used in quotations.

                                                                                                                                                                                                                                                                                                                                                                                              Equations
                                                                                                                                                                                                                                                                                                                                                                                              Instances For
                                                                                                                                                                                                                                                                                                                                                                                                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.
                                                                                                                                                                                                                                                                                                                                                                                                Instances For
                                                                                                                                                                                                                                                                                                                                                                                                  def Array.extract {α : Type u_1} (as : Array α) (start : Nat) (stop : Nat) :

                                                                                                                                                                                                                                                                                                                                                                                                  Returns the slice of as from indices start to stop (exclusive). If start is greater or equal to stop, the result is empty. If stop is greater than the length of as, the length is used instead.

                                                                                                                                                                                                                                                                                                                                                                                                  Equations
                                                                                                                                                                                                                                                                                                                                                                                                  Instances For
                                                                                                                                                                                                                                                                                                                                                                                                    def Array.extract.loop {α : Type u_1} (as : Array α) (i : Nat) (j : Nat) (bs : Array α) :
                                                                                                                                                                                                                                                                                                                                                                                                    Equations
                                                                                                                                                                                                                                                                                                                                                                                                    • One or more equations did not get rendered due to their size.
                                                                                                                                                                                                                                                                                                                                                                                                    Instances For
                                                                                                                                                                                                                                                                                                                                                                                                      @[inline_if_reduce]
                                                                                                                                                                                                                                                                                                                                                                                                      def List.toArrayAux {α : Type u_1} :
                                                                                                                                                                                                                                                                                                                                                                                                      List αArray αArray α

                                                                                                                                                                                                                                                                                                                                                                                                      Auxiliary definition for List.toArray.

                                                                                                                                                                                                                                                                                                                                                                                                      Equations
                                                                                                                                                                                                                                                                                                                                                                                                      Instances For
                                                                                                                                                                                                                                                                                                                                                                                                        @[inline_if_reduce]
                                                                                                                                                                                                                                                                                                                                                                                                        def List.redLength {α : Type u_1} :
                                                                                                                                                                                                                                                                                                                                                                                                        List αNat

                                                                                                                                                                                                                                                                                                                                                                                                        A non-tail-recursive version of List.length, used for List.toArray.

                                                                                                                                                                                                                                                                                                                                                                                                        Equations
                                                                                                                                                                                                                                                                                                                                                                                                        Instances For
                                                                                                                                                                                                                                                                                                                                                                                                          @[match_pattern, inline, export lean_list_to_array]
                                                                                                                                                                                                                                                                                                                                                                                                          def List.toArray {α : Type u_1} (as : List α) :

                                                                                                                                                                                                                                                                                                                                                                                                          Convert a List α into an Array α. This is O(n) in the length of the list.

                                                                                                                                                                                                                                                                                                                                                                                                          Equations
                                                                                                                                                                                                                                                                                                                                                                                                          Instances For
                                                                                                                                                                                                                                                                                                                                                                                                            class Bind (m : Type u → Type v) :
                                                                                                                                                                                                                                                                                                                                                                                                            Type (max (u + 1) v)

                                                                                                                                                                                                                                                                                                                                                                                                            The typeclass which supplies the >>= "bind" function. See Monad.

                                                                                                                                                                                                                                                                                                                                                                                                            • bind : {α β : Type u} → m α(αm β)m β

                                                                                                                                                                                                                                                                                                                                                                                                              If x : m α and f : α → m β, then x >>= f : m β represents the result of executing x to get a value of type α and then passing it to f.

                                                                                                                                                                                                                                                                                                                                                                                                            Instances
                                                                                                                                                                                                                                                                                                                                                                                                              class Pure (f : Type u → Type v) :
                                                                                                                                                                                                                                                                                                                                                                                                              Type (max (u + 1) v)

                                                                                                                                                                                                                                                                                                                                                                                                              The typeclass which supplies the pure function. See Monad.

                                                                                                                                                                                                                                                                                                                                                                                                              • pure : {α : Type u} → αf α

                                                                                                                                                                                                                                                                                                                                                                                                                If a : α, then pure a : f α represents a monadic action that does nothing and returns a.

                                                                                                                                                                                                                                                                                                                                                                                                              Instances
                                                                                                                                                                                                                                                                                                                                                                                                                class Functor (f : Type u → Type v) :
                                                                                                                                                                                                                                                                                                                                                                                                                Type (max (u + 1) v)

                                                                                                                                                                                                                                                                                                                                                                                                                In functional programming, a "functor" is a function on types F : Type u → Type v equipped with an operator called map or <$> such that if f : α → β then map f : F α → F β, so f <$> x : F β if x : F α. This corresponds to the category-theory notion of functor in the special case where the category is the category of types and functions between them, except that this class supplies only the operations and not the laws (see LawfulFunctor).

                                                                                                                                                                                                                                                                                                                                                                                                                • map : {α β : Type u} → (αβ)f αf β

                                                                                                                                                                                                                                                                                                                                                                                                                  If f : α → β and x : F α then f <$> x : F β.

                                                                                                                                                                                                                                                                                                                                                                                                                • mapConst : {α β : Type u} → αf βf α

                                                                                                                                                                                                                                                                                                                                                                                                                  The special case const a <$> x, which can sometimes be implemented more efficiently.

                                                                                                                                                                                                                                                                                                                                                                                                                Instances
                                                                                                                                                                                                                                                                                                                                                                                                                  class Seq (f : Type u → Type v) :
                                                                                                                                                                                                                                                                                                                                                                                                                  Type (max (u + 1) v)

                                                                                                                                                                                                                                                                                                                                                                                                                  The typeclass which supplies the <*> "seq" function. See Applicative.

                                                                                                                                                                                                                                                                                                                                                                                                                  • seq : {α β : Type u} → f (αβ)(Unitf α)f β

                                                                                                                                                                                                                                                                                                                                                                                                                    If mf : F (α → β) and mx : F α, then mf <*> mx : F β. In a monad this is the same as do let f ← mf; x ← mx; pure (f x): it evaluates first the function, then the argument, and applies one to the other.

                                                                                                                                                                                                                                                                                                                                                                                                                    To avoid surprising evaluation semantics, mx is taken "lazily", using a Unit → f α function.

                                                                                                                                                                                                                                                                                                                                                                                                                  Instances
                                                                                                                                                                                                                                                                                                                                                                                                                    class SeqLeft (f : Type u → Type v) :
                                                                                                                                                                                                                                                                                                                                                                                                                    Type (max (u + 1) v)

                                                                                                                                                                                                                                                                                                                                                                                                                    The typeclass which supplies the <* "seqLeft" function. See Applicative.

                                                                                                                                                                                                                                                                                                                                                                                                                    • seqLeft : {α β : Type u} → f α(Unitf β)f α

                                                                                                                                                                                                                                                                                                                                                                                                                      If x : F α and y : F β, then x <* y evaluates x, then y, and returns the result of x.

                                                                                                                                                                                                                                                                                                                                                                                                                      To avoid surprising evaluation semantics, y is taken "lazily", using a Unit → f β function.

                                                                                                                                                                                                                                                                                                                                                                                                                    Instances
                                                                                                                                                                                                                                                                                                                                                                                                                      class SeqRight (f : Type u → Type v) :
                                                                                                                                                                                                                                                                                                                                                                                                                      Type (max (u + 1) v)

                                                                                                                                                                                                                                                                                                                                                                                                                      The typeclass which supplies the *> "seqRight" function. See Applicative.

                                                                                                                                                                                                                                                                                                                                                                                                                      • seqRight : {α β : Type u} → f α(Unitf β)f β

                                                                                                                                                                                                                                                                                                                                                                                                                        If x : F α and y : F β, then x *> y evaluates x, then y, and returns the result of y.

                                                                                                                                                                                                                                                                                                                                                                                                                        To avoid surprising evaluation semantics, y is taken "lazily", using a Unit → f β function.

                                                                                                                                                                                                                                                                                                                                                                                                                      Instances
                                                                                                                                                                                                                                                                                                                                                                                                                        class Applicative (f : Type u → Type v) extends Functor , Pure , Seq , SeqLeft , SeqRight :
                                                                                                                                                                                                                                                                                                                                                                                                                        Type (max (u + 1) v)

                                                                                                                                                                                                                                                                                                                                                                                                                        An applicative functor is an intermediate structure between Functor and Monad. It mainly consists of two operations:

                                                                                                                                                                                                                                                                                                                                                                                                                        • pure : α → F α
                                                                                                                                                                                                                                                                                                                                                                                                                        • seq : F (α → β) → F α → F β (written as <*>)

                                                                                                                                                                                                                                                                                                                                                                                                                        The seq operator gives a notion of evaluation order to the effects, where the first argument is executed before the second, but unlike a monad the results of earlier computations cannot be used to define later actions.

                                                                                                                                                                                                                                                                                                                                                                                                                          Instances
                                                                                                                                                                                                                                                                                                                                                                                                                            class Monad (m : Type u → Type v) extends Applicative , Bind :
                                                                                                                                                                                                                                                                                                                                                                                                                            Type (max (u + 1) v)

                                                                                                                                                                                                                                                                                                                                                                                                                            A monad is a structure which abstracts the concept of sequential control flow. It mainly consists of two operations:

                                                                                                                                                                                                                                                                                                                                                                                                                            • pure : α → F α
                                                                                                                                                                                                                                                                                                                                                                                                                            • bind : F α → (α → F β) → F β (written as >>=)

                                                                                                                                                                                                                                                                                                                                                                                                                            Like many functional programming languages, Lean makes extensive use of monads for structuring programs. In particular, the do notation is a very powerful syntax over monad operations, and it depends on a Monad instance.

                                                                                                                                                                                                                                                                                                                                                                                                                            See the do notation chapter of the manual for details.

                                                                                                                                                                                                                                                                                                                                                                                                                              Instances
                                                                                                                                                                                                                                                                                                                                                                                                                                instance instInhabitedForAll_2 {α : Type u} {m : Type u → Type v} [Monad m] :
                                                                                                                                                                                                                                                                                                                                                                                                                                Inhabited (αm α)
                                                                                                                                                                                                                                                                                                                                                                                                                                Equations
                                                                                                                                                                                                                                                                                                                                                                                                                                • instInhabitedForAll_2 = { default := pure }
                                                                                                                                                                                                                                                                                                                                                                                                                                instance instInhabited {α : Type u} {m : Type u → Type v} [Monad m] [Inhabited α] :
                                                                                                                                                                                                                                                                                                                                                                                                                                Inhabited (m α)
                                                                                                                                                                                                                                                                                                                                                                                                                                Equations
                                                                                                                                                                                                                                                                                                                                                                                                                                • instInhabited = { default := pure default }
                                                                                                                                                                                                                                                                                                                                                                                                                                instance instForAllNonemptyNonempty {m : Type u_1 → Type u_2} {α : Type u_1} [Monad m] [Nonempty α] :
                                                                                                                                                                                                                                                                                                                                                                                                                                Nonempty (m α)
                                                                                                                                                                                                                                                                                                                                                                                                                                Equations
                                                                                                                                                                                                                                                                                                                                                                                                                                • =
                                                                                                                                                                                                                                                                                                                                                                                                                                def Array.sequenceMap {α : Type u} {β : Type v} {m : Type v → Type w} [Monad m] (as : Array α) (f : αm β) :
                                                                                                                                                                                                                                                                                                                                                                                                                                m (Array β)

                                                                                                                                                                                                                                                                                                                                                                                                                                A fusion of Haskell's sequence and map. Used in syntax quotations.

                                                                                                                                                                                                                                                                                                                                                                                                                                Equations
                                                                                                                                                                                                                                                                                                                                                                                                                                Instances For
                                                                                                                                                                                                                                                                                                                                                                                                                                  def Array.sequenceMap.loop {α : Type u} {β : Type v} {m : Type v → Type w} [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.
                                                                                                                                                                                                                                                                                                                                                                                                                                  Instances For
                                                                                                                                                                                                                                                                                                                                                                                                                                    class MonadLift (m : semiOutParam (Type u → Type v)) (n : Type u → Type w) :
                                                                                                                                                                                                                                                                                                                                                                                                                                    Type (max (max (u + 1) v) w)

                                                                                                                                                                                                                                                                                                                                                                                                                                    A function for lifting a computation from an inner Monad to an outer Monad. Like Haskell's MonadTrans, but n does not have to be a monad transformer. Alternatively, an implementation of MonadLayer without layerInvmap (so far).

                                                                                                                                                                                                                                                                                                                                                                                                                                    • monadLift : {α : Type u} → m αn α

                                                                                                                                                                                                                                                                                                                                                                                                                                      Lifts a value from monad m into monad n.

                                                                                                                                                                                                                                                                                                                                                                                                                                    Instances
                                                                                                                                                                                                                                                                                                                                                                                                                                      class MonadLiftT (m : Type u → Type v) (n : Type u → Type w) :
                                                                                                                                                                                                                                                                                                                                                                                                                                      Type (max (max (u + 1) v) w)

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

                                                                                                                                                                                                                                                                                                                                                                                                                                      • monadLift : {α : Type u} → m αn α

                                                                                                                                                                                                                                                                                                                                                                                                                                        Lifts a value from monad m into monad n.

                                                                                                                                                                                                                                                                                                                                                                                                                                      Instances
                                                                                                                                                                                                                                                                                                                                                                                                                                        @[inline, reducible]
                                                                                                                                                                                                                                                                                                                                                                                                                                        abbrev liftM {m : Type u_1 → Type u_2} {n : Type u_1 → Type u_3} [self : MonadLiftT m n] {α : Type u_1} :
                                                                                                                                                                                                                                                                                                                                                                                                                                        m αn α

                                                                                                                                                                                                                                                                                                                                                                                                                                        Lifts a value from monad m into monad n.

                                                                                                                                                                                                                                                                                                                                                                                                                                        Equations
                                                                                                                                                                                                                                                                                                                                                                                                                                        Instances For
                                                                                                                                                                                                                                                                                                                                                                                                                                          @[always_inline]
                                                                                                                                                                                                                                                                                                                                                                                                                                          instance instMonadLiftT (m : Type u_1 → Type u_2) (n : Type u_1 → Type u_3) (o : Type u_1 → Type u_4) [MonadLift n o] [MonadLiftT m n] :
                                                                                                                                                                                                                                                                                                                                                                                                                                          Equations
                                                                                                                                                                                                                                                                                                                                                                                                                                          instance instMonadLiftT_1 (m : Type u_1 → Type u_2) :
                                                                                                                                                                                                                                                                                                                                                                                                                                          Equations
                                                                                                                                                                                                                                                                                                                                                                                                                                          class MonadFunctor (m : semiOutParam (Type u → Type v)) (n : Type u → Type w) :
                                                                                                                                                                                                                                                                                                                                                                                                                                          Type (max (max (u + 1) v) w)

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

                                                                                                                                                                                                                                                                                                                                                                                                                                          • monadMap : {α : Type u} → ({β : Type u} → m βm β)n αn α

                                                                                                                                                                                                                                                                                                                                                                                                                                            Lifts a monad morphism f : {β : Type u} → m β → m β to monadMap f : {α : Type u} → n α → n α.

                                                                                                                                                                                                                                                                                                                                                                                                                                          Instances
                                                                                                                                                                                                                                                                                                                                                                                                                                            class MonadFunctorT (m : Type u → Type v) (n : Type u → Type w) :
                                                                                                                                                                                                                                                                                                                                                                                                                                            Type (max (max (u + 1) v) w)

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

                                                                                                                                                                                                                                                                                                                                                                                                                                            • monadMap : {α : Type u} → ({β : Type u} → m βm β)n αn α

                                                                                                                                                                                                                                                                                                                                                                                                                                              Lifts a monad morphism f : {β : Type u} → m β → m β to monadMap f : {α : Type u} → n α → n α.

                                                                                                                                                                                                                                                                                                                                                                                                                                            Instances
                                                                                                                                                                                                                                                                                                                                                                                                                                              @[always_inline]
                                                                                                                                                                                                                                                                                                                                                                                                                                              instance instMonadFunctorT (m : Type u_1 → Type u_2) (n : Type u_1 → Type u_3) (o : Type u_1 → Type u_4) [MonadFunctor n o] [MonadFunctorT m n] :
                                                                                                                                                                                                                                                                                                                                                                                                                                              Equations
                                                                                                                                                                                                                                                                                                                                                                                                                                              instance monadFunctorRefl (m : Type u_1 → Type u_2) :
                                                                                                                                                                                                                                                                                                                                                                                                                                              Equations
                                                                                                                                                                                                                                                                                                                                                                                                                                              @[unbox]
                                                                                                                                                                                                                                                                                                                                                                                                                                              inductive Except (ε : Type u) (α : Type v) :
                                                                                                                                                                                                                                                                                                                                                                                                                                              Type (max u v)

                                                                                                                                                                                                                                                                                                                                                                                                                                              Except ε α is a type which represents either an error of type ε, or an "ok" value of type α. The error type is listed first because Except ε : Type → Type is a Monad: the pure operation is ok and the bind operation returns the first encountered error.

                                                                                                                                                                                                                                                                                                                                                                                                                                              • error: {ε : Type u} → {α : Type v} → εExcept ε α

                                                                                                                                                                                                                                                                                                                                                                                                                                                A failure value of type ε

                                                                                                                                                                                                                                                                                                                                                                                                                                              • ok: {ε : Type u} → {α : Type v} → αExcept ε α

                                                                                                                                                                                                                                                                                                                                                                                                                                                A success value of type α

                                                                                                                                                                                                                                                                                                                                                                                                                                              Instances For
                                                                                                                                                                                                                                                                                                                                                                                                                                                instance instInhabitedExcept {ε : Type u} {α : Type v} [Inhabited ε] :
                                                                                                                                                                                                                                                                                                                                                                                                                                                Equations
                                                                                                                                                                                                                                                                                                                                                                                                                                                class MonadExceptOf (ε : semiOutParam (Type u)) (m : Type v → Type w) :
                                                                                                                                                                                                                                                                                                                                                                                                                                                Type (max (max u (v + 1)) w)

                                                                                                                                                                                                                                                                                                                                                                                                                                                An implementation of Haskell's MonadError class. A MonadError ε m is a monad m with two operations:

                                                                                                                                                                                                                                                                                                                                                                                                                                                • throw : ε → m α "throws an error" of type ε to the nearest enclosing catch block
                                                                                                                                                                                                                                                                                                                                                                                                                                                • tryCatch (body : m α) (handler : ε → m α) : m α will catch any errors in body and pass the resulting error to handler. Errors in handler will not be caught.

                                                                                                                                                                                                                                                                                                                                                                                                                                                The try ... catch e => ... syntax inside do blocks is sugar for the tryCatch operation.

                                                                                                                                                                                                                                                                                                                                                                                                                                                • throw : {α : Type v} → εm α

                                                                                                                                                                                                                                                                                                                                                                                                                                                  throw : ε → m α "throws an error" of type ε to the nearest enclosing catch block.

                                                                                                                                                                                                                                                                                                                                                                                                                                                • tryCatch : {α : Type v} → m α(εm α)m α

                                                                                                                                                                                                                                                                                                                                                                                                                                                  tryCatch (body : m α) (handler : ε → m α) : m α will catch any errors in body and pass the resulting error to handler. Errors in handler will not be caught.

                                                                                                                                                                                                                                                                                                                                                                                                                                                Instances
                                                                                                                                                                                                                                                                                                                                                                                                                                                  @[inline, reducible]
                                                                                                                                                                                                                                                                                                                                                                                                                                                  abbrev throwThe (ε : Type u) {m : Type v → Type w} [MonadExceptOf ε m] {α : Type v} (e : ε) :
                                                                                                                                                                                                                                                                                                                                                                                                                                                  m α

                                                                                                                                                                                                                                                                                                                                                                                                                                                  This is the same as throw, but allows specifying the particular error type in case the monad supports throwing more than one type of error.

                                                                                                                                                                                                                                                                                                                                                                                                                                                  Equations
                                                                                                                                                                                                                                                                                                                                                                                                                                                  Instances For
                                                                                                                                                                                                                                                                                                                                                                                                                                                    @[inline, reducible]
                                                                                                                                                                                                                                                                                                                                                                                                                                                    abbrev tryCatchThe (ε : Type u) {m : Type v → Type w} [MonadExceptOf ε m] {α : Type v} (x : m α) (handle : εm α) :
                                                                                                                                                                                                                                                                                                                                                                                                                                                    m α

                                                                                                                                                                                                                                                                                                                                                                                                                                                    This is the same as tryCatch, but allows specifying the particular error type in case the monad supports throwing more than one type of error.

                                                                                                                                                                                                                                                                                                                                                                                                                                                    Equations
                                                                                                                                                                                                                                                                                                                                                                                                                                                    Instances For
                                                                                                                                                                                                                                                                                                                                                                                                                                                      class MonadExcept (ε : outParam (Type u)) (m : Type v → Type w) :
                                                                                                                                                                                                                                                                                                                                                                                                                                                      Type (max (max u (v + 1)) w)

                                                                                                                                                                                                                                                                                                                                                                                                                                                      Similar to MonadExceptOf, but ε is an outParam for convenience.

                                                                                                                                                                                                                                                                                                                                                                                                                                                      • throw : {α : Type v} → εm α

                                                                                                                                                                                                                                                                                                                                                                                                                                                        throw : ε → m α "throws an error" of type ε to the nearest enclosing catch block.

                                                                                                                                                                                                                                                                                                                                                                                                                                                      • tryCatch : {α : Type v} → m α(εm α)m α

                                                                                                                                                                                                                                                                                                                                                                                                                                                        tryCatch (body : m α) (handler : ε → m α) : m α will catch any errors in body and pass the resulting error to handler. Errors in handler will not be caught.

                                                                                                                                                                                                                                                                                                                                                                                                                                                      Instances
                                                                                                                                                                                                                                                                                                                                                                                                                                                        def MonadExcept.ofExcept {m : Type u_1 → Type u_2} {ε : Type u_3} {α : Type u_1} [Monad m] [MonadExcept ε m] :
                                                                                                                                                                                                                                                                                                                                                                                                                                                        Except ε αm α

                                                                                                                                                                                                                                                                                                                                                                                                                                                        "Unwraps" an Except ε α to get the α, or throws the exception otherwise.

                                                                                                                                                                                                                                                                                                                                                                                                                                                        Equations
                                                                                                                                                                                                                                                                                                                                                                                                                                                        Instances For
                                                                                                                                                                                                                                                                                                                                                                                                                                                          instance instMonadExcept (ε : outParam (Type u)) (m : Type v → Type w) [MonadExceptOf ε m] :
                                                                                                                                                                                                                                                                                                                                                                                                                                                          Equations
                                                                                                                                                                                                                                                                                                                                                                                                                                                          @[inline]
                                                                                                                                                                                                                                                                                                                                                                                                                                                          def MonadExcept.orElse {ε : Type u} {m : Type v → Type w} [MonadExcept ε m] {α : Type v} (t₁ : m α) (t₂ : Unitm α) :
                                                                                                                                                                                                                                                                                                                                                                                                                                                          m α

                                                                                                                                                                                                                                                                                                                                                                                                                                                          A MonadExcept can implement t₁ <|> t₂ as try t₁ catch _ => t₂.

                                                                                                                                                                                                                                                                                                                                                                                                                                                          Equations
                                                                                                                                                                                                                                                                                                                                                                                                                                                          Instances For
                                                                                                                                                                                                                                                                                                                                                                                                                                                            instance MonadExcept.instOrElse {ε : Type u} {m : Type v → Type w} [MonadExcept ε m] {α : Type v} :
                                                                                                                                                                                                                                                                                                                                                                                                                                                            OrElse (m α)
                                                                                                                                                                                                                                                                                                                                                                                                                                                            Equations
                                                                                                                                                                                                                                                                                                                                                                                                                                                            • MonadExcept.instOrElse = { orElse := MonadExcept.orElse }
                                                                                                                                                                                                                                                                                                                                                                                                                                                            def ReaderT (ρ : Type u) (m : Type u → Type v) (α : Type u) :
                                                                                                                                                                                                                                                                                                                                                                                                                                                            Type (max u v)

                                                                                                                                                                                                                                                                                                                                                                                                                                                            An implementation of Haskell's ReaderT. This is a monad transformer which equips a monad with additional read-only state, of type ρ.

                                                                                                                                                                                                                                                                                                                                                                                                                                                            Equations
                                                                                                                                                                                                                                                                                                                                                                                                                                                            Instances For
                                                                                                                                                                                                                                                                                                                                                                                                                                                              instance instInhabitedReaderT (ρ : Type u) (m : Type u → Type v) (α : Type u) [Inhabited (m α)] :
                                                                                                                                                                                                                                                                                                                                                                                                                                                              Inhabited (ReaderT ρ m α)
                                                                                                                                                                                                                                                                                                                                                                                                                                                              Equations
                                                                                                                                                                                                                                                                                                                                                                                                                                                              @[inline]
                                                                                                                                                                                                                                                                                                                                                                                                                                                              def ReaderT.run {ρ : Type u} {m : Type u → Type v} {α : Type u} (x : ReaderT ρ m α) (r : ρ) :
                                                                                                                                                                                                                                                                                                                                                                                                                                                              m α

                                                                                                                                                                                                                                                                                                                                                                                                                                                              If x : ReaderT ρ m α and r : ρ, then x.run r : ρ runs the monad with the given reader state.

                                                                                                                                                                                                                                                                                                                                                                                                                                                              Equations
                                                                                                                                                                                                                                                                                                                                                                                                                                                              Instances For
                                                                                                                                                                                                                                                                                                                                                                                                                                                                instance ReaderT.instMonadLiftReaderT {ρ : Type u} {m : Type u → Type v} :
                                                                                                                                                                                                                                                                                                                                                                                                                                                                Equations
                                                                                                                                                                                                                                                                                                                                                                                                                                                                • ReaderT.instMonadLiftReaderT = { monadLift := fun {α : Type u} (x : m α) (x_1 : ρ) => x }
                                                                                                                                                                                                                                                                                                                                                                                                                                                                @[always_inline]
                                                                                                                                                                                                                                                                                                                                                                                                                                                                instance ReaderT.instMonadExceptOfReaderT {ρ : Type u} {m : Type u → Type v} (ε : Type u_1) [MonadExceptOf ε m] :
                                                                                                                                                                                                                                                                                                                                                                                                                                                                Equations
                                                                                                                                                                                                                                                                                                                                                                                                                                                                • One or more equations did not get rendered due to their size.
                                                                                                                                                                                                                                                                                                                                                                                                                                                                @[inline]
                                                                                                                                                                                                                                                                                                                                                                                                                                                                def ReaderT.read {ρ : Type u} {m : Type u → Type v} [Monad m] :
                                                                                                                                                                                                                                                                                                                                                                                                                                                                ReaderT ρ m ρ

                                                                                                                                                                                                                                                                                                                                                                                                                                                                (← read) : ρ gets the read-only state of a ReaderT ρ.

                                                                                                                                                                                                                                                                                                                                                                                                                                                                Equations
                                                                                                                                                                                                                                                                                                                                                                                                                                                                • ReaderT.read = pure
                                                                                                                                                                                                                                                                                                                                                                                                                                                                Instances For
                                                                                                                                                                                                                                                                                                                                                                                                                                                                  @[inline]
                                                                                                                                                                                                                                                                                                                                                                                                                                                                  def ReaderT.pure {ρ : Type u} {m : Type u → Type v} [Monad m] {α : Type u} (a : α) :
                                                                                                                                                                                                                                                                                                                                                                                                                                                                  ReaderT ρ m α

                                                                                                                                                                                                                                                                                                                                                                                                                                                                  The pure operation of the ReaderT monad.

                                                                                                                                                                                                                                                                                                                                                                                                                                                                  Equations
                                                                                                                                                                                                                                                                                                                                                                                                                                                                  Instances For
                                                                                                                                                                                                                                                                                                                                                                                                                                                                    @[inline]
                                                                                                                                                                                                                                                                                                                                                                                                                                                                    def ReaderT.bind {ρ : Type u} {m : Type u → Type v} [Monad m] {α : Type u} {β : Type u} (x : ReaderT ρ m α) (f : αReaderT ρ m β) :
                                                                                                                                                                                                                                                                                                                                                                                                                                                                    ReaderT ρ m β

                                                                                                                                                                                                                                                                                                                                                                                                                                                                    The bind operation of the ReaderT monad.

                                                                                                                                                                                                                                                                                                                                                                                                                                                                    Equations
                                                                                                                                                                                                                                                                                                                                                                                                                                                                    Instances For
                                                                                                                                                                                                                                                                                                                                                                                                                                                                      @[always_inline]
                                                                                                                                                                                                                                                                                                                                                                                                                                                                      instance ReaderT.instFunctorReaderT {ρ : Type u} {m : Type u → Type v} [Monad m] :
                                                                                                                                                                                                                                                                                                                                                                                                                                                                      Equations
                                                                                                                                                                                                                                                                                                                                                                                                                                                                      • One or more equations did not get rendered due to their size.
                                                                                                                                                                                                                                                                                                                                                                                                                                                                      @[always_inline]
                                                                                                                                                                                                                                                                                                                                                                                                                                                                      instance ReaderT.instApplicativeReaderT {ρ : Type u} {m : Type u → Type v} [Monad m] :
                                                                                                                                                                                                                                                                                                                                                                                                                                                                      Equations
                                                                                                                                                                                                                                                                                                                                                                                                                                                                      • ReaderT.instApplicativeReaderT = Applicative.mk
                                                                                                                                                                                                                                                                                                                                                                                                                                                                      instance ReaderT.instMonadReaderT {ρ : Type u} {m : Type u → Type v} [Monad m] :
                                                                                                                                                                                                                                                                                                                                                                                                                                                                      Monad (ReaderT ρ m)
                                                                                                                                                                                                                                                                                                                                                                                                                                                                      Equations
                                                                                                                                                                                                                                                                                                                                                                                                                                                                      • ReaderT.instMonadReaderT = Monad.mk
                                                                                                                                                                                                                                                                                                                                                                                                                                                                      instance ReaderT.instMonadFunctorReaderT (ρ : Type u_1) (m : Type u_1 → Type u_2) :
                                                                                                                                                                                                                                                                                                                                                                                                                                                                      Equations
                                                                                                                                                                                                                                                                                                                                                                                                                                                                      @[inline]
                                                                                                                                                                                                                                                                                                                                                                                                                                                                      def ReaderT.adapt {ρ : Type u} {m : Type u → Type v} {ρ' : Type u} {α : Type u} (f : ρ'ρ) :
                                                                                                                                                                                                                                                                                                                                                                                                                                                                      ReaderT ρ m αReaderT ρ' m α

                                                                                                                                                                                                                                                                                                                                                                                                                                                                      adapt (f : ρ' → ρ) precomposes function f on the reader state of a ReaderT ρ, yielding a ReaderT ρ'.

                                                                                                                                                                                                                                                                                                                                                                                                                                                                      Equations
                                                                                                                                                                                                                                                                                                                                                                                                                                                                      Instances For
                                                                                                                                                                                                                                                                                                                                                                                                                                                                        class MonadReaderOf (ρ : semiOutParam (Type u)) (m : Type u → Type v) :

                                                                                                                                                                                                                                                                                                                                                                                                                                                                        An implementation of Haskell's MonadReader (sans functional dependency; see also MonadReader in this module). It does not contain local because this function cannot be lifted using monadLift. local is instead provided by the MonadWithReader class as withReader.

                                                                                                                                                                                                                                                                                                                                                                                                                                                                        Note: This class can be seen as a simplification of the more "principled" definition

                                                                                                                                                                                                                                                                                                                                                                                                                                                                        class MonadReaderOf (ρ : Type u) (n : Type u → Type u) where
                                                                                                                                                                                                                                                                                                                                                                                                                                                                          lift {α : Type u} : ({m : Type u → Type u} → [Monad m] → ReaderT ρ m α) → n α
                                                                                                                                                                                                                                                                                                                                                                                                                                                                        
                                                                                                                                                                                                                                                                                                                                                                                                                                                                        • read : m ρ

                                                                                                                                                                                                                                                                                                                                                                                                                                                                          (← read) : ρ reads the state out of monad m.

                                                                                                                                                                                                                                                                                                                                                                                                                                                                        Instances
                                                                                                                                                                                                                                                                                                                                                                                                                                                                          @[inline]
                                                                                                                                                                                                                                                                                                                                                                                                                                                                          def readThe (ρ : Type u) {m : Type u → Type v} [MonadReaderOf ρ m] :
                                                                                                                                                                                                                                                                                                                                                                                                                                                                          m ρ

                                                                                                                                                                                                                                                                                                                                                                                                                                                                          Like read, but with ρ explicit. This is useful if a monad supports MonadReaderOf for multiple different types ρ.

                                                                                                                                                                                                                                                                                                                                                                                                                                                                          Equations
                                                                                                                                                                                                                                                                                                                                                                                                                                                                          Instances For
                                                                                                                                                                                                                                                                                                                                                                                                                                                                            class MonadReader (ρ : outParam (Type u)) (m : Type u → Type v) :

                                                                                                                                                                                                                                                                                                                                                                                                                                                                            Similar to MonadReaderOf, but ρ is an outParam for convenience.

                                                                                                                                                                                                                                                                                                                                                                                                                                                                            • read : m ρ

                                                                                                                                                                                                                                                                                                                                                                                                                                                                              (← read) : ρ reads the state out of monad m.

                                                                                                                                                                                                                                                                                                                                                                                                                                                                            Instances
                                                                                                                                                                                                                                                                                                                                                                                                                                                                              instance instMonadReader (ρ : Type u) (m : Type u → Type v) [MonadReaderOf ρ m] :
                                                                                                                                                                                                                                                                                                                                                                                                                                                                              Equations
                                                                                                                                                                                                                                                                                                                                                                                                                                                                              instance instMonadReaderOf {ρ : Type u} {m : Type u → Type v} {n : Type u → Type w} [MonadLift m n] [MonadReaderOf ρ m] :
                                                                                                                                                                                                                                                                                                                                                                                                                                                                              Equations
                                                                                                                                                                                                                                                                                                                                                                                                                                                                              • instMonadReaderOf = { read := liftM read }
                                                                                                                                                                                                                                                                                                                                                                                                                                                                              instance instMonadReaderOfReaderT {ρ : Type u} {m : Type u → Type v} [Monad m] :
                                                                                                                                                                                                                                                                                                                                                                                                                                                                              Equations
                                                                                                                                                                                                                                                                                                                                                                                                                                                                              • instMonadReaderOfReaderT = { read := ReaderT.read }
                                                                                                                                                                                                                                                                                                                                                                                                                                                                              class MonadWithReaderOf (ρ : semiOutParam (Type u)) (m : Type u → Type v) :
                                                                                                                                                                                                                                                                                                                                                                                                                                                                              Type (max (u + 1) v)

                                                                                                                                                                                                                                                                                                                                                                                                                                                                              MonadWithReaderOf ρ adds the operation withReader : (ρ → ρ) → m α → m α. This runs the inner x : m α inside a modified context after applying the function f : ρ → ρ. In addition to ReaderT itself, this operation lifts over most monad transformers, so it allows us to apply withReader to monads deeper in the stack.

                                                                                                                                                                                                                                                                                                                                                                                                                                                                              • withReader : {α : Type u} → (ρρ)m αm α

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                withReader (f : ρ → ρ) (x : m α) : m α runs the inner x : m α inside a modified context after applying the function f : ρ → ρ.

                                                                                                                                                                                                                                                                                                                                                                                                                                                                              Instances
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                @[inline]
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                def withTheReader (ρ : Type u) {m : Type u → Type v} [MonadWithReaderOf ρ m] {α : Type u} (f : ρρ) (x : m α) :
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                m α

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                Like withReader, but with ρ explicit. This is useful if a monad supports MonadWithReaderOf for multiple different types ρ.

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                Equations
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                Instances For
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                  class MonadWithReader (ρ : outParam (Type u)) (m : Type u → Type v) :
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                  Type (max (u + 1) v)

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                  Similar to MonadWithReaderOf, but ρ is an outParam for convenience.

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                  • withReader : {α : Type u} → (ρρ)m αm α

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                    withReader (f : ρ → ρ) (x : m α) : m α runs the inner x : m α inside a modified context after applying the function f : ρ → ρ.

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                  Instances
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                    instance instMonadWithReader (ρ : Type u) (m : Type u → Type v) [MonadWithReaderOf ρ m] :
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                    Equations
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                    instance instMonadWithReaderOf {ρ : Type u} {m : Type u → Type v} {n : Type u → Type v} [MonadFunctor m n] [MonadWithReaderOf ρ m] :
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                    Equations
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                    instance instMonadWithReaderOfReaderT {ρ : Type u} {m : Type u → Type v} [Monad m] :
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                    Equations
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                    • instMonadWithReaderOfReaderT = { withReader := fun {α : Type u} (f : ρρ) (x : ReaderT ρ m α) (ctx : ρ) => x (f ctx) }
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                    class MonadStateOf (σ : semiOutParam (Type u)) (m : Type u → Type v) :
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                    Type (max (u + 1) v)

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

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                    • get : m σ

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                      (← get) : σ gets the state out of a monad m.

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                    • set : σm PUnit

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                      set (s : σ) replaces the state with value s.

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                    • modifyGet : {α : Type u} → (σα × σ)m α

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                      modifyGet (f : σ → α × σ) applies f to the current state, replaces the state with the return value, and returns a computed value.

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                      It is equivalent to do let (a, s) := f (← get); put s; pure a, but modifyGet f may be preferable because the former does not use the state linearly (without sufficient inlining).

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                    Instances
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                      @[inline, reducible]
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                      abbrev getThe (σ : Type u) {m : Type u → Type v} [MonadStateOf σ m] :
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                      m σ

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                      Like withReader, but with ρ explicit. This is useful if a monad supports MonadWithReaderOf for multiple different types ρ.

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                      Equations
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                      Instances For
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                        @[inline, reducible]
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                        abbrev modifyThe (σ : Type u) {m : Type u → Type v} [MonadStateOf σ m] (f : σσ) :

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                        Like modify, but with σ explicit. This is useful if a monad supports MonadStateOf for multiple different types σ.

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                        Equations
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                        Instances For
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                          @[inline, reducible]
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                          abbrev modifyGetThe {α : Type u} (σ : Type u) {m : Type u → Type v} [MonadStateOf σ m] (f : σα × σ) :
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                          m α

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                          Like modifyGet, but with σ explicit. This is useful if a monad supports MonadStateOf for multiple different types σ.

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                          Equations
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                          Instances For
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                            class MonadState (σ : outParam (Type u)) (m : Type u → Type v) :
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                            Type (max (u + 1) v)

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                            Similar to MonadStateOf, but σ is an outParam for convenience.

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                            • get : m σ

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                              (← get) : σ gets the state out of a monad m.

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                            • set : σm PUnit

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                              set (s : σ) replaces the state with value s.

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                            • modifyGet : {α : Type u} → (σα × σ)m α

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                              modifyGet (f : σ → α × σ) applies f to the current state, replaces the state with the return value, and returns a computed value.

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                              It is equivalent to do let (a, s) := f (← get); put s; pure a, but modifyGet f may be preferable because the former does not use the state linearly (without sufficient inlining).

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                            Instances
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                              instance instMonadState (σ : Type u) (m : Type u → Type v) [MonadStateOf σ m] :
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                              Equations
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                              @[inline]
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                              def modify {σ : Type u} {m : Type u → Type v} [MonadState σ m] (f : σσ) :

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                              modify (f : σ → σ) applies the function f to the state.

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                              It is equivalent to do set (f (← get)), but modify f may be preferable because the former does not use the state linearly (without sufficient inlining).

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                              Equations
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                              Instances For
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                @[inline]
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                def getModify {σ : Type u} {m : Type u → Type v} [MonadState σ m] [Monad m] (f : σσ) :
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                m σ

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                getModify f gets the state, applies function f, and returns the old value of the state. It is equivalent to get <* modify f but may be more efficient.

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                Equations
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                Instances For
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                  @[always_inline]
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                  instance instMonadStateOf {σ : Type u} {m : Type u → Type v} {n : Type u → Type w} [MonadLift m n] [MonadStateOf σ m] :
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                  Equations
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                  inductive EStateM.Result (ε : Type u) (σ : Type u) (α : Type u) :

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                  Result ε σ α is equivalent to Except ε α × σ, but using a single combined inductive yields a more efficient data representation.

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                  • ok: {ε σ α : Type u} → ασEStateM.Result ε σ α

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                    A success value of type α, and a new state σ.

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                  • error: {ε σ α : Type u} → εσEStateM.Result ε σ α

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                    A failure value of type ε, and a new state σ.

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                  Instances For
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                    instance EStateM.instInhabitedResult {ε : Type u} {σ : Type u} {α : Type u} [Inhabited ε] [Inhabited σ] :
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                    Equations
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                    def EStateM (ε : Type u) (σ : Type u) (α : Type u) :

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                    EStateM ε σ is a combined error and state monad, equivalent to ExceptT ε (StateM σ) but more efficient.

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                    Equations
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                    Instances For
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                      instance EStateM.instInhabitedEStateM {ε : Type u} {σ : Type u} {α : Type u} [Inhabited ε] :
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                      Inhabited (EStateM ε σ α)
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                      Equations
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                      @[inline]
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                      def EStateM.pure {ε : Type u} {σ : Type u} {α : Type u} (a : α) :
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                      EStateM ε σ α

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                      The pure operation of the EStateM monad.

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                      Equations
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                      Instances For
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                        @[inline]
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                        def EStateM.set {ε : Type u} {σ : Type u} (s : σ) :

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                        The set operation of the EStateM monad.

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                        Equations
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                        Instances For
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                          @[inline]
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                          def EStateM.get {ε : Type u} {σ : Type u} :
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                          EStateM ε σ σ

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                          The get operation of the EStateM monad.

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                          Equations
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                          Instances For
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                            @[inline]
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                            def EStateM.modifyGet {ε : Type u} {σ : Type u} {α : Type u} (f : σα × σ) :
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                            EStateM ε σ α

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                            The modifyGet operation of the EStateM monad.

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                            Equations
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                            Instances For
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                              @[inline]
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                              def EStateM.throw {ε : Type u} {σ : Type u} {α : Type u} (e : ε) :
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                              EStateM ε σ α

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                              The throw operation of the EStateM monad.

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                              Equations
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                              Instances For
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                class EStateM.Backtrackable (δ : outParam (Type u)) (σ : Type u) :

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                Auxiliary instance for saving/restoring the "backtrackable" part of the state. Here σ is the state, and δ is some subpart of it, and we have a getter and setter for it (a "lens" in the Haskell terminology).

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                • save : σδ

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                  save s : δ retrieves a copy of the backtracking state out of the state.

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                • restore : σδσ

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                  restore (s : σ) (x : δ) : σ applies the old backtracking state x to the state s to get a backtracked state s'.

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                Instances
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                  @[inline]
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                  def EStateM.tryCatch {ε : Type u} {σ : Type u} {δ : Type u} [EStateM.Backtrackable δ σ] {α : Type u} (x : EStateM ε σ α) (handle : εEStateM ε σ α) :
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                  EStateM ε σ α

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                  Implementation of tryCatch for EStateM where the state is Backtrackable.

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                  Equations
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                  Instances For
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                    @[inline]
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                    def EStateM.orElse {ε : Type u} {σ : Type u} {α : Type u} {δ : Type u} [EStateM.Backtrackable δ σ] (x₁ : EStateM ε σ α) (x₂ : UnitEStateM ε σ α) :
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                    EStateM ε σ α

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                    Implementation of orElse for EStateM where the state is Backtrackable.

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                    Equations
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                    Instances For
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                      @[inline]
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                      def EStateM.adaptExcept {ε : Type u} {σ : Type u} {α : Type u} {ε' : Type u} (f : εε') (x : EStateM ε σ α) :
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                      EStateM ε' σ α

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                      Map the exception type of a EStateM ε σ α by a function f : ε → ε'.

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                      Equations
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                      Instances For
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                        @[inline]
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                        def EStateM.bind {ε : Type u} {σ : Type u} {α : Type u} {β : Type u} (x : EStateM ε σ α) (f : αEStateM ε σ β) :
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                        EStateM ε σ β

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                        The bind operation of the EStateM monad.

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                        Equations
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                        Instances For
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                          @[inline]
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                          def EStateM.map {ε : Type u} {σ : Type u} {α : Type u} {β : Type u} (f : αβ) (x : EStateM ε σ α) :
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                          EStateM ε σ β

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                          The map operation of the EStateM monad.

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                          Equations
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                          Instances For
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                            @[inline]
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                            def EStateM.seqRight {ε : Type u} {σ : Type u} {α : Type u} {β : Type u} (x : EStateM ε σ α) (y : UnitEStateM ε σ β) :
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                            EStateM ε σ β

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                            The seqRight operation of the EStateM monad.

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                            Equations
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                            Instances For
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                              @[always_inline]
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                              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} [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 {α : Type u} => EStateM.modifyGet }
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                              instance EStateM.instMonadExceptOfEStateM {ε : Type u} {σ : Type u} {δ : Type u} [EStateM.Backtrackable δ σ] :
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                              Equations
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                              • EStateM.instMonadExceptOfEStateM = { throw := fun {α : Type u} => EStateM.throw, tryCatch := fun {α : Type u} => EStateM.tryCatch }
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                              @[inline]
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                              def EStateM.run {ε : Type u} {σ : Type u} {α : Type u} (x : EStateM ε σ α) (s : σ) :

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                              Execute an EStateM on initial state s to get a Result.

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                              Equations
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                              Instances For
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                @[inline]
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                def EStateM.run' {ε : Type u} {σ : Type u} {α : Type u} (x : EStateM ε σ α) (s : σ) :

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                Execute an EStateM on initial state s for the returned value α. If the monadic action throws an exception, returns none instead.

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                Equations
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                Instances For
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                  @[inline]
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                  def EStateM.dummySave {σ : Type u} :
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                  σPUnit

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                  The save implementation for Backtrackable PUnit σ.

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                  Equations
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                  Instances For
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                    @[inline]
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                    def EStateM.dummyRestore {σ : Type u} :
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                    σPUnitσ

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                    The restore implementation for Backtrackable PUnit σ.

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                    Equations
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                    Instances For

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                      Dummy default instance. This makes every σ trivially "backtrackable" by doing nothing on backtrack. Because this is the first declared instance of Backtrackable _ σ, it will be picked only if there are no other Backtrackable _ σ instances registered.

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                      Equations
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                      • EStateM.nonBacktrackable = { save := EStateM.dummySave, restore := EStateM.dummyRestore }
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                      class Hashable (α : Sort u) :
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                      Sort (max 1 u)

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                      A class for types that can be hashed into a UInt64.

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                      Instances
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                        @[extern lean_uint64_to_usize]

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                        Converts a UInt64 to a USize by reducing modulo USize.size.

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                        @[extern lean_usize_to_uint64]

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                        Upcast a USize to a UInt64. This is lossless because USize.size is either 2^32 or 2^64. This function is overridden with a native implementation.

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                        Equations
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                        Instances For
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                          @[extern lean_uint64_mix_hash]
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                          opaque mixHash (u₁ : UInt64) (u₂ : UInt64) :

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                          An opaque hash mixing operation, used to implement hashing for tuples.

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                          instance instHashableSubtype {α : Sort u_1} [Hashable α] {p : αProp} :
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                          Equations
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                          • instHashableSubtype = { hash := fun (a : Subtype p) => hash a.val }
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                          @[extern lean_string_hash]
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                          opaque String.hash (s : String) :

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                          An opaque string hash function.

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                          @[implemented_by Lean.Name.hash._override]
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                          noncomputable def Lean.Name.hash :

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                          A hash function for names, which is stored inside the name itself as a computed field.

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                          Equations
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                          Instances For
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                            inductive Lean.Name :

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                            Hierarchical names. We use hierarchical names to name declarations and for creating unique identifiers for free variables and metavariables.

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                            You can create hierarchical names using the following quotation notation.

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                            `Lean.Meta.whnf
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                            

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                            It is short for .str (.str (.str .anonymous "Lean") "Meta") "whnf" You can use double quotes to request Lean to statically check whether the name corresponds to a Lean declaration in scope.

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                            ``Lean.Meta.whnf
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                            

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                            If the name is not in scope, Lean will report an error.

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                            Instances For
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                              @[inline, reducible, export lean_name_mk_string]

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                              .str p s is now the preferred form.

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                              Equations
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                              Instances For
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                @[inline, reducible, export lean_name_mk_numeral]

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                .num p v is now the preferred form.

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                Equations
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                Instances For
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                  @[inline, reducible]

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                  Short for .str .anonymous s.

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                  Equations
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                  Instances For
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                    @[reducible]

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                    Make name s₁

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                    Equations
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                    Instances For
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                      @[reducible]
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                      def Lean.Name.mkStr2 (s₁ : String) (s₂ : String) :

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                      Make name s₁.s₂

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                      Equations
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                      Instances For
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                        @[reducible]
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                        def Lean.Name.mkStr3 (s₁ : String) (s₂ : String) (s₃ : String) :

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                        Make name s₁.s₂.s₃

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                        Equations
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                        Instances For
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                          @[reducible]
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                          def Lean.Name.mkStr4 (s₁ : String) (s₂ : String) (s₃ : String) (s₄ : String) :

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                          Make name s₁.s₂.s₃.s₄

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                          Equations
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                          Instances For
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                            @[reducible]
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                            def Lean.Name.mkStr5 (s₁ : String) (s₂ : String) (s₃ : String) (s₄ : String) (s₅ : String) :

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                            Make name s₁.s₂.s₃.s₄.s₅

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                            Equations
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                            Instances For
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                              @[reducible]
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                              def Lean.Name.mkStr6 (s₁ : String) (s₂ : String) (s₃ : String) (s₄ : String) (s₅ : String) (s₆ : String) :

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                              Make name s₁.s₂.s₃.s₄.s₅.s₆

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                              Equations
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                              Instances For
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                @[reducible]
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                def Lean.Name.mkStr7 (s₁ : String) (s₂ : String) (s₃ : String) (s₄ : String) (s₅ : String) (s₆ : String) (s₇ : String) :

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                Make name s₁.s₂.s₃.s₄.s₅.s₆.s₇

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                Equations
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                • One or more equations did not get rendered due to their size.
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                Instances For
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                  @[reducible]
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                  def Lean.Name.mkStr8 (s₁ : String) (s₂ : String) (s₃ : String) (s₄ : String) (s₅ : String) (s₆ : String) (s₇ : String) (s₈ : String) :

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                  Make name s₁.s₂.s₃.s₄.s₅.s₆.s₇.s₈

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                  Equations
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                  • One or more equations did not get rendered due to their size.
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                  Instances For
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                    @[extern lean_name_eq]

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                    (Boolean) equality comparator for names.

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                    Equations
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                    Instances For

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                      Syntax #

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                      Source information of tokens.

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                      • original: SubstringString.PosSubstringString.PosLean.SourceInfo

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                        Token from original input with whitespace and position information. leading will be inferred after parsing by Syntax.updateLeading. During parsing, it is not at all clear what the preceding token was, especially with backtracking.

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                      • synthetic: String.PosString.PosoptParam Bool falseLean.SourceInfo

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                        Synthesized syntax (e.g. from a quotation) annotated with a span from the original source. In the delaborator, we "misuse" this constructor to store synthetic positions identifying subterms.

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                        The canonical flag on synthetic syntax is enabled for syntax that is not literally part of the original input syntax but should be treated "as if" the user really wrote it for the purpose of hovers and error messages. This is usually used on identifiers, to connect the binding site to the user's original syntax even if the name of the identifier changes during expansion, as well as on tokens where we will attach targeted messages.

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                        The syntax token%$stx in a syntax quotation will annotate the token token with the span from stx and also mark it as canonical.

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                        As a rough guide, a macro expansion should only use a given piece of input syntax in a single canonical token, although this is sometimes violated when the same identifier is used to declare two binders, as in the macro expansion for dependent if:

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                        `(if $h : $cond then $t else $e) ~>
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                        `(dite $cond (fun $h => $t) (fun $h => $t))
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                        

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                        In these cases if the user hovers over h they will see information about both binding sites.

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                      • none: Lean.SourceInfo

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                        Synthesized token without position information.

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                      Instances For

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                        Gets the position information from a SourceInfo, if available. If originalOnly is true, then .synthetic syntax will also return none.

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                        Equations
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                        • One or more equations did not get rendered due to their size.
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                        Instances For
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                          @[inline, reducible]

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                          A SyntaxNodeKind classifies Syntax.node values. It is an abbreviation for Name, and you can use name literals to construct SyntaxNodeKinds, but they need not refer to declarations in the environment. Conventionally, a SyntaxNodeKind will correspond to the Parser or ParserDesc declaration that parses it.

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                          Equations
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                          Instances For

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                            Syntax AST #

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                            Binding information resolved and stored at compile time of a syntax quotation. Note: We do not statically know whether a syntax expects a namespace or term name, so a Syntax.ident may contain both preresolution kinds.

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                            Instances For
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                              inductive Lean.Syntax :

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

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                              • missing: Lean.Syntax

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                A missing syntax corresponds to a portion of the syntax tree that is missing because of a parse error. The indexing operator on Syntax also returns missing for indexing out of bounds.

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                              • node: Lean.SourceInfoLean.SyntaxNodeKindArray Lean.SyntaxLean.Syntax

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                Node in the syntax tree.

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                The info field is used by the delaborator to store the position of the subexpression corresponding to this node. The parser sets the info field to none. The parser sets the info field to none, with position retrieval continuing recursively. Nodes created by quotations use the result from SourceInfo.fromRef so that they are marked as synthetic even when the leading/trailing token is not. The delaborator uses the info field to store the position of the subexpression corresponding to this node.

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                (Remark: the node constructor did not have an info field in previous versions. This caused a bug in the interactive widgets, where the popup for a + b was the same as for a. The delaborator used to associate subexpressions with pretty-printed syntax by setting the (string) position of the first atom/identifier to the (expression) position of the subexpression. For example, both a and a + b have the same first identifier, and so their infos got mixed up.)

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                              • atom: Lean.SourceInfoStringLean.Syntax

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                An atom corresponds to a keyword or piece of literal unquoted syntax. These correspond to quoted strings inside syntax declarations. For example, in (x + y), "(", "+" and ")" are atom and x and y are ident.

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                              • ident: Lean.SourceInfoSubstringLake.NameList Lean.Syntax.PreresolvedLean.Syntax

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                An ident corresponds to an identifier as parsed by the ident or rawIdent parsers.

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                • rawVal is the literal substring from the input file
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                • val is the parsed identifier (with hygiene)
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                • preresolved is the list of possible declarations this could refer to
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                              Instances For

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                Create syntax node with 1 child

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                Equations
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                Instances For

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                  Create syntax node with 2 children

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                  Equations
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                  Instances For

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                    Create syntax node with 3 children

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                    Equations
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                    Instances For

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                      Create syntax node with 4 children

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                      Equations
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                      Instances For

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                        Create syntax node with 5 children

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                        Equations
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                        Instances For
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                          def Lean.Syntax.node6 (info : Lean.SourceInfo) (kind : Lean.SyntaxNodeKind) (a₁ : Lean.Syntax) (a₂ : Lean.Syntax) (a₃ : Lean.Syntax) (a₄ : Lean.Syntax) (a₅ : Lean.Syntax) (a₆ : Lean.Syntax) :

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                          Create syntax node with 6 children

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                          Equations
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                          Instances For
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                            def Lean.Syntax.node7 (info : Lean.SourceInfo) (kind : Lean.SyntaxNodeKind) (a₁ : Lean.Syntax) (a₂ : Lean.Syntax) (a₃ : Lean.Syntax) (a₄ : Lean.Syntax) (a₅ : Lean.Syntax) (a₆ : Lean.Syntax) (a₇ : Lean.Syntax) :

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                            Create syntax node with 7 children

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                            Equations
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                            Instances For
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                              def Lean.Syntax.node8 (info : Lean.SourceInfo) (kind : Lean.SyntaxNodeKind) (a₁ : Lean.Syntax) (a₂ : Lean.Syntax) (a₃ : Lean.Syntax) (a₄ : Lean.Syntax) (a₅ : Lean.Syntax) (a₆ : Lean.Syntax) (a₇ : Lean.Syntax) (a₈ : Lean.Syntax) :

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                              Create syntax node with 8 children

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                              Equations
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                              Instances For

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                SyntaxNodeKinds is a set of SyntaxNodeKind (implemented as a list).

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                Equations
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                Instances For

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                  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.

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                  Instances For
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                    Equations
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                    • Lean.instInhabitedTSyntax = { default := { raw := default } }

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                    Builtin kinds #

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                    @[inline, reducible]

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                    The choice kind is used when a piece of syntax has multiple parses, and the determination of which to use is deferred until typing information is available.

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                    Equations
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                    Instances For
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                      @[inline, reducible]

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                      The null kind is used for raw list parsers like many.

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                      Equations
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                      Instances For
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                        @[inline, reducible]

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                        The group kind is by the group parser, to avoid confusing with the null kind when used inside optional.

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                        Equations
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                        Instances For
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                          @[inline, reducible]

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                          ident is not actually used as a node kind, but it is returned by getKind in the ident case so that things that handle different node kinds can also handle ident.

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                          Equations
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                          Instances For
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                            @[inline, reducible]

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                            str is the node kind of string literals like "foo".

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                            Equations
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                            Instances For
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                              @[inline, reducible]

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                              char is the node kind of character literals like 'A'.

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                              Equations
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                              Instances For
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                @[inline, reducible]

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                num is the node kind of number literals like 42.

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                Equations
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                Instances For
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                  @[inline, reducible]

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                  scientific is the node kind of floating point literals like 1.23e-3.

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                  Equations
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                  Instances For
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                    @[inline, reducible]

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                    name is the node kind of name literals like `foo.

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                    Equations
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                    Instances For
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                      @[inline, reducible]

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                      fieldIdx is the node kind of projection indices like the 2 in x.2.

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                      Equations
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                      Instances For
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                        @[inline, reducible]

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                        hygieneInfo is the node kind of the hygieneInfo parser, which is an "invisible token" which captures the hygiene information at the current point without parsing anything.

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                        They can be used to generate identifiers (with Lean.HygieneInfo.mkIdent) as if they were introduced by the calling context, not the called macro.

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                        Equations
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                        Instances For
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                          @[inline, reducible]

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                          interpolatedStrLitKind is the node kind of interpolated string literal fragments like "value = { and }" in s!"value = {x}".

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                          Equations
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                          Instances For
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                            @[inline, reducible]

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                            interpolatedStrKind is the node kind of an interpolated string literal like "value = {x}" in s!"value = {x}".

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                            Equations
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                            Instances For
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                              @[inline]

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                              Creates an info-less node of the given kind and children.

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                              Equations
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                              Instances For
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                @[inline]

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                Creates an info-less nullKind node with the given children, if any.

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                Equations
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                Instances For

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                  Gets the kind of a Syntax node. For non-node syntax, we use "pseudo kinds": identKind for ident, missing for missing, and the atom's string literal for atoms.

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                  Equations
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                  • One or more equations did not get rendered due to their size.
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                  Instances For

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                    Changes the kind at the root of a Syntax node to k. Does nothing for non-node nodes.

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                    Equations
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                    Instances For

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                      Is this a syntax with node kind k?

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                      Equations
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                      Instances For

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                        Gets the i'th argument of the syntax node. This can also be written stx[i]. Returns missing if i is out of range.

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                        Equations
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                        Instances For

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                          Gets the list of arguments of the syntax node, or #[] if it's not a node.

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                          Equations
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                          Instances For

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                            Gets the number of arguments of the syntax node, or 0 if it's not a node.

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                            Equations
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                            Instances For

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                              Assuming stx was parsed by optional, returns the enclosed syntax if it parsed something and none otherwise.

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                              Equations
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                              • One or more equations did not get rendered due to their size.
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                              Instances For

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                Is this syntax .missing?

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                Equations
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                Instances For

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                  Is this syntax a node with kind k?

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                  Equations
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                  Instances For

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                    stx.isIdent is true iff stx is an identifier.

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                    Equations
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                    Instances For

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                      If this is an ident, return the parsed value, else .anonymous.

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                      Equations
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                      Instances For

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                        Updates the argument list without changing the node kind. Does nothing for non-node nodes.

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                        Equations
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                        Instances For

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                          Updates the i'th argument of the syntax. Does nothing for non-node nodes, or if i is out of bounds of the node list.

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                          Equations
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                          Instances For

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                            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
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                            Instances For

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                              Get the starting position of the syntax, if possible. If canonicalOnly is true, non-canonical synthetic nodes are treated as not carrying position information.

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                              Equations
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                              Instances For
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                partial def Lean.Syntax.getTailPos? (stx : Lean.Syntax) (canonicalOnly : optParam Bool false) :

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                Get the ending position of the syntax, if possible. If canonicalOnly is true, non-canonical synthetic nodes are treated as not carrying position information.

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

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

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                • elemsAndSeps : Array Lean.Syntax

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                  The array of elements and separators, ordered like #[el1, sep1, el2, sep2, el3].

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                Instances For

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                  A typed version of SepArray.

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                  • elemsAndSeps : Array Lean.Syntax

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                    The array of elements and separators, ordered like #[el1, sep1, el2, sep2, el3].

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                  Instances For
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                    @[inline, reducible]

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                    An array of syntaxes of kind ks.

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                    Equations
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                    Instances For

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                      Implementation of TSyntaxArray.raw.

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                      Equations
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                      • Lean.TSyntaxArray.rawImpl = unsafeCast
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                      Instances For
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                        @[implemented_by Lean.TSyntaxArray.rawImpl]

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                        Converts a TSyntaxArray to an Array Syntax, without reallocation.

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                        Implementation of TSyntaxArray.mk.

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                        Equations
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                        • Lean.TSyntaxArray.mkImpl = unsafeCast
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                        Instances For
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                          @[implemented_by Lean.TSyntaxArray.mkImpl]

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                          Converts an Array Syntax to a TSyntaxArray, without reallocation.

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                          Constructs a synthetic SourceInfo using a ref : Syntax for the span.

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                          Equations
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                          • One or more equations did not get rendered due to their size.
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                          Instances For

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                            Constructs a synthetic atom with no source info.

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                            Equations
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                            Instances For
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                              def Lean.mkAtomFrom (src : Lean.Syntax) (val : String) (canonical : optParam Bool false) :

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                              Constructs a synthetic atom with source info coming from src.

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                              Equations
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                              Instances For

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                Parser descriptions #

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                A ParserDescr is a grammar for parsers. This is used by the syntax command to produce parsers without having to import Lean.

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                Instances For
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                  @[inline, reducible]

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                  Although TrailingParserDescr is an abbreviation for ParserDescr, Lean will look at the declared type in order to determine whether to add the parser to the leading or trailing parser table. The determination is done automatically by the syntax command.

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                  Equations
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                  Instances For

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                    Runtime support for making quotation terms auto-hygienic, by mangling identifiers introduced by them with a "macro scope" supplied by the context. Details to appear in a paper soon.

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                    @[inline, reducible]

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                    A macro scope identifier is just a Nat that gets bumped every time we enter a new macro scope. Within a macro scope, all occurrences of identifier x parse to the same thing, but x parsed from different macro scopes will produce different identifiers.

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                    Equations
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                    Instances For

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

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                      Equations
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                      Instances For

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                        First macro scope available for our frontend

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                        Equations
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                        Instances For
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                          class Lean.MonadRef (m : TypeType) :

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                          A MonadRef is a monad that has a ref : Syntax in the read-only state. This is used to keep track of the location where we are working; if an exception is thrown, the ref gives the location where the error will be reported, assuming no more specific location is provided.

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                          Instances
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                            instance Lean.instMonadRef (m : TypeType) (n : TypeType) [MonadLift m n] [MonadFunctor m n] [Lean.MonadRef m] :
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                            Equations

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                            Replaces oldRef with ref, unless ref has no position info. This biases us to having a valid span to report an error on.

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                            Equations
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                            Instances For
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                              @[inline]
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                              def Lean.withRef {m : TypeType} [Monad m] [Lean.MonadRef m] {α : Type} (ref : Lean.Syntax) (x : m α) :
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                              m α

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                              Run x : m α with a modified value for the ref. This is not exactly the same as MonadRef.withRef, because it uses replaceRef to avoid putting syntax with bad spans in the state.

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                              Equations
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                              Instances For
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                class Lean.MonadQuotation (m : TypeType) extends Lean.MonadRef :

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                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.

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                • getRef : m Lean.Syntax
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                • withRef : {α : Type} → Lean.Syntaxm αm α
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                • getCurrMacroScope : m Lean.MacroScope

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                  Get the fresh scope of the current macro invocation

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                • getMainModule : m Lake.Name

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                  Get the module name of the current file. This is used to ensure that hygienic names don't clash across multiple files.

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                • withFreshMacroScope : {α : Type} → m αm α

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                  Execute action in a new macro invocation context. This transformer should be used at all places that morally qualify as the beginning of a "macro call", e.g. elabCommand and elabTerm in the case of the elaborator. However, it can also be used internally inside a "macro" if identifiers introduced by e.g. different recursive calls should be independent and not collide. While returning an intermediate syntax tree that will recursively be expanded by the elaborator can be used for the same effect, doing direct recursion inside the macro guarded by this transformer is often easier because one is not restricted to passing a single syntax tree. Modelling this helper as a transformer and not just a monadic action ensures that the current macro scope before the recursive call is restored after it, as expected.

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                Instances
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                  @[inline]

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                  Construct a synthetic SourceInfo from the ref in the monad state.

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                  Equations
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                  Instances For
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                    Equations

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                    We represent a name with macro scopes as

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                    <actual name>._@.(<module_name>.<scopes>)*.<module_name>._hyg.<scopes>
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                    

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                    Example: suppose the module name is Init.Data.List.Basic, and name is foo.bla, and macroscopes [2, 5]

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                    foo.bla._@.Init.Data.List.Basic._hyg.2.5
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                    

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                    We may have to combine scopes from different files/modules. The main modules being processed is always the right-most one. This situation may happen when we execute a macro generated in an imported file in the current file.

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                    foo.bla._@.Init.Data.List.Basic.2.1.Init.Lean.Expr._hyg.4
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                    

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                    The delimiter _hyg is used just to improve the hasMacroScopes performance.

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                    @[export lean_erase_macro_scopes]

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                    Remove the macro scopes from the name.

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                    Equations
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                    Instances For
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                      @[export lean_simp_macro_scopes]

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                      Helper function we use to create binder names that do not need to be unique.

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                      Equations
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                      Instances For

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                        A MacroScopesView represents a parsed hygienic name. extractMacroScopes will decode it from a Name, and .review will re-encode it. The grammar of a hygienic name is:

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                        <name>._@.(<module_name>.<scopes>)*.<mainModule>._hyg.<scopes>
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                        
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                        • name : Lake.Name

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                          The original (unhygienic) name.

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                        • imported : Lake.Name

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                          All the name components (<module_name>.<scopes>)* from the imports concatenated together.

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                        • mainModule : Lake.Name

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                          The main module in which this identifier was parsed.

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                        • The list of macro scopes.

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                        Instances For
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                          Equations

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                          Encode a hygienic name from the parsed pieces.

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                          Equations
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                          • One or more equations did not get rendered due to their size.
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                          Instances For

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                            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.
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                            Instances For

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                              Add a new macro scope onto the name n, in the given mainModule.

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                              Equations
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                              • One or more equations did not get rendered due to their size.
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                              Instances For

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                Append two names that may have macro scopes. The macro scopes in b are always erased. If a has macro scopes, then they are propagated to the result of append a b.

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                Equations
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                • One or more equations did not get rendered due to their size.
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                Instances For
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                  @[inline]

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                  Add a new macro scope onto the name n, using the monad state to supply the main module and current macro scope.

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                  Equations
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                  Instances For

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                    The default maximum recursion depth. This is adjustable using the maxRecDepth option.

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                    Equations
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                    Instances For

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                      The message to display on stack overflow.

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                      Equations
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                      Instances For

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                        Is this syntax a null node?

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                        Equations
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                        Instances For

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                          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
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                          Instances For

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                            Is this syntax a node kind k wrapping an atom _ val?

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                            Equations
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                            • One or more equations did not get rendered due to their size.
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                            Instances For

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                              The read-only context for the MacroM monad.

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                              • An opaque reference to the Methods object. This is done to break a dependency cycle: the Methods involve MacroM which has not been defined yet.

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                              • mainModule : Lake.Name

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                The currently parsing module.

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                              • currMacroScope : Lean.MacroScope

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                The current macro scope.

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                              • currRecDepth : Nat

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                The current recursion depth.

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                              • maxRecDepth : Nat

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                The maximum recursion depth.

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                              • The syntax which supplies the position of error messages.

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                              Instances For

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                An exception in the MacroM monad.

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                • error: Lean.SyntaxStringLean.Macro.Exception

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                  A general error, given a message and a span (expressed as a Syntax).

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                • unsupportedSyntax: Lean.Macro.Exception

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                  An unsupported syntax exception. We keep this separate because it is used for control flow: if one macro does not support a syntax then we try the next one.

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                Instances For

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                  The mutable state for the MacroM monad.

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                  • macroScope : Lean.MacroScope

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                    The global macro scope counter, used for producing fresh scope names.

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                  • traceMsgs : List (Lake.Name × String)

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                    The list of trace messages that have been produced, each with a trace class and a message.

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                  Instances For
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                    Equations
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                    @[inline, reducible]
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                    abbrev Lean.MacroM (α : Type) :

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                    The MacroM monad is the main monad for macro expansion. It has the information needed to handle hygienic name generation, and is the monad that macro definitions live in.

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                    Notably, this is a (relatively) pure monad: there is no IO and no access to the Environment. That means that things like declaration lookup are impossible here, as well as IO.Ref or other side-effecting operations. For more capabilities, macros can instead be written as elab using adaptExpander.

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                    Equations
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                    Instances For
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                      @[inline, reducible]

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                      A macro has type Macro, which is a SyntaxMacroM Syntax: it receives an input syntax and is supposed to "expand" it into another piece of syntax.

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                      Equations
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                      Instances For
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                        Equations
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                        • One or more equations did not get rendered due to their size.

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                        Add a new macro scope to the name n.

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                        Equations
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                        Instances For

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                          Throw an unsupportedSyntax exception.

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                          Equations
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                          Instances For

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                            Throw an error with the given message, using the ref for the location information.

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                            Equations
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                            Instances For

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                              Throw an error with the given message and location information.

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                              Equations
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                              Instances For
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                @[inline]

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                Increments the macro scope counter so that inside the body of x the macro scope is fresh.

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                Equations
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                • One or more equations did not get rendered due to their size.
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                Instances For
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                  @[inline]

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                  Run x with an incremented recursion depth counter.

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                  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.

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                    The opaque methods that are available to MacroM.

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                    Instances For
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                      Equations
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                      • One or more equations did not get rendered due to their size.

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                      Implementation of mkMethods.

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                      Equations
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                      Instances For
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                        @[implemented_by Lean.Macro.mkMethodsImp]

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                        Make an opaque reference to a Methods.

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                        Implementation of getMethods.

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                        Equations
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                        Instances For
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                          @[implemented_by Lean.Macro.getMethodsImp]

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                          Extract the methods list from the MacroM state.

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

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                          Equations
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                          Instances For

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                            Returns true if the environment contains a declaration with name declName

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                            Equations
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                            Instances For

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                              Gets the current namespace given the position in the file.

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                              Equations
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                              Instances For

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                Resolves the given name to an overload list of namespaces.

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                Equations
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                Instances For

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                  Resolves the given name to an overload list of global definitions. The List String in each alternative is the deduced list of projections (which are ambiguous with name components).

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                  Equations
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                  Instances For

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                    Add a new trace message, with the given trace class and message.

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                    Equations
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                    Instances For
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                      @[inline, reducible]

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                      The unexpander monad, essentially SyntaxOption α. The Syntax is the ref, and it has the possibility of failure without an error message.

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                      Equations
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                      Instances For
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                        @[inline, reducible]

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                        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 [app_unexpander] attribute.

                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                        Equations
                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                                        Instances For