@@ -120,15 +120,6 @@ inductive pgame : Type (u+1)
120
120
121
121
namespace pgame
122
122
123
- /--
124
- Construct a pre-game from list of pre-games describing the available moves for Left and Right.
125
- -/
126
- -- TODO provide some API describing the interaction with
127
- -- `left_moves`, `right_moves`, `move_left` and `move_right` below.
128
- -- TODO define this at the level of games, as well, and perhaps also for finsets of games.
129
- def of_lists (L R : list pgame.{0 }) : pgame.{0 } :=
130
- pgame.mk (fin L.length) (fin R.length) (λ i, L.nth_le i i.is_lt) (λ j, R.nth_le j.val j.is_lt)
131
-
132
123
/-- The indexing type for allowable moves by Left. -/
133
124
def left_moves : pgame → Type u
134
125
| (mk l _ _ _) := l
@@ -148,6 +139,45 @@ def move_right : Π (g : pgame), right_moves g → pgame
148
139
@[simp] lemma right_moves_mk {xl xr xL xR} : (⟨xl, xr, xL, xR⟩ : pgame).right_moves = xr := rfl
149
140
@[simp] lemma move_right_mk {xl xr xL xR} : (⟨xl, xr, xL, xR⟩ : pgame).move_right = xR := rfl
150
141
142
+ /--
143
+ Construct a pre-game from list of pre-games describing the available moves for Left and Right.
144
+ -/
145
+ -- TODO define this at the level of games, as well, and perhaps also for finsets of games.
146
+ def of_lists (L R : list pgame.{u}) : pgame.{u} :=
147
+ mk (ulift (fin L.length)) (ulift (fin R.length))
148
+ (λ i, L.nth_le i.down i.down.is_lt) (λ j, R.nth_le j.down.val j.down.is_lt)
149
+
150
+ lemma left_moves_of_lists (L R : list pgame) : (of_lists L R).left_moves = ulift (fin L.length) :=
151
+ rfl
152
+ lemma right_moves_of_lists (L R : list pgame) : (of_lists L R).right_moves = ulift (fin R.length) :=
153
+ rfl
154
+
155
+ /-- Converts a number into a left move for `of_lists`. -/
156
+ def to_of_lists_left_moves {L R : list pgame} : fin L.length ≃ (of_lists L R).left_moves :=
157
+ ((equiv.cast (left_moves_of_lists L R).symm).trans equiv.ulift).symm
158
+
159
+ /-- Converts a number into a right move for `of_lists`. -/
160
+ def to_of_lists_right_moves {L R : list pgame} : fin R.length ≃ (of_lists L R).right_moves :=
161
+ ((equiv.cast (right_moves_of_lists L R).symm).trans equiv.ulift).symm
162
+
163
+ theorem of_lists_move_left {L R : list pgame} (i : fin L.length) :
164
+ (of_lists L R).move_left (to_of_lists_left_moves i) = L.nth_le i i.is_lt :=
165
+ rfl
166
+
167
+ @[simp] theorem of_lists_move_left' {L R : list pgame} (i : (of_lists L R).left_moves) :
168
+ (of_lists L R).move_left i =
169
+ L.nth_le (to_of_lists_left_moves.symm i) (to_of_lists_left_moves.symm i).is_lt :=
170
+ rfl
171
+
172
+ theorem of_lists_move_right {L R : list pgame} (i : fin R.length) :
173
+ (of_lists L R).move_right (to_of_lists_right_moves i) = R.nth_le i i.is_lt :=
174
+ rfl
175
+
176
+ @[simp] theorem of_lists_move_right' {L R : list pgame} (i : (of_lists L R).right_moves) :
177
+ (of_lists L R).move_right i =
178
+ R.nth_le (to_of_lists_right_moves.symm i) (to_of_lists_right_moves.symm i).is_lt :=
179
+ rfl
180
+
151
181
/-- A variant of `pgame.rec_on` expressed in terms of `pgame.move_left` and `pgame.move_right`.
152
182
153
183
Both this and `pgame.rec_on` describe Conway induction on games. -/
@@ -217,22 +247,23 @@ meta def pgame_wf_tac :=
217
247
/-- The pre-game `zero` is defined by `0 = { | }`. -/
218
248
instance : has_zero pgame := ⟨⟨pempty, pempty, pempty.elim, pempty.elim⟩⟩
219
249
220
- @[simp] lemma zero_left_moves : ( 0 : pgame). left_moves = pempty := rfl
221
- @[simp] lemma zero_right_moves : ( 0 : pgame). right_moves = pempty := rfl
250
+ @[simp] lemma zero_left_moves : left_moves 0 = pempty := rfl
251
+ @[simp] lemma zero_right_moves : right_moves 0 = pempty := rfl
222
252
223
- instance is_empty_zero_left_moves : is_empty (0 : pgame).left_moves := pempty.is_empty
224
- instance is_empty_zero_right_moves : is_empty (0 : pgame).right_moves := pempty.is_empty
253
+ instance is_empty_zero_left_moves : is_empty (left_moves 0 ) := pempty.is_empty
254
+ instance is_empty_zero_right_moves : is_empty (right_moves 0 ) := pempty.is_empty
225
255
226
256
instance : inhabited pgame := ⟨0 ⟩
227
257
228
258
/-- The pre-game `one` is defined by `1 = { 0 | }`. -/
229
259
instance : has_one pgame := ⟨⟨punit, pempty, λ _, 0 , pempty.elim⟩⟩
230
260
231
- @[simp] lemma one_left_moves : ( 1 : pgame). left_moves = punit := rfl
232
- @[simp] lemma one_move_left : ( 1 : pgame). move_left punit.star = 0 := rfl
233
- @[simp] lemma one_right_moves : ( 1 : pgame). right_moves = pempty := rfl
261
+ @[simp] lemma one_left_moves : left_moves 1 = punit := rfl
262
+ @[simp] lemma one_move_left (x) : move_left 1 x = 0 := rfl
263
+ @[simp] lemma one_right_moves : right_moves 1 = pempty := rfl
234
264
235
- instance is_empty_one_right_moves : is_empty (1 : pgame).right_moves := pempty.is_empty
265
+ instance unique_one_left_moves : unique (left_moves 1 ) := punit.unique
266
+ instance is_empty_one_right_moves : is_empty (right_moves 1 ) := pempty.is_empty
236
267
237
268
/-- Define simultaneously by mutual induction the `<=` and `<`
238
269
relation on pre-games. The ZFC definition says that `x = {xL | xR}`
@@ -308,38 +339,28 @@ end
308
339
/-- The definition of `x ≤ 0` on pre-games, in terms of `≤ 0` two moves later. -/
309
340
theorem le_zero {x : pgame} : x ≤ 0 ↔
310
341
∀ i : x.left_moves, ∃ j : (x.move_left i).right_moves, (x.move_left i).move_right j ≤ 0 :=
311
- begin
312
- rw le_def,
313
- dsimp,
314
- simp [forall_pempty, exists_pempty]
315
- end
342
+ by { rw le_def, dsimp, simp [forall_pempty, exists_pempty] }
316
343
317
344
/-- The definition of `0 ≤ x` on pre-games, in terms of `0 ≤` two moves later. -/
318
345
theorem zero_le {x : pgame} : 0 ≤ x ↔
319
346
∀ j : x.right_moves, ∃ i : (x.move_right j).left_moves, 0 ≤ (x.move_right j).move_left i :=
320
- begin
321
- rw le_def,
322
- dsimp,
323
- simp [forall_pempty, exists_pempty]
324
- end
347
+ by { rw le_def, dsimp, simp [forall_pempty, exists_pempty] }
325
348
326
349
/-- The definition of `x < 0` on pre-games, in terms of `< 0` two moves later. -/
327
350
theorem lt_zero {x : pgame} : x < 0 ↔
328
351
∃ j : x.right_moves, ∀ i : (x.move_right j).left_moves, (x.move_right j).move_left i < 0 :=
329
- begin
330
- rw lt_def,
331
- dsimp,
332
- simp [forall_pempty, exists_pempty]
333
- end
352
+ by { rw lt_def, dsimp, simp [forall_pempty, exists_pempty] }
334
353
335
354
/-- The definition of `0 < x` on pre-games, in terms of `< x` two moves later. -/
336
355
theorem zero_lt {x : pgame} : 0 < x ↔
337
356
∃ i : x.left_moves, ∀ j : (x.move_left i).right_moves, 0 < (x.move_left i).move_right j :=
338
- begin
339
- rw lt_def,
340
- dsimp,
341
- simp [forall_pempty, exists_pempty]
342
- end
357
+ by { rw lt_def, dsimp, simp [forall_pempty, exists_pempty] }
358
+
359
+ @[simp] theorem le_zero_of_is_empty_left_moves (x : pgame) [is_empty x.left_moves] : x ≤ 0 :=
360
+ le_zero.2 is_empty_elim
361
+
362
+ @[simp] theorem zero_le_of_is_empty_right_moves (x : pgame) [is_empty x.right_moves] : 0 ≤ x :=
363
+ zero_le.2 is_empty_elim
343
364
344
365
/-- Given a right-player-wins game, provide a response to any move by left. -/
345
366
noncomputable def right_response {x : pgame} (h : x ≤ 0 ) (i : x.left_moves) :
@@ -1105,54 +1126,56 @@ theorem lt_iff_sub_pos {x y : pgame} : x < y ↔ 0 < y - x :=
1105
1126
... ≤ y : (add_zero_relabelling y).le⟩
1106
1127
1107
1128
/-- The pre-game `star`, which is fuzzy/confused with zero. -/
1108
- def star : pgame := pgame. of_lists [0 ] [0 ]
1129
+ def star : pgame.{u} := of_lists [0 ] [0 ]
1109
1130
1110
- instance inhabited_star_left_moves : inhabited star.left_moves :=
1111
- show (inhabited (fin 1 )), by apply_instance
1131
+ @[simp] theorem star_left_moves : star.left_moves = ulift (fin 1 ) := rfl
1132
+ @[simp] theorem star_right_moves : star.right_moves = ulift (fin 1 ) := rfl
1112
1133
1113
- instance inhabited_star_right_moves : inhabited star.right_moves :=
1114
- show (inhabited (fin 1 )), by apply_instance
1134
+ @[simp] theorem star_move_left (x) : star.move_left x = 0 :=
1135
+ show (of_lists _ _).move_left x = 0 , by simp
1136
+ @[simp] theorem star_move_right (x) : star.move_right x = 0 :=
1137
+ show (of_lists _ _).move_right x = 0 , by simp
1115
1138
1116
- theorem star_lt_zero : star < 0 :=
1117
- by rw lt_def; exact
1118
- or.inr ⟨⟨0 , zero_lt_one⟩, (by split; rintros ⟨⟩)⟩
1139
+ instance unique_star_left_moves : unique star.left_moves :=
1140
+ @equiv.unique _ (fin 1 ) _ equiv.ulift
1141
+ instance unique_star_right_moves : unique star.right_moves :=
1142
+ @equiv.unique _ (fin 1 ) _ equiv.ulift
1119
1143
1144
+ theorem star_lt_zero : star < 0 :=
1145
+ by { rw lt_zero, use default, rintros ⟨⟩ }
1120
1146
theorem zero_lt_star : 0 < star :=
1121
- by rw lt_def; exact
1122
- or.inl ⟨⟨0 , zero_lt_one⟩, (by split; rintros ⟨⟩)⟩
1147
+ by { rw zero_lt, use default, rintros ⟨⟩ }
1123
1148
1124
1149
/-- The pre-game `ω`. (In fact all ordinals have game and surreal representatives.) -/
1125
1150
def omega : pgame := ⟨ulift ℕ, pempty, λ n, ↑n.1 , pempty.elim⟩
1126
1151
1127
- theorem zero_lt_one : (0 : pgame) < 1 :=
1128
- begin
1129
- rw lt_def,
1130
- left,
1131
- use ⟨punit.star, by split; rintro ⟨ ⟩⟩,
1132
- end
1152
+ @[simp] theorem zero_lt_one : (0 : pgame) < 1 :=
1153
+ by { rw zero_lt, use default, rintro ⟨⟩ }
1154
+
1155
+ theorem zero_le_one : (0 : pgame) ≤ 1 :=
1156
+ zero_le_of_is_empty_right_moves 1
1133
1157
1134
1158
/-- The pre-game `half` is defined as `{0 | 1}`. -/
1135
1159
def half : pgame := ⟨punit, punit, 0 , 1 ⟩
1136
1160
1137
- @[simp] lemma half_move_left : half.move_left punit.star = 0 := rfl
1161
+ @[simp] theorem half_left_moves : half.left_moves = punit := rfl
1162
+ @[simp] theorem half_right_moves : half.right_moves = punit := rfl
1163
+ @[simp] lemma half_move_left (x) : half.move_left x = 0 := rfl
1164
+ @[simp] lemma half_move_right (x) : half.move_right x = 1 := rfl
1138
1165
1139
- @[simp] lemma half_move_right : half.move_right punit.star = 1 := rfl
1166
+ instance unique_half_left_moves : unique half.left_moves := punit.unique
1167
+ instance unique_half_right_moves : unique half.right_moves := punit.unique
1140
1168
1141
1169
protected theorem zero_lt_half : 0 < half :=
1142
- begin
1143
- rw lt_def,
1144
- left,
1145
- use punit.star,
1146
- split; rintro ⟨ ⟩,
1147
- end
1170
+ by { rw zero_lt, use default, rintro ⟨⟩ }
1148
1171
1149
1172
theorem half_lt_one : half < 1 :=
1150
1173
begin
1151
1174
rw lt_def,
1152
1175
right,
1153
- use punit.star ,
1154
- split; rintro ⟨ ⟩,
1155
- exact zero_lt_one,
1176
+ use default ,
1177
+ split; rintro ⟨⟩,
1178
+ exact zero_lt_one
1156
1179
end
1157
1180
1158
1181
end pgame
0 commit comments