Combinatorial (pre-)games. #
The basic theory of combinatorial games, following Conway's book On Numbers and Games
. We
construct "pregames", define an ordering and arithmetic operations on them, then show that the
operations descend to "games", defined via the equivalence relation p ≈ q ↔ p ≤ q ∧ q ≤ p
.
The surreal numbers will be built as a quotient of a subtype of pregames.
A pregame (SetTheory.PGame
below) is axiomatised via an inductive type, whose sole constructor
takes two types (thought of as indexing the possible moves for the players Left and Right), and a
pair of functions out of these types to SetTheory.PGame
(thought of as describing the resulting
game after making a move).
We may denote a game as $\{L | R\}$, where $L$ and $R$ stand for the collections of left and right moves. This notation is not currently used in Mathlib.
Combinatorial games themselves, as a quotient of pregames, are constructed in Game.lean
.
Conway induction #
By construction, the induction principle for pregames is exactly "Conway induction". That is, to
prove some predicate SetTheory.PGame → Prop
holds for all pregames, it suffices to prove
that for every pregame g
, if the predicate holds for every game resulting from making a move,
then it also holds for g
.
While it is often convenient to work "by induction" on pregames, in some situations this becomes
awkward, so we also define accessor functions SetTheory.PGame.LeftMoves
,
SetTheory.PGame.RightMoves
, SetTheory.PGame.moveLeft
and SetTheory.PGame.moveRight
.
There is a relation PGame.Subsequent p q
, saying that
p
can be reached by playing some non-empty sequence of moves starting from q
, an instance
WellFounded Subsequent
, and a local tactic pgame_wf_tac
which is helpful for discharging proof
obligations in inductive proofs relying on this relation.
Order properties #
Pregames have both a ≤
and a <
relation, satisfying the usual properties of a Preorder
. The
relation 0 < x
means that x
can always be won by Left, while 0 ≤ x
means that x
can be won
by Left as the second player.
It turns out to be quite convenient to define various relations on top of these. We define the "less
or fuzzy" relation x ⧏ y
as ¬ y ≤ x
, the equivalence relation x ≈ y
as x ≤ y ∧ y ≤ x
, and
the fuzzy relation x ‖ y
as x ⧏ y ∧ y ⧏ x
. If 0 ⧏ x
, then x
can be won by Left as the
first player. If x ≈ 0
, then x
can be won by the second player. If x ‖ 0
, then x
can be won
by the first player.
Statements like zero_le_lf
, zero_lf_le
, etc. unfold these definitions. The theorems le_def
and
lf_def
give a recursive characterisation of each relation in terms of themselves two moves later.
The theorems zero_le
, zero_lf
, etc. also take into account that 0
has no moves.
Later, games will be defined as the quotient by the ≈
relation; that is to say, the
Antisymmetrization
of SetTheory.PGame
.
Algebraic structures #
We next turn to defining the operations necessary to make games into a commutative additive 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, we introduce the notion of a Relabelling
of a
game, and show, for example, that there is a relabelling between x + (y + z)
and (x + y) + z
.
Future work #
- The theory of dominated and reversible positions, and unique normal form for short games.
- Analysis of basic domineering positions.
- Hex.
- Temperature.
- The development of surreal numbers, based on this development of combinatorial games, is still quite incomplete.
References #
The material here is all drawn from
- [Conway, On numbers and games][conway2001]
An interested reader may like to formalise some of the material from
- [Andreas Blass, A game semantics for linear logic][MR1167694]
- [André Joyal, Remarques sur la théorie des jeux à deux personnes][joyal1997]
Pre-game moves #
The type of pre-games, before we have quotiented
by equivalence (PGame.Setoid
). 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 pre-game is built
inductively from two families of pre-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.
- mk: (α β : Type u) → (α → SetTheory.PGame) → (β → SetTheory.PGame) → SetTheory.PGame
Instances For
The indexing type for allowable moves by Left.
Equations
- (SetTheory.PGame.mk l β a a_1).LeftMoves = l
Instances For
The indexing type for allowable moves by Right.
Equations
- (SetTheory.PGame.mk l β a a_1).RightMoves = β
Instances For
The new game after Left makes an allowed move.
Equations
- (SetTheory.PGame.mk l β a a_1).moveLeft = a
Instances For
The new game after Right makes an allowed move.
Equations
- (SetTheory.PGame.mk l β a a_1).moveRight = a_1
Instances For
Construct a pre-game from list of pre-games describing the available moves for Left and Right.
Equations
- One or more equations did not get rendered due to their size.
Instances For
Converts a number into a left move for ofLists
.
This is just an abbreviation for Equiv.ulift.symm
Equations
- SetTheory.PGame.toOfListsLeftMoves = Equiv.ulift.symm
Instances For
Converts a number into a right move for ofLists
.
This is just an abbreviation for Equiv.ulift.symm
Equations
- SetTheory.PGame.toOfListsRightMoves = Equiv.ulift.symm
Instances For
A variant of PGame.recOn
expressed in terms of PGame.moveLeft
and PGame.moveRight
.
Both this and PGame.recOn
describe Conway induction on games.
Equations
- x.moveRecOn IH = SetTheory.PGame.recOn x fun (yl yr : Type ?u.50) (yL : yl → SetTheory.PGame) (yR : yr → SetTheory.PGame) => IH (SetTheory.PGame.mk yl yr yL yR)
Instances For
IsOption x y
means that x
is either a left or right option for y
.
- moveLeft: ∀ {x : SetTheory.PGame} (i : x.LeftMoves), (x.moveLeft i).IsOption x
- moveRight: ∀ {x : SetTheory.PGame} (i : x.RightMoves), (x.moveRight i).IsOption x
Instances For
Subsequent x y
says that x
can be obtained by playing some nonempty sequence of moves from
y
. It is the transitive closure of IsOption
.
Instances For
Discharges proof obligations of the form ⊢ Subsequent ..
arising in termination proofs
of definitions using well-founded recursion on PGame
.
Equations
- SetTheory.PGame.tacticPgame_wf_tac = Lean.ParserDescr.node `SetTheory.PGame.tacticPgame_wf_tac 1024 (Lean.ParserDescr.nonReservedSymbol "pgame_wf_tac" false)
Instances For
Basic pre-games #
The pre-game Zero
is defined by 0 = { | }
.
Equations
- SetTheory.PGame.instZero = { zero := SetTheory.PGame.mk PEmpty.{?u.3 + 1} PEmpty.{?u.3 + 1} PEmpty.elim PEmpty.elim }
Equations
- SetTheory.PGame.instInhabited = { default := 0 }
The pre-game One
is defined by 1 = { 0 | }
.
Equations
- SetTheory.PGame.instOnePGame = { one := SetTheory.PGame.mk PUnit.{?u.4 + 1} PEmpty.{?u.4 + 1} (fun (x : PUnit.{?u.4 + 1} ) => 0) PEmpty.elim }
Pre-game order relations #
The less or equal relation on pre-games.
If 0 ≤ x
, then Left can win x
as the second player. x ≤ y
means that 0 ≤ y - x
.
See PGame.le_iff_sub_nonneg
.
Equations
- One or more equations did not get rendered due to their size.
The less or fuzzy relation on pre-games. x ⧏ y
is defined as ¬ y ≤ x
.
If 0 ⧏ x
, then Left can win x
as the first player. x ⧏ y
means that 0 ⧏ y - x
.
See PGame.lf_iff_sub_zero_lf
.
Instances For
The less or fuzzy relation on pre-games. x ⧏ y
is defined as ¬ y ≤ x
.
If 0 ⧏ x
, then Left can win x
as the first player. x ⧏ y
means that 0 ⧏ y - x
.
See PGame.lf_iff_sub_zero_lf
.
Equations
- SetTheory.PGame.«term_⧏_» = Lean.ParserDescr.trailingNode `SetTheory.PGame.term_⧏_ 50 50 (Lean.ParserDescr.binary `andthen (Lean.ParserDescr.symbol " ⧏ ") (Lean.ParserDescr.cat `term 51))
Instances For
Definition of x ≤ y
on pre-games, in terms of ⧏
.
The ordering here is chosen so that And.left
refer to moves by Left, and And.right
refer to
moves by Right.
Definition of x ≤ y
on pre-games built using the constructor.
Definition of x ⧏ y
on pre-games, in terms of ≤
.
The ordering here is chosen so that or.inl
refer to moves by Left, and or.inr
refer to
moves by Right.
Definition of x ⧏ y
on pre-games built using the constructor.
Alias of SetTheory.PGame.moveLeft_lf_of_le
.
Alias of SetTheory.PGame.lf_moveRight_of_le
.
Alias of SetTheory.PGame.lf_of_lt
.
Equations
Equations
- SetTheory.PGame.instTransLeLF = { trans := ⋯ }
Equations
- SetTheory.PGame.instTransLFLe = { trans := ⋯ }
Alias of SetTheory.PGame.lf_of_le_of_lf
.
Alias of SetTheory.PGame.lf_of_lf_of_le
.
Alias of SetTheory.PGame.lf_of_lt_of_lf
.
Alias of SetTheory.PGame.lf_of_lf_of_lt
.
This special case of PGame.le_of_forall_lf
is useful when dealing with surreals, where <
is
preferred over ⧏
.
The definition of x ≤ y
on pre-games, in terms of ≤
two moves later.
The definition of x ⧏ y
on pre-games, in terms of ⧏
two moves later.
The definition of 0 ≤ x
on pre-games, in terms of 0 ⧏
.
The definition of x ≤ 0
on pre-games, in terms of ⧏ 0
.
The definition of 0 ⧏ x
on pre-games, in terms of 0 ≤
.
The definition of x ⧏ 0
on pre-games, in terms of ≤ 0
.
The definition of 0 ≤ x
on pre-games, in terms of 0 ≤
two moves later.
The definition of x ≤ 0
on pre-games, in terms of ≤ 0
two moves later.
The definition of 0 ⧏ x
on pre-games, in terms of 0 ⧏
two moves later.
The definition of x ⧏ 0
on pre-games, in terms of ⧏ 0
two moves later.
Given a game won by the right player when they play second, provide a response to any move by left.
Equations
Instances For
Show that the response for right provided by rightResponse
preserves the right-player-wins
condition.
Given a game won by the left player when they play second, provide a response to any move by right.
Equations
Instances For
Show that the response for left provided by leftResponse
preserves the left-player-wins
condition.
A small family of pre-games is bounded above.
A small set of pre-games is bounded above.
A small family of pre-games is bounded below.
A small set of pre-games is bounded below.
The equivalence relation on pre-games. Two pre-games x
, y
are equivalent if x ≤ y
and
y ≤ x
.
If x ≈ 0
, then the second player can always win x
.
Instances For
Equations
- SetTheory.PGame.setoid = { r := SetTheory.PGame.Equiv, iseqv := SetTheory.PGame.setoid.proof_1 }
Equations
- SetTheory.PGame.instTransLeEquiv = { trans := ⋯ }
Equations
- SetTheory.PGame.instTransEquivLe = { trans := ⋯ }
Equations
- SetTheory.PGame.instTransLFEquiv = { trans := ⋯ }
Equations
- SetTheory.PGame.instTransEquivLF = { trans := ⋯ }
Equations
- SetTheory.PGame.instTransLtEquiv = { trans := ⋯ }
Equations
- SetTheory.PGame.instTransEquivLt = { trans := ⋯ }
Alias of SetTheory.PGame.Equiv.of_equiv
.
The fuzzy, confused, or incomparable relation on pre-games.
If x ‖ 0
, then the first player can always win x
.
Instances For
The fuzzy, confused, or incomparable relation on pre-games.
If x ‖ 0
, then the first player can always win x
.
Equations
- SetTheory.PGame.«term_‖_» = Lean.ParserDescr.trailingNode `SetTheory.PGame.term_‖_ 50 50 (Lean.ParserDescr.binary `andthen (Lean.ParserDescr.symbol " ‖ ") (Lean.ParserDescr.cat `term 51))
Instances For
Equations
Equations
Alias of SetTheory.PGame.lf_of_fuzzy
.
Exactly one of the following is true (although we don't prove this here).
Relabellings #
Relabelling x y
says that x
and y
are really the same game, just dressed up differently.
Specifically, there is a bijection between the moves for Left in x
and in y
, and similarly
for Right, and under these bijections we inductively have Relabelling
s for the consequent games.
- mk: {x y : SetTheory.PGame} → (L : x.LeftMoves ≃ y.LeftMoves) → (R : x.RightMoves ≃ y.RightMoves) → ((i : x.LeftMoves) → (x.moveLeft i).Relabelling (y.moveLeft (L i))) → ((j : x.RightMoves) → (x.moveRight j).Relabelling (y.moveRight (R j))) → x.Relabelling y
Instances For
Relabelling x y
says that x
and y
are really the same game, just dressed up differently.
Specifically, there is a bijection between the moves for Left in x
and in y
, and similarly
for Right, and under these bijections we inductively have Relabelling
s for the consequent games.
Equations
- SetTheory.PGame.«term_≡r_» = Lean.ParserDescr.trailingNode `SetTheory.PGame.term_≡r_ 50 50 (Lean.ParserDescr.binary `andthen (Lean.ParserDescr.symbol " ≡r ") (Lean.ParserDescr.cat `term 51))
Instances For
A constructor for relabellings swapping the equivalences.
Equations
- SetTheory.PGame.Relabelling.mk' L R hL hR = SetTheory.PGame.Relabelling.mk L.symm R.symm (fun (i : x.LeftMoves) => ⋯.mp (hL (L.symm i))) fun (j : x.RightMoves) => ⋯.mp (hR (R.symm j))
Instances For
The equivalence between left moves of x
and y
given by the relabelling.
Equations
- (SetTheory.PGame.Relabelling.mk L R a a_1).leftMovesEquiv = L
Instances For
The equivalence between right moves of x
and y
given by the relabelling.
Equations
- (SetTheory.PGame.Relabelling.mk L R a a_1).rightMovesEquiv = R
Instances For
A left move of x
is a relabelling of a left move of y
.
Equations
- (SetTheory.PGame.Relabelling.mk L R a a_1).moveLeft = a
Instances For
A left move of y
is a relabelling of a left move of x
.
Equations
- (SetTheory.PGame.Relabelling.mk L R hL hR).moveLeftSymm x = id (⋯.mp (hL (L.symm x)))
Instances For
A right move of x
is a relabelling of a right move of y
.
Equations
- (SetTheory.PGame.Relabelling.mk L R a a_1).moveRight = a_1
Instances For
A right move of y
is a relabelling of a right move of x
.
Equations
- (SetTheory.PGame.Relabelling.mk L R hL hR).moveRightSymm x = id (⋯.mp (hR (R.symm x)))
Instances For
The identity relabelling.
Equations
- One or more equations did not get rendered due to their size.
Instances For
Equations
- SetTheory.PGame.Relabelling.instInhabited x = { default := SetTheory.PGame.Relabelling.refl x }
Flip a relabelling.
Equations
- (SetTheory.PGame.Relabelling.mk L R hL hR).symm = SetTheory.PGame.Relabelling.mk' L R (fun (i : x✝.LeftMoves) => (hL i).symm) fun (j : x✝.RightMoves) => (hR j).symm
Instances For
A relabelling lets us prove equivalence of games.
Transitivity of relabelling.
Equations
- One or more equations did not get rendered due to their size.
Instances For
Any game without left or right moves is a relabelling of 0.
Equations
- One or more equations did not get rendered due to their size.
Instances For
Equations
- SetTheory.PGame.instCoeRelabellingEquiv = { coe := ⋯ }
Replace the types indexing the next moves for Left and Right by equivalent types.
Equations
- SetTheory.PGame.relabel el er = SetTheory.PGame.mk xl' xr' (x.moveLeft ∘ ⇑el) (x.moveRight ∘ ⇑er)
Instances For
The game obtained by relabelling the next moves is a relabelling of the original game.
Equations
- One or more equations did not get rendered due to their size.
Instances For
Negation #
The negation of {L | R}
is {-R | -L}
.
Equations
- (SetTheory.PGame.mk l β a a_1).neg = SetTheory.PGame.mk β l (fun (i : β) => (a_1 i).neg) fun (i : l) => (a i).neg
Instances For
Equations
- SetTheory.PGame.instNeg = { neg := SetTheory.PGame.neg }
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
- SetTheory.PGame.toLeftMovesNeg = Equiv.cast ⋯
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
- SetTheory.PGame.toRightMovesNeg = Equiv.cast ⋯
Instances For
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 }
Equations
- ⋯ = ⋯
Equations
- ⋯ = ⋯
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
.
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
- SetTheory.PGame.toLeftMovesAdd = Equiv.cast ⋯
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
- SetTheory.PGame.toRightMovesAdd = Equiv.cast ⋯
Instances For
Equations
- ⋯ = ⋯
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 }
Alias of SetTheory.PGame.sub_zero_eq_add_zero
.
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
- h₁.subCongr h₂ = h₁.addCongr h₂.negCongr
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
Equations
Equations
Inserting an option #
The pregame constructed by inserting x'
as a new left option into x.
Equations
- (SetTheory.PGame.mk l β a a_1).insertLeft x' = SetTheory.PGame.mk (l ⊕ PUnit.{?u.3 + 1} ) β (Sum.elim a fun (x : PUnit.{?u.3 + 1} ) => x') a_1
Instances For
A new left option cannot hurt Left.
Adding a gift horse left option does not change the value of x
. A gift horse left option is
a game x'
with x' ⧏ x
. It is called "gift horse" because it seems like Left has gotten the
"gift" of a new option, but actually the value of the game did not change.
The pregame constructed by inserting x'
as a new right option into x.
Equations
- (SetTheory.PGame.mk l β a a_1).insertRight x' = SetTheory.PGame.mk l (β ⊕ PUnit.{?u.143 + 1} ) a (Sum.elim a_1 fun (x : PUnit.{?u.143 + 1} ) => x')
Instances For
A new right option cannot hurt Right.
Adding a gift horse right option does not change the value of x
. A gift horse right option is
a game x'
with x ⧏ x'
. It is called "gift horse" because it seems like Right has gotten the
"gift" of a new option, but actually the value of the game did not change.
Inserting on the left and right commutes.
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