Documentation

Lean.Log

class Lean.MonadLog (m : TypeType) extends Lean.MonadFileMap m :

The MonadLog interface for logging error messages.

  • getRef : m Syntax

    Return the current reference syntax being used to provide position information.

  • getFileName : m String

    The name of the file being processed.

  • hasErrors : m Bool

    Return true if errors have been logged.

  • logMessage : Messagem Unit

    Log a new message.

Instances
    def Lean.getRefPos {m : TypeType} [Monad m] [MonadLog m] :

    Return the position (as String.pos) associated with the current reference syntax (i.e., the syntax object returned by getRef.)

    Equations
    Instances For

      Return the line and column numbers associated with the current reference syntax (i.e., the syntax object returned by getRef.)

      Equations
      Instances For

        If warningAsError is set to true, then warning messages are treated as errors.

        def Lean.logAt {m : TypeType} [Monad m] [MonadLog m] [AddMessageContext m] [MonadOptions m] (ref : Syntax) (msgData : MessageData) (severity : MessageSeverity := MessageSeverity.error) :

        Log the message msgData at the position provided by ref with the given severity. If getRef has position information but ref does not, we use getRef. We use the fileMap to find the line and column numbers for the error message.

        Equations
        • One or more equations did not get rendered due to their size.
        Instances For
          def Lean.logErrorAt {m : TypeType} [Monad m] [MonadLog m] [AddMessageContext m] [MonadOptions m] (ref : Syntax) (msgData : MessageData) :

          Log a new error message using the given message data. The position is provided by ref.

          Equations
          Instances For
            def Lean.logWarningAt {m : TypeType} [Monad m] [MonadLog m] [AddMessageContext m] [MonadOptions m] (ref : Syntax) (msgData : MessageData) :

            Log a new warning message using the given message data. The position is provided by ref.

            Equations
            Instances For
              def Lean.logInfoAt {m : TypeType} [Monad m] [MonadLog m] [AddMessageContext m] [MonadOptions m] (ref : Syntax) (msgData : MessageData) :

              Log a new information message using the given message data. The position is provided by ref.

              Equations
              Instances For

                Log a new error/warning/information message using the given message data and severity. The position is provided by getRef.

                Equations
                Instances For
                  def Lean.logError {m : TypeType} [Monad m] [MonadLog m] [AddMessageContext m] [MonadOptions m] (msgData : MessageData) :

                  Log a new error message using the given message data. The position is provided by getRef.

                  Equations
                  Instances For

                    Log a new warning message using the given message data. The position is provided by getRef.

                    Equations
                    • One or more equations did not get rendered due to their size.
                    Instances For
                      def Lean.logInfo {m : TypeType} [Monad m] [MonadLog m] [AddMessageContext m] [MonadOptions m] (msgData : MessageData) :

                      Log a new information message using the given message data. The position is provided by getRef.

                      Equations
                      Instances For
                        def Lean.logUnknownDecl {m : TypeType} [Monad m] [MonadLog m] [AddMessageContext m] [MonadOptions m] (declName : Name) :

                        Log the error message "unknown declaration"

                        Equations
                        Instances For