Documentation

Mathlib.SetTheory.Game.Nim

Nim and the Sprague-Grundy theorem #

This file contains the definition for nim for any ordinal o. In the game of nim o₁ both players may move to nim o₂ for any o₂ < o₁. We also define a Grundy value for an impartial game G and prove the Sprague-Grundy theorem, that G is equivalent to nim (grundyValue G). Finally, we prove that the grundy value of a sum G + H corresponds to the nimber sum of the individual grundy values.

Implementation details #

The pen-and-paper definition of nim defines the possible moves of nim o to be Set.Iio o. However, this definition does not work for us because it would make the type of nim Ordinal.{u} → SetTheory.PGame.{u + 1}, which would make it impossible for us to state the Sprague-Grundy theorem, since that requires the type of nim to be Ordinal.{u} → SetTheory.PGame.{u}. For this reason, we instead use o.toType for the possible moves. We expose toLeftMovesNim and toRightMovesNim to conveniently convert an ordinal less than o into a left or right move of nim o, and vice versa.

@[irreducible]
noncomputable def SetTheory.PGame.nim (o : Ordinal.{u}) :

The definition of single-heap nim, which can be viewed as a pile of stones where each player can take a positive number of stones from it on their turn.

Equations
  • One or more equations did not get rendered due to their size.
Instances For
    @[deprecated "you can use `rw [nim]` directly" (since := "2025-01-23")]
    theorem SetTheory.PGame.nim_def (o : Ordinal.{u_1}) :
    nim o = mk o.toType o.toType (fun (x : o.toType) => nim (o.enumIsoToType.symm x)) fun (x : o.toType) => nim (o.enumIsoToType.symm x)

    Turns an ordinal less than o into a left move for nim o and vice versa.

    Equations
    Instances For

      Turns an ordinal less than o into a right move for nim o and vice versa.

      Equations
      Instances For
        @[deprecated SetTheory.PGame.moveLeft_nim (since := "2024-10-30")]

        Alias of SetTheory.PGame.moveLeft_nim.

        @[deprecated SetTheory.PGame.moveRight_nim (since := "2024-10-30")]

        Alias of SetTheory.PGame.moveRight_nim.

        def SetTheory.PGame.leftMovesNimRecOn {o : Ordinal.{u_2}} {P : (nim o).LeftMovesSort u_1} (i : (nim o).LeftMoves) (H : (a : Ordinal.{u_2}) → (H : a < o) → P (toLeftMovesNim a, H)) :
        P i

        A recursion principle for left moves of a nim game.

        Equations
        Instances For
          def SetTheory.PGame.rightMovesNimRecOn {o : Ordinal.{u_2}} {P : (nim o).RightMovesSort u_1} (i : (nim o).RightMoves) (H : (a : Ordinal.{u_2}) → (H : a < o) → P (toRightMovesNim a, H)) :
          P i

          A recursion principle for right moves of a nim game.

          Equations
          Instances For

            nim 1 has exactly the same moves as star.

            Equations
            • One or more equations did not get rendered due to their size.
            Instances For
              @[simp]
              theorem SetTheory.PGame.nim_add_equiv_zero_iff (o₁ o₂ : Ordinal.{u_1}) :
              nim o₁ + nim o₂ 0 o₁ = o₂
              @[simp]
              theorem SetTheory.PGame.nim_add_fuzzy_zero_iff {o₁ o₂ : Ordinal.{u_1}} :
              (nim o₁ + nim o₂).Fuzzy 0 o₁ o₂
              @[simp]
              theorem SetTheory.PGame.nim_equiv_iff_eq {o₁ o₂ : Ordinal.{u_1}} :
              nim o₁ nim o₂ o₁ = o₂
              @[irreducible]
              noncomputable def SetTheory.PGame.grundyValue (G : PGame) :

              The Grundy value of an impartial game is recursively defined as the minimum excluded value (the infimum of the complement) of the Grundy values of either its left or right options.

              This is the ordinal which corresponds to the game of nim that the game is equivalent to.

              This function takes a value in Nimber. This is a type synonym for the ordinals which has the same ordering, but addition in Nimber is such that it corresponds to the grundy value of the addition of games. See that file for more information on nimbers and their arithmetic.

              Equations
              Instances For
                @[deprecated SetTheory.PGame.grundyValue_eq_sInf_moveLeft (since := "2024-09-16")]
                @[irreducible]

                The Sprague-Grundy theorem states that every impartial game is equivalent to a game of nim, namely the game of nim corresponding to the game's Grundy value.

                @[deprecated SetTheory.PGame.grundyValue_eq_sInf_moveRight (since := "2024-09-16")]
                @[irreducible]

                The Grundy value of the sum of two nim games equals their nimber addition.