Skip to content
This repository was archived by the owner on Jul 24, 2024. It is now read-only.

Commit b03ce61

Browse files
committed
chore(set_theory/game): various cleanup and code golf (#3939)
1 parent 878c44f commit b03ce61

File tree

3 files changed

+90
-114
lines changed

3 files changed

+90
-114
lines changed

src/set_theory/game/impartial.lean

Lines changed: 12 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -117,6 +117,18 @@ begin
117117
{ right, assumption }
118118
end
119119

120+
lemma impartial_not_first_wins {G : pgame} (hG : G.impartial) : ¬G.first_wins ↔ G.first_loses :=
121+
begin
122+
cases impartial_winner_cases hG,
123+
{ have := not_first_wins_of_first_loses h,
124+
tauto },
125+
{ have := not_first_loses_of_first_wins h,
126+
tauto }
127+
end
128+
129+
lemma impartial_not_first_loses {G : pgame} (hG : G.impartial) : ¬G.first_loses ↔ G.first_wins :=
130+
iff.symm $ iff_not_comm.1 $ iff.symm $ impartial_not_first_wins hG
131+
120132
lemma impartial_add_self {G : pgame} (hG : G.impartial) : (G + G).first_loses :=
121133
first_loses_is_zero.2 $ equiv_trans (add_congr (impartial_neg_equiv_self hG) G.equiv_refl)
122134
add_left_neg_equiv

src/set_theory/game/nim.lean

Lines changed: 67 additions & 113 deletions
Original file line numberDiff line numberDiff line change
@@ -17,6 +17,16 @@ This file contains the definition for nim for any ordinal `O`. In the game of `n
1717
may move to `nim O₂` for any `O₂ < O₁`.
1818
We also define a Grundy value for an impartial game `G` and prove the Sprague-Grundy theorem, that
1919
`G` is equivalent to `nim (Grundy_value G)`.
20+
21+
## Implementation details
22+
23+
The pen-and-paper definition of nim defines the possible moves of `nim O` to be `{ O' | O' < O}`.
24+
However, this definition does not work for us because it would make the type of nim
25+
`ordinal.{u} → pgame.{u + 1}`, which would make it impossible for us to state the Sprague-Grundy
26+
theorem, since that requires the type of `nim` to be `ordinal.{u} → pgame.{u}`. For this reason, we
27+
instead use `O.out.α` for the possible moves, which makes proofs significantly more messy and
28+
tedious, but avoids the universe bump.
29+
2030
-/
2131

2232
open pgame
@@ -41,12 +51,11 @@ def nim : ordinal → pgame
4151
λ O₂, have hwf : (ordinal.typein O₁.out.r O₂) < O₁,
4252
from begin nth_rewrite_rhs 0 ←ordinal.type_out' O₁, exact ordinal.typein_lt_type _ _ end,
4353
nim (ordinal.typein O₁.out.r O₂)⟩
44-
using_well_founded {dec_tac := tactic.assumption}
54+
using_well_founded { dec_tac := tactic.assumption }
4555

4656
namespace nim
4757

48-
lemma nim_def (O : ordinal) : nim O = pgame.mk
49-
O.out.α O.out.α
58+
lemma nim_def (O : ordinal) : nim O = pgame.mk O.out.α O.out.α
5059
(λ O₂, nim (ordinal.typein O.out.r O₂))
5160
(λ O₂, nim (ordinal.typein O.out.r O₂)) :=
5261
by rw nim
@@ -67,107 +76,68 @@ begin
6776
split,
6877
{ intro i,
6978
let hwf : (ordinal.typein O.out.r i) < O := nim_wf_lemma i,
70-
left,
71-
exact ⟨ i, (impartial_neg_equiv_self $ nim_impartial $ ordinal.typein O.out.r i).1 ⟩ },
79+
exact or.inl ⟨i, (impartial_neg_equiv_self $ nim_impartial $ ordinal.typein O.out.r i).1⟩ },
7280
{ intro j,
7381
let hwf : (ordinal.typein O.out.r j) < O := nim_wf_lemma j,
74-
right,
75-
exact ⟨ j, (impartial_neg_equiv_self $ nim_impartial $ ordinal.typein O.out.r j).1 ⟩ } },
82+
exact or.inr ⟨j, (impartial_neg_equiv_self $ nim_impartial $ ordinal.typein O.out.r j).1⟩ } },
7683
{ rw pgame.le_def,
7784
split,
7885
{ intro i,
7986
let hwf : (ordinal.typein O.out.r i) < O := nim_wf_lemma i,
80-
left,
81-
exact ⟨ i, (impartial_neg_equiv_self $ nim_impartial $ ordinal.typein O.out.r i).2 ⟩ },
87+
exact or.inl ⟨i, (impartial_neg_equiv_self $ nim_impartial $ ordinal.typein O.out.r i).2⟩ },
8288
{ intro j,
8389
let hwf : (ordinal.typein O.out.r j) < O := nim_wf_lemma j,
84-
right,
85-
exact ⟨ j, (impartial_neg_equiv_self $ nim_impartial $ ordinal.typein O.out.r j).2 ⟩ } },
90+
exact or.inr ⟨j, (impartial_neg_equiv_self $ nim_impartial $ ordinal.typein O.out.r j).2⟩ } },
8691
split,
8792
{ intro i,
8893
let hwf : (ordinal.typein O.out.r i) < O := nim_wf_lemma i,
89-
rw move_left_mk,
90-
exact nim_impartial (ordinal.typein O.out.r i) },
94+
simpa using nim_impartial (ordinal.typein O.out.r i) },
9195
{ intro j,
9296
let hwf : (ordinal.typein O.out.r j) < O := nim_wf_lemma j,
93-
rw move_right_mk,
94-
exact nim_impartial (ordinal.typein O.out.r j) }
97+
simpa using nim_impartial (ordinal.typein O.out.r j) }
9598
end
96-
using_well_founded {dec_tac := tactic.assumption}
99+
using_well_founded { dec_tac := tactic.assumption }
97100

98-
lemma nim_zero_first_loses : (nim (0:ordinal)).first_loses :=
101+
lemma nim_zero_first_loses : (nim (0 : ordinal)).first_loses :=
99102
begin
100-
rw nim_def,
101-
split;
102-
rw le_def_lt;
103-
split;
104-
intro i;
105-
try {rw left_moves_mk at i};
106-
try {rw right_moves_mk at i};
107-
try { have h := ordinal.typein_lt_type (quotient.out (0:ordinal)).r i,
103+
rw [impartial_first_loses_symm (nim_impartial _), nim_def, le_def_lt],
104+
split,
105+
{ rintro (i : (0 : ordinal).out.α),
106+
have h := ordinal.typein_lt_type _ i,
108107
rw ordinal.type_out at h,
109-
have hcontra := ordinal.zero_le (ordinal.typein (quotient.out (0:ordinal)).r i),
110-
have h := not_le_of_lt h,
111-
contradiction };
112-
try { exact pempty.elim i }
108+
exact false.elim (not_le_of_lt h (ordinal.zero_le (ordinal.typein _ i))) },
109+
{ tidy }
113110
end
114111

115112
lemma nim_non_zero_first_wins (O : ordinal) (hO : O ≠ 0) : (nim O).first_wins :=
116113
begin
117-
rw nim_def,
114+
rw [impartial_first_wins_symm (nim_impartial _), nim_def, lt_def_le],
118115
rw ←ordinal.pos_iff_ne_zero at hO,
119-
split;
120-
rw lt_def_le,
121-
{ left,
122-
use (ordinal.principal_seg_out hO).top,
123-
rw [move_left_mk, ordinal.typein_top, ordinal.type_out],
124-
exact nim_zero_first_loses.2 },
125-
{ right,
126-
use (ordinal.principal_seg_out hO).top,
127-
rw [move_right_mk, ordinal.typein_top, ordinal.type_out],
128-
exact nim_zero_first_loses.1 }
116+
exact or.inr ⟨(ordinal.principal_seg_out hO).top, by simpa using nim_zero_first_loses.1
129117
end
130118

131119
lemma nim_sum_first_loses_iff_eq (O₁ O₂ : ordinal) : (nim O₁ + nim O₂).first_loses ↔ O₁ = O₂ :=
132120
begin
133121
split,
134122
{ contrapose,
135-
intros hneq hp,
136-
wlog h : O₁ ≤ O₂ using [O₁ O₂, O₂ O₁],
137-
exact ordinal.le_total O₁ O₂,
138-
{ have h : O₁ < O₂ := lt_of_le_of_ne h hneq,
139-
rw ←(no_good_left_moves_iff_first_loses $ impartial_add (nim_impartial O₁) (nim_impartial O₂))
140-
at hp,
141-
equiv_rw left_moves_add (nim O₁) (nim O₂) at hp,
142-
rw nim_def O₂ at hp,
143-
specialize hp (sum.inr (ordinal.principal_seg_out h).top),
144-
rw [id, add_move_left_inr, move_left_mk, ordinal.typein_top, ordinal.type_out] at hp,
145-
cases hp,
146-
have hcontra := (impartial_add_self $ nim_impartial O₁).1,
147-
rw ←pgame.not_lt at hcontra,
148-
contradiction },
149-
exact this (λ h, hneq h.symm) (first_loses_of_equiv add_comm_equiv hp) },
150-
{ intro h,
151-
rw h,
152-
exact impartial_add_self (nim_impartial O₂) }
123+
intro h,
124+
rw [impartial_not_first_loses (impartial_add (nim_impartial _) (nim_impartial _))],
125+
wlog h' : O₁ ≤ O₂ using [O₁ O₂, O₂ O₁],
126+
{ exact ordinal.le_total O₁ O₂ },
127+
{ have h : O₁ < O₂ := lt_of_le_of_ne h' h,
128+
rw [impartial_first_wins_symm' (impartial_add (nim_impartial _) (nim_impartial _)),
129+
lt_def_le, nim_def O₂],
130+
refine or.inl ⟨(left_moves_add (nim O₁) _).symm (sum.inr _), _⟩,
131+
{ exact (ordinal.principal_seg_out h).top },
132+
{ simpa using (impartial_add_self (nim_impartial _)).2 } },
133+
{ exact first_wins_of_equiv add_comm_equiv (this (ne.symm h)) } },
134+
{ rintro rfl,
135+
exact impartial_add_self (nim_impartial O₁) }
153136
end
154137

155138
lemma nim_sum_first_wins_iff_neq (O₁ O₂ : ordinal) : (nim O₁ + nim O₂).first_wins ↔ O₁ ≠ O₂ :=
156-
begin
157-
split,
158-
{ intros hn heq,
159-
cases hn,
160-
rw ←nim_sum_first_loses_iff_eq at heq,
161-
cases heq with h,
162-
rw ←pgame.not_lt at h,
163-
contradiction },
164-
{ contrapose,
165-
intro hnp,
166-
rw [not_not, ←nim_sum_first_loses_iff_eq],
167-
cases impartial_winner_cases (impartial_add (nim_impartial O₁) (nim_impartial O₂)),
168-
assumption,
169-
contradiction }
170-
end
139+
by rw [iff_not_comm, impartial_not_first_wins (impartial_add (nim_impartial _) (nim_impartial _)),
140+
nim_sum_first_loses_iff_eq]
171141

172142
lemma nim_equiv_iff_eq (O₁ O₂ : ordinal) : nim O₁ ≈ nim O₂ ↔ O₁ = O₂ :=
173143
⟨λ h, (nim_sum_first_loses_iff_eq _ _).1 $
@@ -185,26 +155,19 @@ lemma nonmoves_nonempty {α : Type u} (M : α → ordinal.{u}) : ∃ O : ordinal
185155
begin
186156
classical,
187157
by_contra h,
188-
rw nonmoves at h,
189-
simp only [not_exists, not_forall, set.mem_set_of_eq, not_not] at h,
158+
simp only [nonmoves, not_exists, not_forall, set.mem_set_of_eq, not_not] at h,
190159

191160
have hle : cardinal.univ.{u (u+1)} ≤ cardinal.lift.{u (u+1)} (cardinal.mk α),
192-
{ split,
193-
fconstructor,
194-
{ rintro ⟨ O ⟩,
195-
exact ⟨ (classical.indefinite_description _ $ h O).val ⟩ },
196-
{ rintros ⟨ O₁ ⟩ ⟨ O₂ ⟩ heq,
197-
ext,
198-
transitivity,
199-
symmetry,
200-
exact (classical.indefinite_description _ (h O₁)).prop,
201-
injection heq with heq,
202-
rw subtype.val_eq_coe at heq,
203-
rw heq,
204-
exact (classical.indefinite_description _ (h O₂)).prop } },
161+
{ refine ⟨⟨λ ⟨O⟩, ⟨classical.some (h O)⟩, _⟩⟩,
162+
rintros ⟨O₁⟩ ⟨O₂⟩ heq,
163+
ext,
164+
refine eq.trans (classical.some_spec (h O₁)).symm _,
165+
injection heq with heq,
166+
rw heq,
167+
exact classical.some_spec (h O₂) },
205168

206169
have hlt : cardinal.lift.{u (u+1)} (cardinal.mk α) < cardinal.univ.{u (u+1)} :=
207-
cardinal.lt_univ.2 cardinal.mk α, rfl ⟩,
170+
cardinal.lt_univ.2 ⟨cardinal.mk α, rfl⟩,
208171

209172
cases hlt,
210173
contradiction
@@ -216,7 +179,7 @@ noncomputable def Grundy_value : Π {G : pgame.{u}}, G.impartial → ordinal.{u}
216179
| G :=
217180
λ hG, ordinal.omin (nonmoves (λ i, Grundy_value $ impartial_move_left_impartial hG i))
218181
(nonmoves_nonempty (λ i, Grundy_value (impartial_move_left_impartial hG i)))
219-
using_well_founded {dec_tac := pgame_wf_tac}
182+
using_well_founded { dec_tac := pgame_wf_tac }
220183

221184
lemma Grundy_value_def {G : pgame} (hG : G.impartial) :
222185
Grundy_value hG = ordinal.omin (nonmoves (λ i, (Grundy_value $ impartial_move_left_impartial hG i)))
@@ -239,23 +202,20 @@ begin
239202
equiv_rw left_moves_add G (nim $ Grundy_value hG) at i,
240203
cases i with i₁ i₂,
241204
{ rw add_move_left_inl,
242-
apply first_wins_of_equiv,
243-
exact (add_congr (Sprague_Grundy $ impartial_move_left_impartial hG i₁).symm (equiv_refl _)),
205+
apply first_wins_of_equiv
206+
(add_congr (Sprague_Grundy $ impartial_move_left_impartial hG i₁).symm (equiv_refl _)),
244207
rw nim_sum_first_wins_iff_neq,
245208
intro heq,
246-
have heq := symm heq,
247-
rw Grundy_value_def hG at heq,
209+
rw [eq_comm, Grundy_value_def hG] at heq,
248210
have h := ordinal.omin_mem
249211
(nonmoves (λ (i : G.left_moves), Grundy_value (impartial_move_left_impartial hG i)))
250212
(nonmoves_nonempty _),
251213
rw heq at h,
252214
have hcontra : ∃ (i' : G.left_moves),
253215
(λ (i'' : G.left_moves), Grundy_value $ impartial_move_left_impartial hG i'') i' =
254-
Grundy_value (impartial_move_left_impartial hG i₁) :=
255-
⟨ i₁, rfl ⟩,
216+
Grundy_value (impartial_move_left_impartial hG i₁) := ⟨i₁, rfl⟩,
256217
contradiction },
257-
{ rw [add_move_left_inr,
258-
←good_left_move_iff_first_wins
218+
{ rw [add_move_left_inr, ←good_left_move_iff_first_wins
259219
(impartial_add hG $ impartial_move_left_impartial (nim_impartial _) _)],
260220
revert i₂,
261221
rw nim_def,
@@ -264,35 +224,28 @@ begin
264224
have h' : ∃ i : G.left_moves, (Grundy_value $ impartial_move_left_impartial hG i) =
265225
ordinal.typein (quotient.out $ Grundy_value hG).r i₂,
266226
{ have hlt : ordinal.typein (quotient.out $ Grundy_value hG).r i₂ <
267-
ordinal.type (quotient.out $ Grundy_value hG).r :=
268-
ordinal.typein_lt_type _ _,
227+
ordinal.type (quotient.out $ Grundy_value hG).r := ordinal.typein_lt_type _ _,
269228
rw ordinal.type_out at hlt,
270229
revert i₂ hlt,
271230
rw Grundy_value_def,
272231
intros i₂ hlt,
273-
have hnotin :
274-
ordinal.typein (quotient.out (ordinal.omin
275-
(nonmoves (λ i, Grundy_value (impartial_move_left_impartial hG i))) _)).r i₂ ∉
232+
have hnotin : ordinal.typein (quotient.out (ordinal.omin
233+
(nonmoves (λ i, Grundy_value (impartial_move_left_impartial hG i))) _)).r i₂ ∉
276234
(nonmoves (λ (i : G.left_moves), Grundy_value (impartial_move_left_impartial hG i))),
277235
{ intro hin,
278236
have hge := ordinal.omin_le hin,
279237
have hcontra := (le_not_le_of_lt hlt).2,
280238
contradiction },
281-
unfold nonmoves at hnotin,
282-
simpa using hnotin },
239+
simpa [nonmoves] using hnotin },
283240

284241
cases h' with i hi,
285242
use (left_moves_add _ _).symm (sum.inl i),
286243
rw [add_move_left_inl, move_left_mk],
287-
apply first_loses_of_equiv,
288-
apply add_congr,
289-
symmetry,
290-
exact Sprague_Grundy (impartial_move_left_impartial hG i),
291-
refl,
292-
rw hi,
293-
exact impartial_add_self (nim_impartial _) }
244+
apply first_loses_of_equiv
245+
(add_congr (equiv_symm (Sprague_Grundy (impartial_move_left_impartial hG i))) (equiv_refl _)),
246+
simpa only [hi] using impartial_add_self (nim_impartial _) }
294247
end
295-
using_well_founded {dec_tac := pgame_wf_tac}
248+
using_well_founded { dec_tac := pgame_wf_tac }
296249

297250
lemma equiv_nim_iff_Grundy_value_eq {G : pgame.{u}} (hG : impartial G) (O : ordinal) :
298251
G ≈ nim O ↔ Grundy_value hG = O :=
@@ -303,3 +256,4 @@ lemma Grundy_value_nim (O : ordinal.{u}) : Grundy_value (nim_impartial O) = O :=
303256
by rw ←equiv_nim_iff_Grundy_value_eq
304257

305258
end nim
259+
#lint

src/set_theory/game/winner.lean

Lines changed: 11 additions & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -9,7 +9,7 @@ import set_theory.pgame
99
# Basic definitions about who has a winning stratergy
1010
1111
We define `G.first_loses`, `G.first_wins`, `G.left_wins` and `G.right_wins` for a pgame `G`, which
12-
means the second, first, left and right players have a winning stratergy respectively.
12+
means the second, first, left and right players have a winning strategy respectively.
1313
These are defined by inequalities which can be unfolded with `pgame.lt_def` and `pgame.le_def`.
1414
-/
1515

@@ -79,4 +79,14 @@ lemma left_wins_of_equiv_iff {G H : pgame} (h : G ≈ H) : G.left_wins ↔ H.lef
7979
lemma right_wins_of_equiv_iff {G H : pgame} (h : G ≈ H) : G.right_wins ↔ H.right_wins :=
8080
⟨ right_wins_of_equiv h, right_wins_of_equiv h.symm ⟩
8181

82+
lemma not_first_wins_of_first_loses {G : pgame} : G.first_loses → ¬G.first_wins :=
83+
begin
84+
rw first_loses_is_zero,
85+
rintros h ⟨h₀, -⟩,
86+
exact lt_irrefl 0 (lt_of_lt_of_equiv h₀ h)
87+
end
88+
89+
lemma not_first_loses_of_first_wins {G : pgame} : G.first_wins → ¬G.first_loses :=
90+
imp_not_comm.1 $ not_first_wins_of_first_loses
91+
8292
end pgame

0 commit comments

Comments
 (0)