Documentation

Lean.Elab.Binders

Auxiliary datatype for elaborating binders.

  • Position information provider for the Info Tree. We currently do not track binder "macro expansion" steps in the info tree. For example, suppose we expand a _ into a fresh identifier. The fresh identifier has synthetic position since it was not written by the user, and we would not get hover information for the _ because we also don't have this macro expansion step stored in the info tree. Thus, we store the original Syntax in ref, and use it when storing the binder information in the info tree.

    Potential better solution: add a binder syntax category, an extensible elabBinder (like we have elabTerm), and perform all macro expansion steps at elabBinder and record them in the infro tree.

Instances For

    Determines the local declaration kind depending on the variable name.

    The __x in let __x := 42; body gets kind .implDetail.

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

        Like elabBinders, but also pass syntax node per binder. elabBinders(Ex) automatically adds binder info nodes for the produced fvars, but storing the syntax nodes might be necessary when later adding the same binders back to the local context so that info nodes can manually be added for the new fvars; see MutualDef for an example.

        Equations
        Instances For

          Elaborate the given binders (i.e., Syntax objects for bracketedBinder), update the local context, set of local instances, reset instance cache (if needed), and then execute k with the updated context. The local context will only be included inside k.

          For example, suppose you have binders [(a : α), (b : β a)], then the elaborator will create two new free variables a and b, push these to the context and pass to k #[a,b].

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

            Same as elabBinder with a single binder.

            Equations
            Instances For

              If binder is a _ or an identifier, return a bracketedBinder using type otherwise throw an exception.

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

                        The dependent arrow. (x : α) → β is equivalent to ∀ x : α, β, but we usually reserve the latter for propositions. Also written as Π x : α, β (the "Pi-type") in the literature.

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

                          Auxiliary function for expanding fun notation binders. Recall that fun parser is defined as

                          def funBinder : Parser := implicitBinder <|> instBinder <|> termParser maxPrec
                          leading_parser unicodeSymbol "λ" "fun" >> many1 funBinder >> "=>" >> termParser
                          

                          to allow notation such as fun (a, b) => a + b, where (a, b) should be treated as a pattern. The result is a pair (explicitBinders, newBody), where explicitBinders is syntax of the form

                          `(` ident `:` term `)`
                          

                          which can be elaborated using elabBinders, and newBody is the updated body syntax. We update the body syntax when expanding the pattern notation. Example: fun (a, b) => a + b expands into fun _a_1 => match _a_1 with | (a, b) => a + b. See local function processAsPattern at expandFunBindersAux.

                          The resulting Bool is true if a pattern was found. We use it "mark" a macro expansion.

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

                                    Expand matchAlts syntax into a full match-expression. Example:

                                    | 0, true => alt_1
                                    | i, _    => alt_2
                                    

                                    expands into (for tactic == false)

                                    fun x_1 x_2 =>
                                    match @x_1, @x_2 with
                                    | 0, true => alt_1
                                    | i, _    => alt_2
                                    

                                    and (for tactic == true)

                                    intro x_1; intro x_2;
                                    match @x_1, @x_2 with
                                    | 0, true => alt_1
                                    | i, _    => alt_2
                                    

                                    If useExplicit = true, we add a @ before fun to disable implicit lambdas. We disable them when processing let and let rec declarations to make sure the behavior is consistent with top-level declarations where we can write

                                    def f : {α : Type} → α → α
                                      | _, a => a
                                    

                                    We use useExplicit = false when we are elaborating the fun | ... => ... | ... notation. See issue #1132. If @fun is used with this notation, the we set useExplicit = true. We also use useExplicit = false when processing instance ... where notation declarations. The motivation is to have compact declarations such as

                                    instance [Alternative m] : MonadLiftT Option m where
                                    monadLift -- We don't want to provide the implicit arguments of `monadLift` here. One should use `monadLift := @fun ...` if they want to provide them.
                                      | some a => pure a
                                      | none => failure
                                    

                                    Remark: we add @ at discriminants to make sure we don't consume implicit arguments, and to make the behavior consistent with fun. Example:

                                    inductive T : Type 1 :=
                                    | mkT : (forall {a : Type}, a -> a) -> T
                                    
                                    def makeT (f : forall {a : Type}, a -> a) : T :=
                                      mkT f
                                    
                                    def makeT' : (forall {a : Type}, a -> a) -> T
                                    | f => mkT f
                                    

                                    The two definitions should be elaborated without errors and be equivalent.

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

                                        Similar to expandMatchAltsIntoMatch, but supports an optional where clause.

                                        Expand matchAltsWhereDecls into let rec + match-expression. Example

                                        | 0, true => ... f 0 ...
                                        | i, _    => ... f i + g i ...
                                        where
                                          f x := g x + 1
                                        
                                          g : NatNat
                                            | 0   => 1
                                            | x+1 => f x
                                        

                                        expands into

                                        fux x_1 x_2 =>
                                          let rec
                                            f x := g x + 1,
                                            g : NatNat
                                              | 0   => 1
                                              | x+1 => f x
                                          match x_1, x_2 with
                                          | 0, true => ... f 0 ...
                                          | i, _    => ... f i + g i ...
                                        
                                        Equations
                                        • One or more equations did not get rendered due to their size.
                                        Instances For
                                          Equations
                                          • One or more equations did not get rendered due to their size.
                                          Instances For
                                            Equations
                                            • One or more equations did not get rendered due to their size.
                                            Instances For
                                              Equations
                                              • One or more equations did not get rendered due to their size.
                                              Instances For
                                                Equations
                                                • One or more equations did not get rendered due to their size.
                                                Instances For
                                                  Equations
                                                  • One or more equations did not get rendered due to their size.
                                                  Instances For
                                                    def Lean.Elab.Term.elabLetDeclAux (id : Lean.Syntax) (binders : Array Lean.Syntax) (typeStx : Lean.Syntax) (valStx : Lean.Syntax) (body : Lean.Syntax) (expectedType? : Option Lean.Expr) (useLetExpr : Bool) (elabBodyFirst : Bool) (usedLetOnly : Bool) :

                                                    If useLetExpr is true, then a kernel let-expression let x : type := val; body is created. Otherwise, we create a term of the form (fun (x : type) => body) val

                                                    The default elaboration order is binders, typeStx, valStx, and body. If elabBodyFirst == true, then we use the order binders, typeStx, body, and valStx.

                                                    Equations
                                                    • One or more equations did not get rendered due to their size.
                                                    Instances For
                                                      Equations
                                                      • One or more equations did not get rendered due to their size.
                                                      Instances For
                                                        Equations
                                                        • One or more equations did not get rendered due to their size.
                                                        Instances For
                                                          def Lean.Elab.Term.elabLetDeclCore (stx : Lean.Syntax) (expectedType? : Option Lean.Expr) (useLetExpr : Bool) (elabBodyFirst : Bool) (usedLetOnly : Bool) :
                                                          Equations
                                                          • One or more equations did not get rendered due to their size.
                                                          Instances For