Documentation

Mathlib.SetTheory.Game.Basic

Combinatorial games. #

In this file we construct an instance OrderedAddCommGroup SetTheory.Game.

Multiplication on pre-games #

We define the operations of multiplication and inverse on pre-games, and prove a few basic theorems about them. Multiplication is not well-behaved under equivalence of pre-games i.e. x ≈ y does not imply x * z ≈ y * z. Hence, multiplication is not a well-defined operation on games. Nevertheless, the abelian group structure on games allows us to simplify many proofs for pre-games.

@[reducible, inline]
abbrev SetTheory.Game :
Type (u_1 + 1)

The type of combinatorial games. In ZFC, a combinatorial game is constructed from two sets of combinatorial games that have been constructed at an earlier stage. To do this in type theory, we say that a combinatorial pre-game is built inductively from two families of combinatorial games indexed over any type in Type u. The resulting type PGame.{u} lives in Type (u+1), reflecting that it is a proper class in ZFC. A combinatorial game is then constructed by quotienting by the equivalence x ≈ y ↔ x ≤ y ∧ y ≤ x.

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

        The less or fuzzy relation on games.

        If 0 ⧏ x (less or fuzzy with), then Left can win x as the first player.

        Equations
        Instances For
          theorem SetTheory.Game.not_le {x y : Game} :
          Iff (Not (LE.le x y)) (y.LF x)

          On Game, simp-normal inequalities should use as few negations as possible.

          theorem SetTheory.Game.not_lf {x y : Game} :
          Iff (Not (x.LF y)) (LE.le y x)

          On Game, simp-normal inequalities should use as few negations as possible.

          The fuzzy, confused, or incomparable relation on games.

          If x ‖ 0, then the first player can always win x.

          Equations
          Instances For

            It can be useful to use these lemmas to turn PGame inequalities into Game inequalities, as the AddCommGroup structure on Game often simplifies many proofs.

            theorem SetTheory.Game.add_lf_add_right {b c : Game} :
            b.LF c∀ (a : Game), (HAdd.hAdd b a).LF (HAdd.hAdd c a)
            theorem SetTheory.Game.add_lf_add_left {b c : Game} :
            b.LF c∀ (a : Game), (HAdd.hAdd a b).LF (HAdd.hAdd a c)
            Equations
            • One or more equations did not get rendered due to their size.
            Instances For
              theorem SetTheory.Game.bddAbove_range_of_small {ι : Type u_1} [Small ι] (f : ιGame) :

              A small family of games is bounded above.

              A small set of games is bounded above.

              theorem SetTheory.Game.bddBelow_range_of_small {ι : Type u_1} [Small ι] (f : ιGame) :

              A small family of games is bounded below.

              A small set of games is bounded below.

              Multiplicative operations can be defined at the level of pre-games, but to prove their properties we need to use the abelian group structure of games. Hence we define them here.

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

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

                Turns two left or right moves for x and 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

                  Turns a left and a right move for x and 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
                    theorem SetTheory.PGame.mk_mul_moveLeft_inl {xl xr yl yr : Type u_1} {xL : xlPGame} {xR : xrPGame} {yL : ylPGame} {yR : yrPGame} {i : xl} {j : yl} :
                    Eq ((HMul.hMul (mk xl xr xL xR) (mk yl yr yL yR)).moveLeft (Sum.inl { fst := i, snd := j })) (HSub.hSub (HAdd.hAdd (HMul.hMul (xL i) (mk yl yr yL yR)) (HMul.hMul (mk xl xr xL xR) (yL j))) (HMul.hMul (xL i) (yL j)))
                    theorem SetTheory.PGame.mk_mul_moveLeft_inr {xl xr yl yr : Type u_1} {xL : xlPGame} {xR : xrPGame} {yL : ylPGame} {yR : yrPGame} {i : xr} {j : yr} :
                    Eq ((HMul.hMul (mk xl xr xL xR) (mk yl yr yL yR)).moveLeft (Sum.inr { fst := i, snd := j })) (HSub.hSub (HAdd.hAdd (HMul.hMul (xR i) (mk yl yr yL yR)) (HMul.hMul (mk xl xr xL xR) (yR j))) (HMul.hMul (xR i) (yR j)))
                    theorem SetTheory.PGame.mk_mul_moveRight_inl {xl xr yl yr : Type u_1} {xL : xlPGame} {xR : xrPGame} {yL : ylPGame} {yR : yrPGame} {i : xl} {j : yr} :
                    Eq ((HMul.hMul (mk xl xr xL xR) (mk yl yr yL yR)).moveRight (Sum.inl { fst := i, snd := j })) (HSub.hSub (HAdd.hAdd (HMul.hMul (xL i) (mk yl yr yL yR)) (HMul.hMul (mk xl xr xL xR) (yR j))) (HMul.hMul (xL i) (yR j)))
                    theorem SetTheory.PGame.mk_mul_moveRight_inr {xl xr yl yr : Type u_1} {xL : xlPGame} {xR : xrPGame} {yL : ylPGame} {yR : yrPGame} {i : xr} {j : yl} :
                    Eq ((HMul.hMul (mk xl xr xL xR) (mk yl yr yL yR)).moveRight (Sum.inr { fst := i, snd := j })) (HSub.hSub (HAdd.hAdd (HMul.hMul (xR i) (mk yl yr yL yR)) (HMul.hMul (mk xl xr xL xR) (yL j))) (HMul.hMul (xR i) (yL j)))
                    theorem SetTheory.PGame.neg_mk_mul_moveLeft_inl {xl xr yl yr : Type u_1} {xL : xlPGame} {xR : xrPGame} {yL : ylPGame} {yR : yrPGame} {i : xl} {j : yr} :
                    Eq ((Neg.neg (HMul.hMul (mk xl xr xL xR) (mk yl yr yL yR))).moveLeft (Sum.inl { fst := i, snd := j })) (Neg.neg (HSub.hSub (HAdd.hAdd (HMul.hMul (xL i) (mk yl yr yL yR)) (HMul.hMul (mk xl xr xL xR) (yR j))) (HMul.hMul (xL i) (yR j))))
                    theorem SetTheory.PGame.neg_mk_mul_moveLeft_inr {xl xr yl yr : Type u_1} {xL : xlPGame} {xR : xrPGame} {yL : ylPGame} {yR : yrPGame} {i : xr} {j : yl} :
                    Eq ((Neg.neg (HMul.hMul (mk xl xr xL xR) (mk yl yr yL yR))).moveLeft (Sum.inr { fst := i, snd := j })) (Neg.neg (HSub.hSub (HAdd.hAdd (HMul.hMul (xR i) (mk yl yr yL yR)) (HMul.hMul (mk xl xr xL xR) (yL j))) (HMul.hMul (xR i) (yL j))))
                    theorem SetTheory.PGame.neg_mk_mul_moveRight_inl {xl xr yl yr : Type u_1} {xL : xlPGame} {xR : xrPGame} {yL : ylPGame} {yR : yrPGame} {i : xl} {j : yl} :
                    Eq ((Neg.neg (HMul.hMul (mk xl xr xL xR) (mk yl yr yL yR))).moveRight (Sum.inl { fst := i, snd := j })) (Neg.neg (HSub.hSub (HAdd.hAdd (HMul.hMul (xL i) (mk yl yr yL yR)) (HMul.hMul (mk xl xr xL xR) (yL j))) (HMul.hMul (xL i) (yL j))))
                    theorem SetTheory.PGame.neg_mk_mul_moveRight_inr {xl xr yl yr : Type u_1} {xL : xlPGame} {xR : xrPGame} {yL : ylPGame} {yR : yrPGame} {i : xr} {j : yr} :
                    Eq ((Neg.neg (HMul.hMul (mk xl xr xL xR) (mk yl yr yL yR))).moveRight (Sum.inr { fst := i, snd := j })) (Neg.neg (HSub.hSub (HAdd.hAdd (HMul.hMul (xR i) (mk yl yr yL yR)) (HMul.hMul (mk xl xr xL xR) (yR j))) (HMul.hMul (xR i) (yR j))))
                    theorem SetTheory.PGame.leftMoves_mul_cases {x y : PGame} (k : (HMul.hMul x y).LeftMoves) {P : (HMul.hMul x y).LeftMovesProp} (hl : ∀ (ix : x.LeftMoves) (iy : y.LeftMoves), P (DFunLike.coe toLeftMovesMul (Sum.inl { fst := ix, snd := iy }))) (hr : ∀ (jx : x.RightMoves) (jy : y.RightMoves), P (DFunLike.coe toLeftMovesMul (Sum.inr { fst := jx, snd := jy }))) :
                    P k
                    theorem SetTheory.PGame.rightMoves_mul_cases {x y : PGame} (k : (HMul.hMul x y).RightMoves) {P : (HMul.hMul x y).RightMovesProp} (hl : ∀ (ix : x.LeftMoves) (jy : y.RightMoves), P (DFunLike.coe toRightMovesMul (Sum.inl { fst := ix, snd := jy }))) (hr : ∀ (jx : x.RightMoves) (iy : y.LeftMoves), P (DFunLike.coe toRightMovesMul (Sum.inr { fst := jx, snd := iy }))) :
                    P k
                    @[irreducible]

                    x * y and y * x have the same moves.

                    @[irreducible]

                    x * y and y * x have the same moves.

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

                      x * y is equivalent to y * x.

                      x * 0 has exactly the same moves as 0.

                      x * 0 has exactly the same moves as 0.

                      Equations
                      Instances For

                        x * 0 is equivalent to 0.

                        0 * x has exactly the same moves as 0.

                        0 * x has exactly the same moves as 0.

                        Equations
                        Instances For

                          0 * x is equivalent to 0.

                          @[irreducible]

                          -x * y and -(x * y) have the same moves.

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

                            x * -y and -(x * y) have the same moves.

                            -x * y and -(x * y) have the same moves.

                            x * -y and -(x * y) have the same moves.

                            Equations
                            Instances For

                              x * (y + z) is equivalent to x * y + x * z.

                              (x + y) * z is equivalent to x * z + y * z.

                              x * 1 has the same moves as x.

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

                                1 * x has the same moves as x.

                                x * 1 has the same moves as x.

                                x * 1 is equivalent to x.

                                1 * x is equivalent to x.

                                x * y * z is equivalent to x * (y * z).

                                The left options of x * y of the first kind, i.e. of the form xL * y + x * yL - xL * yL.

                                Equations
                                Instances For

                                  Any left option of x * y of the first kind is also a left option of x * -(-y) of the first kind.

                                  The left options of x * y agree with that of y * x up to equivalence.

                                  theorem SetTheory.PGame.leftMoves_mul_iff {x y : PGame} (P : GameProp) :
                                  Iff (∀ (k : (HMul.hMul x y).LeftMoves), P (Quotient.mk setoid ((HMul.hMul x y).moveLeft k))) (And (∀ (i : x.LeftMoves) (j : y.LeftMoves), P (Quotient.mk setoid (x.mulOption y i j))) (∀ (i : (Neg.neg x).LeftMoves) (j : (Neg.neg y).LeftMoves), P (Quotient.mk setoid ((Neg.neg x).mulOption (Neg.neg y) i j))))

                                  The left options of x * y of the second kind are the left options of (-x) * (-y) of the first kind, up to equivalence.

                                  theorem SetTheory.PGame.rightMoves_mul_iff {x y : PGame} (P : GameProp) :
                                  Iff (∀ (k : (HMul.hMul x y).RightMoves), P (Quotient.mk setoid ((HMul.hMul x y).moveRight k))) (And (∀ (i : x.LeftMoves) (j : (Neg.neg y).LeftMoves), P (Neg.neg (Quotient.mk setoid (x.mulOption (Neg.neg y) i j)))) (∀ (i : (Neg.neg x).LeftMoves) (j : y.LeftMoves), P (Neg.neg (Quotient.mk setoid ((Neg.neg x).mulOption y i j)))))

                                  The right options of x * y are the left options of x * (-y) and of (-x) * y of the first kind, up to equivalence.

                                  inductive SetTheory.PGame.InvTy (l r : Type u) :
                                  BoolType u

                                  Because the two halves of the definition of inv produce more elements on each side, we have to define the two families inductively. This is the indexing set for the function, and invVal is the function part.

                                  Instances For
                                    Equations
                                    Instances For
                                      def SetTheory.PGame.invVal {l r : Type u_1} (L : lPGame) (R : rPGame) (IHl : lPGame) (IHr : rPGame) (x : PGame) {b : Bool} :
                                      InvTy l r bPGame

                                      Because the two halves of the definition of inv produce more elements of each side, we have to define the two families inductively. This is the function part, defined by recursion on InvTy.

                                      Equations
                                      • One or more equations did not get rendered due to their size.
                                      Instances For
                                        theorem SetTheory.PGame.invVal_isEmpty {l r : Type u} {b : Bool} (L : lPGame) (R : rPGame) (IHl : lPGame) (IHr : rPGame) (i : InvTy l r b) (x : PGame) [IsEmpty l] [IsEmpty r] :
                                        Eq (invVal L R IHl IHr x i) 0

                                        The inverse of a positive surreal number x = {L | R} is given by x⁻¹ = {0, (1 + (R - x) * x⁻¹L) * R, (1 + (L - x) * x⁻¹R) * L | (1 + (L - x) * x⁻¹L) * L, (1 + (R - x) * x⁻¹R) * R}. Because the two halves x⁻¹L, x⁻¹R of x⁻¹ are used in their own definition, the sets and elements are inductively generated.

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

                                          inv' 0 has exactly the same moves as 1.

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

                                            inv' 1 has exactly the same moves as 1.

                                            inv' 1 has exactly the same moves as 1.

                                            Equations
                                            • One or more equations did not get rendered due to their size.
                                            Instances For
                                              noncomputable def SetTheory.PGame.instInv :

                                              The inverse of a pre-game in terms of the inverse on positive pre-games.

                                              Equations
                                              Instances For
                                                noncomputable def SetTheory.PGame.instDiv :
                                                Equations
                                                Instances For

                                                  1⁻¹ has exactly the same moves as 1.