@@ -17,6 +17,16 @@ This file contains the definition for nim for any ordinal `O`. In the game of `n
17
17
may move to `nim O₂` for any `O₂ < O₁`.
18
18
We also define a Grundy value for an impartial game `G` and prove the Sprague-Grundy theorem, that
19
19
`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
+
20
30
-/
21
31
22
32
open pgame
@@ -41,12 +51,11 @@ def nim : ordinal → pgame
41
51
λ O₂, have hwf : (ordinal.typein O₁.out.r O₂) < O₁,
42
52
from begin nth_rewrite_rhs 0 ←ordinal.type_out' O₁, exact ordinal.typein_lt_type _ _ end ,
43
53
nim (ordinal.typein O₁.out.r O₂)⟩
44
- using_well_founded {dec_tac := tactic.assumption}
54
+ using_well_founded { dec_tac := tactic.assumption }
45
55
46
56
namespace nim
47
57
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.α
50
59
(λ O₂, nim (ordinal.typein O.out.r O₂))
51
60
(λ O₂, nim (ordinal.typein O.out.r O₂)) :=
52
61
by rw nim
@@ -67,107 +76,68 @@ begin
67
76
split,
68
77
{ intro i,
69
78
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 ⟩ },
72
80
{ intro j,
73
81
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 ⟩ } },
76
83
{ rw pgame.le_def,
77
84
split,
78
85
{ intro i,
79
86
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 ⟩ },
82
88
{ intro j,
83
89
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 ⟩ } },
86
91
split,
87
92
{ intro i,
88
93
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) },
91
95
{ intro j,
92
96
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) }
95
98
end
96
- using_well_founded {dec_tac := tactic.assumption}
99
+ using_well_founded { dec_tac := tactic.assumption }
97
100
98
- lemma nim_zero_first_loses : (nim (0 : ordinal)).first_loses :=
101
+ lemma nim_zero_first_loses : (nim (0 : ordinal)).first_loses :=
99
102
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,
108
107
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 }
113
110
end
114
111
115
112
lemma nim_non_zero_first_wins (O : ordinal) (hO : O ≠ 0 ) : (nim O).first_wins :=
116
113
begin
117
- rw nim_def,
114
+ rw [impartial_first_wins_symm (nim_impartial _), nim_def, lt_def_le] ,
118
115
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 ⟩
129
117
end
130
118
131
119
lemma nim_sum_first_loses_iff_eq (O₁ O₂ : ordinal) : (nim O₁ + nim O₂).first_loses ↔ O₁ = O₂ :=
132
120
begin
133
121
split,
134
122
{ 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₁) }
153
136
end
154
137
155
138
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]
171
141
172
142
lemma nim_equiv_iff_eq (O₁ O₂ : ordinal) : nim O₁ ≈ nim O₂ ↔ O₁ = O₂ :=
173
143
⟨λ h, (nim_sum_first_loses_iff_eq _ _).1 $
@@ -185,26 +155,19 @@ lemma nonmoves_nonempty {α : Type u} (M : α → ordinal.{u}) : ∃ O : ordinal
185
155
begin
186
156
classical,
187
157
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,
190
159
191
160
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₂) },
205
168
206
169
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⟩,
208
171
209
172
cases hlt,
210
173
contradiction
@@ -216,7 +179,7 @@ noncomputable def Grundy_value : Π {G : pgame.{u}}, G.impartial → ordinal.{u}
216
179
| G :=
217
180
λ hG, ordinal.omin (nonmoves (λ i, Grundy_value $ impartial_move_left_impartial hG i))
218
181
(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 }
220
183
221
184
lemma Grundy_value_def {G : pgame} (hG : G.impartial) :
222
185
Grundy_value hG = ordinal.omin (nonmoves (λ i, (Grundy_value $ impartial_move_left_impartial hG i)))
@@ -239,23 +202,20 @@ begin
239
202
equiv_rw left_moves_add G (nim $ Grundy_value hG) at i,
240
203
cases i with i₁ i₂,
241
204
{ 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 _)),
244
207
rw nim_sum_first_wins_iff_neq,
245
208
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,
248
210
have h := ordinal.omin_mem
249
211
(nonmoves (λ (i : G.left_moves), Grundy_value (impartial_move_left_impartial hG i)))
250
212
(nonmoves_nonempty _),
251
213
rw heq at h,
252
214
have hcontra : ∃ (i' : G.left_moves),
253
215
(λ (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⟩,
256
217
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
259
219
(impartial_add hG $ impartial_move_left_impartial (nim_impartial _) _)],
260
220
revert i₂,
261
221
rw nim_def,
@@ -264,35 +224,28 @@ begin
264
224
have h' : ∃ i : G.left_moves, (Grundy_value $ impartial_move_left_impartial hG i) =
265
225
ordinal.typein (quotient.out $ Grundy_value hG).r i₂,
266
226
{ 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 _ _,
269
228
rw ordinal.type_out at hlt,
270
229
revert i₂ hlt,
271
230
rw Grundy_value_def,
272
231
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₂ ∉
276
234
(nonmoves (λ (i : G.left_moves), Grundy_value (impartial_move_left_impartial hG i))),
277
235
{ intro hin,
278
236
have hge := ordinal.omin_le hin,
279
237
have hcontra := (le_not_le_of_lt hlt).2 ,
280
238
contradiction },
281
- unfold nonmoves at hnotin,
282
- simpa using hnotin },
239
+ simpa [nonmoves] using hnotin },
283
240
284
241
cases h' with i hi,
285
242
use (left_moves_add _ _).symm (sum.inl i),
286
243
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 _) }
294
247
end
295
- using_well_founded {dec_tac := pgame_wf_tac}
248
+ using_well_founded { dec_tac := pgame_wf_tac }
296
249
297
250
lemma equiv_nim_iff_Grundy_value_eq {G : pgame.{u}} (hG : impartial G) (O : ordinal) :
298
251
G ≈ nim O ↔ Grundy_value hG = O :=
@@ -303,3 +256,4 @@ lemma Grundy_value_nim (O : ordinal.{u}) : Grundy_value (nim_impartial O) = O :=
303
256
by rw ←equiv_nim_iff_Grundy_value_eq
304
257
305
258
end nim
259
+ #lint
0 commit comments