@@ -57,6 +57,35 @@ lemma nim_def (O : ordinal) : nim O = pgame.mk O.out.α O.out.α
57
57
(λ O₂, nim (ordinal.typein (<) O₂)) :=
58
58
by { rw nim, refl }
59
59
60
+ instance : is_empty (left_moves (nim 0 )) :=
61
+ by { rw nim_def, exact α.is_empty }
62
+
63
+ instance : is_empty (right_moves (nim 0 )) :=
64
+ by { rw nim_def, exact α.is_empty }
65
+
66
+ noncomputable instance : unique (left_moves (nim 1 )) :=
67
+ by { rw nim_def, exact α.unique }
68
+
69
+ noncomputable instance : unique (right_moves (nim 1 )) :=
70
+ by { rw nim_def, exact α.unique }
71
+
72
+ /-- `0` has exactly the same moves as `nim 0`. -/
73
+ def nim_zero_relabelling : relabelling 0 (nim 0 ) :=
74
+ (relabelling.is_empty _).symm
75
+
76
+ @[simp] theorem nim_zero_equiv : 0 ≈ nim 0 := nim_zero_relabelling.equiv
77
+
78
+ /-- `nim 1` has exactly the same moves as `star`. -/
79
+ noncomputable def nim_one_relabelling : relabelling star (nim 1 ) :=
80
+ begin
81
+ rw nim_def,
82
+ refine ⟨_, _, λ i, _, λ j, _⟩,
83
+ any_goals { dsimp, apply equiv_of_unique_of_unique },
84
+ all_goals { simp, exact nim_zero_relabelling }
85
+ end
86
+
87
+ @[simp] theorem nim_one_equiv : star ≈ nim 1 := nim_one_relabelling.equiv
88
+
60
89
@[simp] lemma nim_birthday (O : ordinal) : (nim O).birthday = O :=
61
90
begin
62
91
induction O using ordinal.induction with O IH,
@@ -100,7 +129,7 @@ lemma non_zero_first_wins {O : ordinal} (hO : O ≠ 0) : (nim O).first_wins :=
100
129
begin
101
130
rw [impartial.first_wins_symm, nim_def, lt_def_le],
102
131
rw ←ordinal.pos_iff_ne_zero at hO,
103
- exact or.inr ⟨(ordinal.principal_seg_out hO).top, by simpa using zero_first_loses. 1 ⟩
132
+ exact or.inr ⟨(ordinal.principal_seg_out hO).top, by simp ⟩
104
133
end
105
134
106
135
@[simp] lemma sum_first_loses_iff_eq (O₁ O₂ : ordinal) : (nim O₁ + nim O₂).first_loses ↔ O₁ = O₂ :=
@@ -194,12 +223,15 @@ by simp
194
223
grundy_value G = grundy_value H ↔ G ≈ H :=
195
224
(grundy_value_eq_iff_equiv_nim _ _).trans (equiv_congr_left.1 (equiv_nim_grundy_value H) _).symm
196
225
197
- @[simp] lemma grundy_value_zero : grundy_value 0 = 0 :=
198
- by rw [(grundy_value_eq_iff_equiv 0 (nim 0 )). 2 (equiv_symm nim.zero_first_loses), nim.grundy_value]
226
+ lemma grundy_value_zero : grundy_value 0 = 0 :=
227
+ by simp
199
228
200
229
@[simp] lemma grundy_value_iff_equiv_zero (G : pgame) [G.impartial] : grundy_value G = 0 ↔ G ≈ 0 :=
201
230
by rw [←grundy_value_eq_iff_equiv, grundy_value_zero]
202
231
232
+ lemma grundy_value_star : grundy_value star = 1 :=
233
+ by simp
234
+
203
235
@[simp] lemma grundy_value_nim_add_nim (n m : ℕ) :
204
236
grundy_value (nim.{u} n + nim.{u} m) = nat.lxor n m :=
205
237
begin
0 commit comments