Documentation

Mathlib.SetTheory.PGame.Algebra

Algebraic structure on pregames #

This file defines the operations necessary to make games into an additive commutative group. Addition is defined for $x = \{xL | xR\}$ and $y = \{yL | yR\}$ by $x + y = \{xL + y, x + yL | xR + y, x + yR\}$. Negation is defined by $\{xL | xR\} = \{-xR | -xL\}$.

The order structures interact in the expected way with addition, so we have

theorem le_iff_sub_nonneg {x y : PGame} : x ≤ y ↔ 0 ≤ y - x := sorry
theorem lt_iff_sub_pos {x y : PGame} : x < y ↔ 0 < y - x := sorry

We show that these operations respect the equivalence relation, and hence descend to games. At the level of games, these operations satisfy all the laws of a commutative group. To prove the necessary equivalence relations at the level of pregames, the notion of a Relabelling of a game is used (defined in Mathlib.SetTheory.PGame.Basic); for example, there is a relabelling between x + (y + z) and (x + y) + z.

Negation #

The negation of {L | R} is {-R | -L}.

Equations
Instances For
    @[simp]
    theorem SetTheory.PGame.neg_def {xl xr : Type u_1} {xL : xlPGame} {xR : xrPGame} :
    -mk xl xr xL xR = mk xr xl (fun (x : xr) => -xR x) fun (x : xl) => -xL x
    @[simp]
    theorem SetTheory.PGame.neg_ofLists (L R : List PGame) :
    -ofLists L R = ofLists (List.map (fun (x : PGame) => -x) R) (List.map (fun (x : PGame) => -x) L)

    Use toLeftMovesNeg to cast between these two types.

    Use toRightMovesNeg to cast between these two types.

    Turns a right move for x into a left move for -x and vice versa.

    Even though these types are the same (not definitionally so), this is the preferred way to convert between them.

    Equations
    Instances For

      Turns a left move for x into a right move for -x and vice versa.

      Even though these types are the same (not definitionally so), this is the preferred way to convert between them.

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

        Alias of SetTheory.PGame.moveLeft_neg.

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

        Alias of SetTheory.PGame.moveRight_neg.

        @[deprecated SetTheory.PGame.moveRight_neg (since := "2024-10-30")]
        @[deprecated SetTheory.PGame.moveRight_neg (since := "2024-10-30")]
        @[deprecated SetTheory.PGame.moveLeft_neg (since := "2024-10-30")]
        @[deprecated SetTheory.PGame.moveLeft_neg (since := "2024-10-30")]
        @[simp]
        theorem SetTheory.PGame.forall_leftMoves_neg {x : PGame} {p : (-x).LeftMovesProp} :
        (∀ (i : (-x).LeftMoves), p i) ∀ (i : x.RightMoves), p (toLeftMovesNeg i)
        @[simp]
        theorem SetTheory.PGame.exists_leftMoves_neg {x : PGame} {p : (-x).LeftMovesProp} :
        (∃ (i : (-x).LeftMoves), p i) ∃ (i : x.RightMoves), p (toLeftMovesNeg i)
        @[simp]
        theorem SetTheory.PGame.forall_rightMoves_neg {x : PGame} {p : (-x).RightMovesProp} :
        (∀ (i : (-x).RightMoves), p i) ∀ (i : x.LeftMoves), p (toRightMovesNeg i)
        @[simp]
        theorem SetTheory.PGame.exists_rightMoves_neg {x : PGame} {p : (-x).RightMovesProp} :
        (∃ (i : (-x).RightMoves), p i) ∃ (i : x.LeftMoves), p (toRightMovesNeg i)
        theorem SetTheory.PGame.leftMoves_neg_cases {x : PGame} (k : (-x).LeftMoves) {P : (-x).LeftMovesProp} (h : ∀ (i : x.RightMoves), P (toLeftMovesNeg i)) :
        P k
        theorem SetTheory.PGame.rightMoves_neg_cases {x : PGame} (k : (-x).RightMoves) {P : (-x).RightMovesProp} (h : ∀ (i : x.LeftMoves), P (toRightMovesNeg i)) :
        P k
        theorem SetTheory.PGame.Identical.neg {x₁ x₂ : PGame} :
        x₁.Identical x₂(-x₁).Identical (-x₂)

        If x has the same moves as y, then -x has the sames moves as -y.

        theorem SetTheory.PGame.Identical.of_neg {x₁ x₂ : PGame} :
        (-x₁).Identical (-x₂)x₁.Identical x₂

        If -x has the same moves as -y, then x has the sames moves as y.

        theorem SetTheory.PGame.memₗ_neg_iff {x y : PGame} :
        x.memₗ (-y) ∃ (z : PGame), z.memᵣ y x.Identical (-z)
        theorem SetTheory.PGame.memᵣ_neg_iff {x y : PGame} :
        x.memᵣ (-y) ∃ (z : PGame), z.memₗ y x.Identical (-z)

        If x has the same moves as y, then -x has the same moves as -y.

        Equations
        • One or more equations did not get rendered due to their size.
        Instances For
          @[simp]
          @[simp]
          theorem SetTheory.PGame.neg_lf_neg_iff {x y : PGame} :
          (-y).LF (-x) x.LF y
          @[simp]
          theorem SetTheory.PGame.neg_lt_neg_iff {x y : PGame} :
          -y < -x x < y
          @[simp]
          @[simp]
          theorem SetTheory.PGame.neg_lf_iff {x y : PGame} :
          (-y).LF x (-x).LF y
          theorem SetTheory.PGame.lf_neg_iff {x y : PGame} :
          y.LF (-x) x.LF (-y)
          @[simp]
          @[simp]
          @[simp]
          @[simp]
          @[simp]
          @[simp]

          Addition and subtraction #

          The sum of x = {xL | xR} and y = {yL | yR} is {xL + y, x + yL | xR + y, x + yR}.

          Equations
          • One or more equations did not get rendered due to their size.
          theorem SetTheory.PGame.mk_add_moveLeft {xl xr yl yr : Type u_1} {xL : xlPGame} {xR : xrPGame} {yL : ylPGame} {yR : yrPGame} {i : (mk xl xr xL xR + mk yl yr yL yR).LeftMoves} :
          (mk xl xr xL xR + mk yl yr yL yR).moveLeft i = Sum.rec (fun (x : xl) => xL x + mk yl yr yL yR) (fun (x : yl) => mk xl xr xL xR + yL x) i
          theorem SetTheory.PGame.mk_add_moveRight {xl xr yl yr : Type u_1} {xL : xlPGame} {xR : xrPGame} {yL : ylPGame} {yR : yrPGame} {i : (mk xl xr xL xR + mk yl yr yL yR).RightMoves} :
          (mk xl xr xL xR + mk yl yr yL yR).moveRight i = Sum.rec (fun (x : xr) => xR x + mk yl yr yL yR) (fun (x : yr) => mk xl xr xL xR + yR x) i

          The pre-game ((0 + 1) + ⋯) + 1.

          Note that this is not the usual recursive definition n = {0, 1, … | }. For instance, 2 = 0 + 1 + 1 = {0 + 0 + 1, 0 + 1 + 0 | } does not contain any left option equivalent to 0. For an implementation of said definition, see Ordinal.toPGame. For the proof that these games are equivalent, see Ordinal.toPGame_natCast.

          Equations
          @[simp]
          theorem SetTheory.PGame.nat_succ (n : ) :
          ↑(n + 1) = n + 1
          @[irreducible]

          x + 0 has exactly the same moves as x.

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

            x + 0 is equivalent to x.

            0 + x has exactly the same moves as x.

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

              0 + x is equivalent to x.

              Use toLeftMovesAdd to cast between these two types.

              Use toRightMovesAdd to cast between these two types.

              Converts a left move for x or y into a left move for x + y and vice versa.

              Even though these types are the same (not definitionally so), this is the preferred way to convert between them.

              Equations
              Instances For

                Converts a right move for x or y into a right move for x + y and vice versa.

                Even though these types are the same (not definitionally so), this is the preferred way to convert between them.

                Equations
                Instances For
                  @[simp]
                  theorem SetTheory.PGame.mk_add_moveLeft_inl {xl xr yl yr : Type u_1} {xL : xlPGame} {xR : xrPGame} {yL : ylPGame} {yR : yrPGame} {i : xl} :
                  (mk xl xr xL xR + mk yl yr yL yR).moveLeft (Sum.inl i) = (mk xl xr xL xR).moveLeft i + mk yl yr yL yR
                  @[simp]
                  @[simp]
                  theorem SetTheory.PGame.mk_add_moveRight_inl {xl xr yl yr : Type u_1} {xL : xlPGame} {xR : xrPGame} {yL : ylPGame} {yR : yrPGame} {i : xr} :
                  (mk xl xr xL xR + mk yl yr yL yR).moveRight (Sum.inl i) = (mk xl xr xL xR).moveRight i + mk yl yr yL yR
                  @[simp]
                  theorem SetTheory.PGame.mk_add_moveLeft_inr {xl xr yl yr : Type u_1} {xL : xlPGame} {xR : xrPGame} {yL : ylPGame} {yR : yrPGame} {i : yl} :
                  (mk xl xr xL xR + mk yl yr yL yR).moveLeft (Sum.inr i) = mk xl xr xL xR + (mk yl yr yL yR).moveLeft i
                  @[simp]
                  @[simp]
                  theorem SetTheory.PGame.mk_add_moveRight_inr {xl xr yl yr : Type u_1} {xL : xlPGame} {xR : xrPGame} {yL : ylPGame} {yR : yrPGame} {i : yr} :
                  (mk xl xr xL xR + mk yl yr yL yR).moveRight (Sum.inr i) = mk xl xr xL xR + (mk yl yr yL yR).moveRight i
                  theorem SetTheory.PGame.leftMoves_add_cases {x y : PGame} (k : (x + y).LeftMoves) {P : (x + y).LeftMovesProp} (hl : ∀ (i : x.LeftMoves), P (toLeftMovesAdd (Sum.inl i))) (hr : ∀ (i : y.LeftMoves), P (toLeftMovesAdd (Sum.inr i))) :
                  P k

                  Case on possible left moves of x + y.

                  theorem SetTheory.PGame.rightMoves_add_cases {x y : PGame} (k : (x + y).RightMoves) {P : (x + y).RightMovesProp} (hl : ∀ (j : x.RightMoves), P (toRightMovesAdd (Sum.inl j))) (hr : ∀ (j : y.RightMoves), P (toRightMovesAdd (Sum.inr j))) :
                  P k

                  Case on possible right moves of x + y.

                  @[irreducible]
                  theorem SetTheory.PGame.add_comm (x y : PGame) :
                  (x + y).Identical (y + x)

                  x + y has exactly the same moves as y + x.

                  @[irreducible]
                  theorem SetTheory.PGame.add_assoc (x y z : PGame) :
                  (x + y + z).Identical (x + (y + z))

                  (x + y) + z has exactly the same moves as x + (y + z).

                  theorem SetTheory.PGame.add_zero (x : PGame) :
                  (x + 0).Identical x

                  x + 0 has exactly the same moves as x.

                  theorem SetTheory.PGame.zero_add (x : PGame) :
                  (0 + x).Identical x

                  0 + x has exactly the same moves as x.

                  @[irreducible]
                  theorem SetTheory.PGame.neg_add (x y : PGame) :
                  -(x + y) = -x + -y

                  -(x + y) has exactly the same moves as -x + -y.

                  theorem SetTheory.PGame.neg_add_rev (x y : PGame) :
                  (-(x + y)).Identical (-y + -x)

                  -(x + y) has exactly the same moves as -y + -x.

                  Any game without left or right moves is identical to 0.

                  @[irreducible]
                  theorem SetTheory.PGame.Identical.add_right {x₁ x₂ y : PGame} :
                  x₁.Identical x₂(x₁ + y).Identical (x₂ + y)
                  theorem SetTheory.PGame.Identical.add_left {x y₁ y₂ : PGame} (hy : y₁.Identical y₂) :
                  (x + y₁).Identical (x + y₂)
                  theorem SetTheory.PGame.Identical.add {x₁ x₂ y₁ y₂ : PGame} (hx : x₁.Identical x₂) (hy : y₁.Identical y₂) :
                  (x₁ + y₁).Identical (x₂ + y₂)

                  If w has the same moves as x and y has the same moves as z, then w + y has the same moves as x + z.

                  theorem SetTheory.PGame.memₗ_add_iff {x y₁ y₂ : PGame} :
                  x.memₗ (y₁ + y₂) (∃ (z : PGame), z.memₗ y₁ x.Identical (z + y₂)) ∃ (z : PGame), z.memₗ y₂ x.Identical (y₁ + z)
                  theorem SetTheory.PGame.memᵣ_add_iff {x y₁ y₂ : PGame} :
                  x.memᵣ (y₁ + y₂) (∃ (z : PGame), z.memᵣ y₁ x.Identical (z + y₂)) ∃ (z : PGame), z.memᵣ y₂ x.Identical (y₁ + z)
                  @[irreducible]
                  def SetTheory.PGame.Relabelling.addCongr {w x y z : PGame} :
                  w.Relabelling xy.Relabelling z(w + y).Relabelling (x + z)

                  If w has the same moves as x and y has the same moves as z, then w + y has the same moves as x + z.

                  Equations
                  • One or more equations did not get rendered due to their size.
                  Instances For
                    Equations
                    @[simp]
                    theorem SetTheory.PGame.neg_sub' (x y : PGame) :
                    -(x - y) = -x - -y

                    This lemma is named to match neg_sub'.

                    theorem SetTheory.PGame.Identical.sub {x₁ x₂ y₁ y₂ : PGame} (hx : x₁.Identical x₂) (hy : y₁.Identical y₂) :
                    (x₁ - y₁).Identical (x₂ - y₂)

                    If w has the same moves as x and y has the same moves as z, then w - y has the same moves as x - z.

                    def SetTheory.PGame.Relabelling.subCongr {w x y z : PGame} (h₁ : w.Relabelling x) (h₂ : y.Relabelling z) :
                    (w - y).Relabelling (x - z)

                    If w has the same moves as x and y has the same moves as z, then w - y has the same moves as x - z.

                    Equations
                    Instances For
                      @[irreducible]

                      -(x + y) has exactly the same moves as -x + -y.

                      Equations
                      • One or more equations did not get rendered due to their size.
                      Instances For
                        theorem SetTheory.PGame.neg_add_le {x y : PGame} :
                        -(x + y) -x + -y
                        @[irreducible]

                        x + y has exactly the same moves as y + x.

                        Equations
                        • One or more equations did not get rendered due to their size.
                        Instances For
                          theorem SetTheory.PGame.add_comm_le {x y : PGame} :
                          x + y y + x
                          @[irreducible]
                          def SetTheory.PGame.addAssocRelabelling (x y z : PGame) :
                          (x + y + z).Relabelling (x + (y + z))

                          (x + y) + z has exactly the same moves as x + (y + z).

                          Equations
                          • One or more equations did not get rendered due to their size.
                          Instances For
                            theorem SetTheory.PGame.add_assoc_equiv {x y z : PGame} :
                            x + y + z x + (y + z)
                            theorem SetTheory.PGame.add_lf_add_right {y z : PGame} (h : y.LF z) (x : PGame) :
                            (y + x).LF (z + x)
                            theorem SetTheory.PGame.add_lf_add_left {y z : PGame} (h : y.LF z) (x : PGame) :
                            (x + y).LF (x + z)
                            theorem SetTheory.PGame.add_lf_add_of_lf_of_le {w x y z : PGame} (hwx : w.LF x) (hyz : y z) :
                            (w + y).LF (x + z)
                            theorem SetTheory.PGame.add_lf_add_of_le_of_lf {w x y z : PGame} (hwx : w x) (hyz : y.LF z) :
                            (w + y).LF (x + z)
                            theorem SetTheory.PGame.add_congr {w x y z : PGame} (h₁ : w x) (h₂ : y z) :
                            w + y x + z
                            theorem SetTheory.PGame.add_congr_left {x y z : PGame} (h : x y) :
                            x + z y + z
                            theorem SetTheory.PGame.add_congr_right {x y z : PGame} :
                            y zx + y x + z
                            theorem SetTheory.PGame.sub_congr {w x y z : PGame} (h₁ : w x) (h₂ : y z) :
                            w - y x - z
                            theorem SetTheory.PGame.sub_congr_left {x y z : PGame} (h : x y) :
                            x - z y - z
                            theorem SetTheory.PGame.sub_congr_right {x y z : PGame} :
                            y zx - y x - z
                            theorem SetTheory.PGame.lt_iff_sub_pos {x y : PGame} :
                            x < y 0 < y - x

                            Interaction of option insertion with negation #

                            Special pre-games #

                            The pre-game star, which is fuzzy with zero.

                            Equations
                            Instances For
                              @[simp]