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

Commit edf6cef

Browse files
committed
feat(set_theory/game/nim): nim 0 is a relabelling of 0 and nim 1 is a relabelling of star (#13846)
1 parent fd8474f commit edf6cef

File tree

2 files changed

+52
-7
lines changed

2 files changed

+52
-7
lines changed

src/set_theory/game/nim.lean

Lines changed: 35 additions & 3 deletions
Original file line numberDiff line numberDiff line change
@@ -57,6 +57,35 @@ lemma nim_def (O : ordinal) : nim O = pgame.mk O.out.α O.out.α
5757
(λ O₂, nim (ordinal.typein (<) O₂)) :=
5858
by { rw nim, refl }
5959

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+
6089
@[simp] lemma nim_birthday (O : ordinal) : (nim O).birthday = O :=
6190
begin
6291
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 :=
100129
begin
101130
rw [impartial.first_wins_symm, nim_def, lt_def_le],
102131
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
104133
end
105134

106135
@[simp] lemma sum_first_loses_iff_eq (O₁ O₂ : ordinal) : (nim O₁ + nim O₂).first_loses ↔ O₁ = O₂ :=
@@ -194,12 +223,15 @@ by simp
194223
grundy_value G = grundy_value H ↔ G ≈ H :=
195224
(grundy_value_eq_iff_equiv_nim _ _).trans (equiv_congr_left.1 (equiv_nim_grundy_value H) _).symm
196225

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
199228

200229
@[simp] lemma grundy_value_iff_equiv_zero (G : pgame) [G.impartial] : grundy_value G = 0 ↔ G ≈ 0 :=
201230
by rw [←grundy_value_eq_iff_equiv, grundy_value_zero]
202231

232+
lemma grundy_value_star : grundy_value star = 1 :=
233+
by simp
234+
203235
@[simp] lemma grundy_value_nim_add_nim (n m : ℕ) :
204236
grundy_value (nim.{u} n + nim.{u} m) = nat.lxor n m :=
205237
begin

src/set_theory/ordinal/arithmetic.lean

Lines changed: 17 additions & 4 deletions
Original file line numberDiff line numberDiff line change
@@ -176,12 +176,25 @@ type_ne_zero_iff_nonempty.2 ⟨punit.star⟩
176176
instance : nontrivial ordinal.{u} :=
177177
⟨⟨1, 0, ordinal.one_ne_zero⟩⟩
178178

179-
instance : nonempty (1 : ordinal).out.α :=
180-
out_nonempty_iff_ne_zero.2 ordinal.one_ne_zero
181-
182-
theorem zero_lt_one : (0 : ordinal) < 1 :=
179+
@[simp] theorem zero_lt_one : (0 : ordinal) < 1 :=
183180
lt_iff_le_and_ne.2 ⟨ordinal.zero_le _, ne.symm $ ordinal.one_ne_zero⟩
184181

182+
instance : unique (1 : ordinal).out.α :=
183+
{ default := enum (<) 0 (by simp),
184+
uniq := λ a, begin
185+
rw ←enum_typein (<) a,
186+
unfold default,
187+
congr,
188+
rw ←lt_one_iff_zero,
189+
apply typein_lt_self
190+
end }
191+
192+
theorem one_out_eq (x : (1 : ordinal).out.α) : x = enum (<) 0 (by simp) :=
193+
unique.eq_default x
194+
195+
@[simp] theorem typein_one_out (x : (1 : ordinal).out.α) : typein (<) x = 0 :=
196+
by rw [one_out_eq x, typein_enum]
197+
185198
theorem le_one_iff {a : ordinal} : a ≤ 1 ↔ a = 0 ∨ a = 1 :=
186199
begin
187200
refine ⟨λ ha, _, _⟩,

0 commit comments

Comments
 (0)