Documentation

Lean.Parser.Term.Basic

This module contains the bare minimum of term syntax that's required to get documentation syntax to parse, namely structure initializers and their dependencies.

This matters because some term syntax (such as let rec) includes docstrings, but docstrings include metadata blocks that themselves include a structure initializer. Separating these layers prevents an import cycle.

The remaining term syntax is found in Lean.Parser.Term. It may freely make use of documentation comments.

@[inline]
Equations
Instances For
    @[inline]
    Equations
    Instances For

      sepByIndentSemicolon(p) parses a sequence of p optionally followed by ;, similar to manyIndent(p ";"?), except that if two occurrences of p occur on the same line, the ; is mandatory. This is used by tactic parsing, so that

      example := by
        skip
        skip
      

      is legal, but by skip skip is not - it must be written as by skip; skip.

      Equations
      Instances For

        sepBy1IndentSemicolon(p) parses a (nonempty) sequence of p optionally followed by ;, similar to many1Indent(p ";"?), except that if two occurrences of p occur on the same line, the ; is mandatory. This is used by tactic parsing, so that

        example := by
          skip
          skip
        

        is legal, but by skip skip is not - it must be written as by skip; skip.

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

            The syntax { tacs } is an alternative syntax for · tacs. It runs the tactics in sequence, and fails if the goal is not solved.

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

              A sequence of tactics in brackets, or a delimiter-free indented sequence of tactics. Delimiter-free indentation is determined by the first tactic of the sequence.

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

                Same as [tacticSeq] but requires delimiter-free tactic sequence to have strict indentation. The strict indentation requirement only apply to nested bys, as top-level bys do not have a position set.

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

                  A hole (or placeholder term), which stands for an unknown term that is expected to be inferred based on context. For example, in @id _ Nat.zero, the _ must be the type of Nat.zero, which is Nat.

                  The way this works is that holes create fresh metavariables. The elaborator is allowed to assign terms to metavariables while it is checking definitional equalities. This is often known as unification.

                  Normally, all holes must be solved for. However, there are a few contexts where this is not necessary:

                  • In match patterns, holes are catch-all patterns.
                  • In some tactics, such as refine' and apply, unsolved-for placeholders become new goals.

                  Related concept: implicit parameters are automatically filled in with holes during the elaboration process.

                  See also ?m syntax (synthetic holes).

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

                    A synthetic hole (or synthetic placeholder), which stands for an unknown term that should be synthesized using tactics.

                    • ?_ creates a fresh metavariable with an auto-generated name.
                    • ?m either refers to a pre-existing metavariable named m or creates a fresh metavariable with that name.

                    In particular, the synthetic hole syntax creates "synthetic opaque metavariables", the same kind of metavariable used to represent goals in the tactic state.

                    Synthetic holes are similar to holes in that _ also creates metavariables, but synthetic opaque metavariables have some different properties:

                    • In tactics such as refine, only synthetic holes yield new goals.
                    • During elaboration, unification will not solve for synthetic opaque metavariables, they are "opaque". This is to prevent counterintuitive behavior such as disappearing goals.
                    • When synthetic holes appear under binders, they capture local variables using a more complicated mechanism known as delayed assignment.

                    Delayed assigned metavariables #

                    This section gives an overview of some technical details of synthetic holes, which you should feel free to skip. Understanding delayed assignments is mainly useful for those who are working on tactics and other metaprogramming. It is included here until there is a suitable place for it in the reference manual.

                    When a synthetic hole appears under a binding construct, such as for example fun (x : α) (y : β) => ?s, the system creates a delayed assignment. This consists of

                    1. A metavariable ?m of type (x : α) → (y : β) → γ x y whose local context is the local context outside the fun, where γ x y is the type of ?s. Recall that x and y appear in the local context of ?s.
                    2. A delayed assignment record associating ?m to ?s and the variables #[x, y] in the local context of ?s

                    Then, this function elaborates as fun (x : α) (y : β) => ?m x y, where one should understand x and y here as being De Bruijn indexes, since Lean uses the locally nameless encoding of lambda calculus.

                    Once ?s is fully solved for, in the sense that after metavariable instantiation it is a metavariable-free term e, then we can make the assignment ?m := fun (x' : α) (y' : β) => e[x := x', y := y']. (Implementation note: Lean only instantiates full applications ?m x' y' of delayed assigned metavariables, to skip forming this function.) This delayed assignment mechanism is essential to the operation of basic tactics like intro, and a good mental model is that it is a way to "apply" the metavariable ?s by substituting values in for some of its local variables. While it would be easier to immediately assign ?s := ?m x y, delayed assignment preserves ?s as an unsolved-for metavariable with a local context that still contains x and y, which is exactly what tactics like intro need.

                    By default, delayed assigned metavariables pretty print with what they are delayed assigned to. The delayed assigned metavariables themselves can be pretty printed using set_option pp.mvars.delayed true.

                    For more information, see the "Gruesome details" module docstrings in Lean.MetavarContext.

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

                      The term denotes a term that was omitted by the pretty printer. The presence of in pretty printer output is controlled by the pp.deepTerms and pp.proofs options, and these options can be further adjusted using pp.deepTerms.threshold and pp.proofs.threshold.

                      It is only meant to be used for pretty printing. However, in case it is copied and pasted from the Infoview, logs a warning and elaborates like _.

                      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

                                Explicit binder, like (x y : A) or (x y). Default values can be specified using (x : A := v) syntax, and tactics using (x : A := by tac).

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

                                  Implicit binder, like {x y : A} or {x y}. In regular applications, whenever all parameters before it have been specified, then a _ placeholder is automatically inserted for this parameter. Implicit parameters should be able to be determined from the other arguments and the return type by unification.

                                  In @ explicit mode, implicit binders behave like explicit binders.

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

                                    Strict-implicit binder, like ⦃x y : A⦄ or ⦃x y⦄. In contrast to { ... } implicit binders, strict-implicit binders do not automatically insert a _ placeholder until at least one subsequent explicit parameter is specified. Do not use strict-implicit binders unless there is a subsequent explicit parameter. Assuming this rule is followed, for fully applied expressions implicit and strict-implicit binders have the same behavior.

                                    Example: If h : ∀ ⦃x : A⦄, x ∈ s → p x and hs : y ∈ s, then h by itself elaborates to itself without inserting _ for the x : A parameter, and h hs has type p y. In contrast, if h' : ∀ {x : A}, x ∈ s → p x, then h by itself elaborates to have type ?m ∈ s → p ?m with ?m a fresh metavariable.

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

                                      Instance-implicit binder, like [C] or [inst : C]. In regular applications without @ explicit mode, it is automatically inserted and solved for by typeclass inference for the specified class C. In @ explicit mode, if _ is used for an instance-implicit parameter, then it is still solved for by typeclass inference; use (_) to inhibit this and have it be solved for by unification instead, like an implicit argument.

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

                                        A bracketedBinder matches any kind of binder group that uses some kind of brackets:

                                        • An explicit binder like (x y : A)
                                        • An implicit binder like {x y : A}
                                        • A strict implicit binder, ⦃y z : A⦄ or its ASCII alternative {{y z : A}}
                                        • An instance binder [A] or [x : A] (multiple variables are not allowed here)
                                        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
                                                    Equations
                                                    • One or more equations did not get rendered due to their size.
                                                    Instances For
                                                      Equations
                                                      Instances For