Documentation

Mathlib.Tactic.Linter.CommandStart

The commandStart linter #

The commandStart linter emits a warning if

The commandStart linter emits a warning if

  • either a command does not start at the beginning of a line;
  • or the "hypotheses segment" of a declaration does not coincide with its pretty-printed version.

In practice, this makes sure that the spacing in a typical declaration looks like

example (a : Nat) {R : Type} [Add R] : <not linted part>

as opposed to

example (a: Nat) {R:Type}  [Add  R] : <not linted part>

If the linter.style.commandStart.verbose option is true, the commandStart linter reports some helpful diagnostic information.

CommandStart.endPos stx returns the position up until the commandStart linter checks the formatting. This is every declaration until the type-specification, if there is one, or the value, as well as all variable commands.

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

    A FormatError is the main structure tracking how some surface syntax differs from its pretty-printed version.

    In case of deviations, it contains the deviation's location within an ambient string.

    • srcNat : Nat

      The distance to the end of the source string, as number of characters

    • srcEndPos : String.Pos

      The distance to the end of the source string, as number of string positions

    • fmtPos : Nat

      The distance to the end of the formatted string, as number of characters

    • msg : String

      The kind of formatting error. For example: extra space, remove line break or missing space.

      Strings starting with Oh no indicate an internal error.

    • length : Nat

      The length of the mismatch, as number of characters.

    • srcStartPos : String.Pos

      The starting position of the mismatch, as a String.pos.

    Instances For
      Equations
      Equations
      • One or more equations did not get rendered due to their size.
      def Mathlib.Linter.mkFormatError (ls ms msg : String) (length : Nat := 1) :

      Produces a FormatError from the input data. It expects

      • ls to be a "user-typed" string;
      • ms to be a "pretty-printed" string;
      • msg to be a custom error message, such as extra space or remove line break;
      • length (optional with default 1), how many characters the error spans.

      In particular, it extracts the position information within the string, both as number of characters and as String.Pos.

      Equations
      Instances For

        Add a new FormatError f to the array fs, trying, as much as possible, to merge the new FormatError with the last entry of fs.

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

          Scan the two input strings L and M, assuming M is the pretty-printed version of L. This almost means that L and M only differ in whitespace.

          While scanning the two strings, accumulate any discrepancies --- with some heuristics to avoid flagging some line-breaking changes. (The pretty-printer does not always produce desirably formatted code.)

          Scan the two input strings L and M, assuming M is the pretty-printed version of L. This almost means that L and M only differ in whitespace.

          While scanning the two strings, accumulate any discrepancies --- with some heuristics to avoid flagging some line-breaking changes. (The pretty-printer does not always produce desirably formatted code.)

          Equations
          Instances For
            @[reducible, inline]

            unlintedNodes contains the SyntaxNodeKinds for which there is no clear formatting preference: if they appear in surface syntax, the linter will ignore formatting.

            Currently, the unlined nodes are mostly related to Subtype, Set and Finset notation and list notation.

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

              Given an array a of SyntaxNodeKinds, we accumulate the ranges of the syntax nodes of the input syntax whose kind is in a.

              The linter uses this information to avoid emitting a warning for nodes with kind contained in unlintedNodes.

              Equations
              Instances For

                Given a HashSet of String.Ranges rgs and a further String.Range rg, isOutside rgs rg returns false if and only if rgs contains a range that completely contains rg.

                The linter uses this to figure out which nodes should be ignored.

                Equations
                Instances For

                  mkWindow orig start ctx extracts from orig a string that starts at the first non-whitespace character before start, then expands to cover ctx more characters and continues still until the first non-whitespace character.

                  In essence, it extracts the substring of orig that begins at start, continues for ctx characters plus expands left and right until it encounters the first whitespace character, to avoid cutting into "words".

                  Note. start is the number of characters from the right where our focus is!

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

                    The commandStart linter emits a warning if

                    • either a command does not start at the beginning of a line;
                    • or the "hypotheses segment" of a declaration does not coincide with its pretty-printed version.

                    In practice, this makes sure that the spacing in a typical declaration looks like

                    example (a : Nat) {R : Type} [Add R] : <not linted part>
                    

                    as opposed to

                    example (a: Nat) {R:Type}  [Add  R] : <not linted part>
                    
                    Equations
                    • One or more equations did not get rendered due to their size.
                    Instances For