Documentation

Lean.Language.Basic

MessageLog with interactive diagnostics.

  • Non-interactive message log.

  • interactiveDiagsRef? : Option (IO.Ref (Option Dynamic))

    Dynamic mutable slot usable by the language server for memorizing interactive diagnostics. If none, interactive diagnostics are not remembered, which should only be used for messages not containing any interactive elements as client-side state will be lost on recreating a diagnostic.

    See also section "Communication" in Lean/Server/README.md.

Instances For

    The empty set of diagnostics.

    Equations
    Instances For

      The base class of all snapshots: all the generic information the language server needs about a snapshot.

      • The messages produced by this step. The union of message logs of all finished snapshots is reported to the user.

      • General elaboration metadata produced by this step.

      • isFatal : Bool

        Whether it should be indicated to the user that a fatal error (which should be part of diagnostics) occurred that prevents processing of the remainder of the file.

      Instances For
        Equations

        A task producing some snapshot type (usually a subclass of Snapshot).

        • Range that is marked as being processed by the server while the task is running. If none, the range of the outer task if some or else the entire file is reported.

        • task : Task α

          Underlying task producing the snapshot.

        Instances For

          Creates a snapshot task from a reporting range and a BaseIO action.

          Equations
          Instances For

            Creates a finished snapshot task.

            Equations
            Instances For

              Explicitly cancels a tasks. Like with basic Taskss, cancellation happens implicitly when the last reference to the task is dropped if it is not an I/O task.

              Equations
              Instances For

                Transforms a task's output without changing the reporting range.

                Equations
                Instances For

                  Chains two snapshot tasks. The range is taken from the first task if not specified; the range of the second task is discarded.

                  Equations
                  • t.bind act range? sync = { range? := range?, task := t.task.bind (fun (x : α) => (act x).task) Task.Priority.default sync }
                  Instances For

                    Chains two snapshot tasks. The range is taken from the first task if not specified; the range of the second task is discarded.

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

                      Synchronously waits on the result of the task.

                      Equations
                      • t.get = t.task.get
                      Instances For

                        Returns task result if already finished or else none.

                        Equations
                        Instances For

                          Arbitrary value paired with a syntax that should be inspected when considering the value for reuse.

                          • Syntax to be inspected for reuse.

                          • val : α

                            Potentially reusable value.

                          Instances For

                            Pair of (optional) old snapshot task usable for incremental reuse and new snapshot promise for incremental reporting. Inside the elaborator, we build snapshots by carrying such bundles and then checking if we can reuse old? if set or else redoing the corresponding elaboration step. In either case, we derive new bundles for nested snapshots, if any, and finally resolve new to the result.

                            Note that failing to resolve a created promise will block the language server indefinitely! Corresponding IO.Promise.new calls should come with a "definitely resolved in ..." comment explaining how this is avoided in each case.

                            In the future, the 1-element history old? may be replaced with a global cache indexed by strong hashes but the promise will still need to be passed through the elaborator.

                            • Snapshot task of corresponding elaboration in previous document version if any, paired with its old syntax to be considered for reuse. Should be set to none as soon as reuse can be ruled out.

                            • new : IO.Promise α

                              Promise of snapshot value for the current document. When resolved, the language server will report its result even before the current elaborator invocation has finished.

                            Instances For

                              Tree of snapshots where each snapshot comes with an array of asynchronous further subtrees. Used for asynchronously collecting information about the entirety of snapshots in the language server. The involved tasks may form a DAG on the Task dependency level but this is not captured by this data structure.

                              Instances For
                                @[reducible, inline]

                                The immediately available element of the snapshot tree node.

                                Equations
                                Instances For
                                  @[reducible, inline]

                                  The asynchronously available children of the snapshot tree node.

                                  Equations
                                  Instances For

                                    Helper class for projecting a heterogeneous hierarchy of snapshot classes to a homogeneous representation.

                                    Instances
                                      Equations

                                      Snapshot type without child nodes.

                                        Instances For

                                          Arbitrary snapshot type, used for extensibility.

                                          Instances For

                                            Returns the original snapshot value if it is of the given type.

                                            Equations
                                            Instances For
                                              @[specialize #[]]

                                              Runs a tree of snapshots to conclusion, incrementally performing f on each snapshot in tree preorder.

                                              Option for printing end position of each message in addition to start position. Used for testing message ranges in the test suite.

                                              Reports messages on stdout. If json is true, prints messages as JSON (one per line).

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

                                                Runs a tree of snapshots to conclusion and incrementally report messages on stdout. Messages are reported in tree preorder. This function is used by the cmdline driver; see Lean.Server.FileWorker.reportSnapshots for how the language server reports snapshots asynchronously.

                                                Equations
                                                Instances For

                                                  Waits on and returns all snapshots in the tree.

                                                  Equations
                                                  Instances For

                                                    Metadata that does not change during the lifetime of the language processing process.

                                                    • mainModuleName : Lake.Name

                                                      Module name of the file being processed.

                                                    • Options provided outside of the file content, e.g. on the cmdline or in the lakefile.

                                                    • trustLevel : UInt32

                                                      Kernel trust level.

                                                    Instances For

                                                      Context of an input processing invocation.

                                                        Instances For
                                                          @[reducible, inline]
                                                          abbrev Lean.Language.ProcessingT (m : TypeType u_1) (α : Type) :
                                                          Type u_1

                                                          Monad transformer holding all relevant data for processing.

                                                          Equations
                                                          Instances For
                                                            @[reducible, inline]

                                                            Monad holding all relevant data for processing.

                                                            Equations
                                                            Instances For

                                                              Creates snapshot message log from non-interactive message log, also allocating a mutable cell that can be used by the server to memorize interactive diagnostics derived from the log.

                                                              Equations
                                                              Instances For

                                                                Creates diagnostics from a single error message that should span the whole file.

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

                                                                  Adds unexpected exceptions from header processing to the message log as a last resort; standard errors should already have been caught earlier.

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

                                                                    Builds a function for processing a language using incremental snapshots by passing the previous snapshot to Language.process on subsequent invocations.

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