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
- (SetTheory.PGame.mk l r L R).neg = SetTheory.PGame.mk r l (fun (i : r) => (R i).neg) fun (i : l) => (L i).neg
Instances For
Equations
- SetTheory.PGame.instNeg = { neg := SetTheory.PGame.neg }
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
Alias of SetTheory.PGame.moveLeft_neg
.
Alias of SetTheory.PGame.moveRight_neg
.
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
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.
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
- SetTheory.PGame.instNatCast = { natCast := Nat.unaryCast }
x + 0
has exactly the same moves as x
.
Equations
- One or more equations did not get rendered due to their size.
Instances For
0 + x
has exactly the same moves as x
.
Equations
- One or more equations did not get rendered due to their size.
Instances For
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
Case on possible left moves of x + y
.
Case on possible right moves of x + y
.
x + 0
has exactly the same moves as x
.
0 + x
has exactly the same moves as x
.
Any game without left or right moves is identical to 0.
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
- SetTheory.PGame.instSub = { sub := fun (x y : SetTheory.PGame) => 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
.
Instances For
-(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
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
(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
Interaction of option insertion with negation #
Special pre-games #
The pre-game star
, which is fuzzy with zero.
Equations
- SetTheory.PGame.star = SetTheory.PGame.mk PUnit.{?u.1 + 1} PUnit.{?u.1 + 1} (fun (x : PUnit.{?u.1 + 1}) => 0) fun (x : PUnit.{?u.1 + 1}) => 0
Instances For
The pre-game up
Equations
- SetTheory.PGame.up = SetTheory.PGame.mk PUnit.{?u.1 + 1} PUnit.{?u.1 + 1} (fun (x : PUnit.{?u.1 + 1}) => 0) fun (x : PUnit.{?u.1 + 1}) => SetTheory.PGame.star
Instances For
The pre-game down
Equations
- SetTheory.PGame.down = SetTheory.PGame.mk PUnit.{?u.1 + 1} PUnit.{?u.1 + 1} (fun (x : PUnit.{?u.1 + 1}) => SetTheory.PGame.star) fun (x : PUnit.{?u.1 + 1}) => 0