@@ -63,21 +63,29 @@ namespace pgame
63
63
and all the elements of L and R are also numeric. -/
64
64
def numeric : pgame → Prop
65
65
| ⟨l, r, L, R⟩ :=
66
- (∀ i j, L i < R j) ∧ (∀ i, numeric (L i)) ∧ (∀ i , numeric (R i ))
66
+ (∀ i j, L i < R j) ∧ (∀ i, numeric (L i)) ∧ (∀ j , numeric (R j ))
67
67
68
- lemma numeric_def ( x : pgame) : numeric x ↔ (∀ i j, x.move_left i < x.move_right j) ∧
69
- (∀ i, numeric (x.move_left i)) ∧ (∀ i , numeric (x.move_right i )) :=
68
+ lemma numeric_def { x : pgame} : numeric x ↔ (∀ i j, x.move_left i < x.move_right j) ∧
69
+ (∀ i, numeric (x.move_left i)) ∧ (∀ j , numeric (x.move_right j )) :=
70
70
by { cases x, refl }
71
71
72
- lemma numeric.left_lt_right {x : pgame} (o : numeric x) (i : x.left_moves) (j : x.right_moves) :
72
+ namespace numeric
73
+
74
+ lemma mk {x : pgame} (h₁ : ∀ i j, x.move_left i < x.move_right j)
75
+ (h₂ : ∀ i, numeric (x.move_left i)) (h₃ : ∀ j, numeric (x.move_right j)) : numeric x :=
76
+ numeric_def.2 ⟨h₁, h₂, h₃⟩
77
+
78
+ lemma left_lt_right {x : pgame} (o : numeric x) (i : x.left_moves) (j : x.right_moves) :
73
79
x.move_left i < x.move_right j :=
74
- by { cases x with xl xr xL xR , exact o.1 i j }
75
- lemma numeric. move_left {x : pgame} (o : numeric x) (i : x.left_moves) :
80
+ by { cases x, exact o.1 i j }
81
+ lemma move_left {x : pgame} (o : numeric x) (i : x.left_moves) :
76
82
numeric (x.move_left i) :=
77
- by { cases x with xl xr xL xR , exact o.2 .1 i }
78
- lemma numeric. move_right {x : pgame} (o : numeric x) (j : x.right_moves) :
83
+ by { cases x, exact o.2 .1 i }
84
+ lemma move_right {x : pgame} (o : numeric x) (j : x.right_moves) :
79
85
numeric (x.move_right j) :=
80
- by { cases x with xl xr xL xR, exact o.2 .2 j }
86
+ by { cases x, exact o.2 .2 j }
87
+
88
+ end numeric
81
89
82
90
@[elab_as_eliminator]
83
91
theorem numeric_rec {C : pgame → Prop }
@@ -151,33 +159,35 @@ theorem lt_or_equiv_or_gt {x y : pgame} (ox : numeric x) (oy : numeric y) : x <
151
159
152
160
theorem numeric_of_is_empty (x : pgame) [is_empty x.left_moves] [is_empty x.right_moves] :
153
161
numeric x :=
154
- (numeric_def x). 2 ⟨ is_empty_elim, is_empty_elim, is_empty_elim⟩
162
+ numeric.mk is_empty_elim is_empty_elim is_empty_elim
155
163
156
- theorem numeric_of_is_empty_left_moves (x : pgame) [is_empty x.left_moves]
157
- (H : ∀ j, numeric (x.move_right j)) : numeric x :=
158
- (numeric_def x). 2 ⟨ is_empty_elim, is_empty_elim, H⟩
164
+ theorem numeric_of_is_empty_left_moves (x : pgame) [is_empty x.left_moves] :
165
+ (∀ j, numeric (x.move_right j)) → numeric x :=
166
+ numeric.mk is_empty_elim is_empty_elim
159
167
160
168
theorem numeric_of_is_empty_right_moves (x : pgame) [is_empty x.right_moves]
161
169
(H : ∀ i, numeric (x.move_left i)) : numeric x :=
162
- (numeric_def x). 2 ⟨ λ _, is_empty_elim, H, is_empty_elim⟩
170
+ numeric.mk ( λ _, is_empty_elim) H is_empty_elim
163
171
164
172
theorem numeric_zero : numeric 0 := numeric_of_is_empty 0
165
173
theorem numeric_one : numeric 1 := numeric_of_is_empty_right_moves 1 $ λ _, numeric_zero
166
174
167
175
theorem numeric.neg : Π {x : pgame} (o : numeric x), numeric (-x)
168
176
| ⟨l, r, L, R⟩ o := ⟨λ j i, neg_lt_neg_iff.2 (o.1 i j), λ j, (o.2 .2 j).neg, λ i, (o.2 .1 i).neg⟩
169
177
170
- theorem numeric.move_left_lt {x : pgame} (o : numeric x) (i) : x.move_left i < x :=
178
+ namespace numeric
179
+
180
+ theorem move_left_lt {x : pgame} (o : numeric x) (i) : x.move_left i < x :=
171
181
(pgame.move_left_lf i).lt (o.move_left i) o
172
- theorem numeric. move_left_le {x : pgame} (o : numeric x) (i) : x.move_left i ≤ x :=
182
+ theorem move_left_le {x : pgame} (o : numeric x) (i) : x.move_left i ≤ x :=
173
183
(o.move_left_lt i).le
174
184
175
- theorem numeric. lt_move_right {x : pgame} (o : numeric x) (j) : x < x.move_right j :=
185
+ theorem lt_move_right {x : pgame} (o : numeric x) (j) : x < x.move_right j :=
176
186
(pgame.lf_move_right j).lt o (o.move_right j)
177
- theorem numeric. le_move_right {x : pgame} (o : numeric x) (j) : x ≤ x.move_right j :=
187
+ theorem le_move_right {x : pgame} (o : numeric x) (j) : x ≤ x.move_right j :=
178
188
(o.lt_move_right j).le
179
189
180
- theorem numeric. add : Π {x y : pgame} (ox : numeric x) (oy : numeric y), numeric (x + y)
190
+ theorem add : Π {x y : pgame} (ox : numeric x) (oy : numeric y), numeric (x + y)
181
191
| ⟨xl, xr, xL, xR⟩ ⟨yl, yr, yL, yR⟩ ox oy :=
182
192
⟨begin
183
193
rintros (ix|iy) (jx|jy),
@@ -199,7 +209,9 @@ theorem numeric.add : Π {x y : pgame} (ox : numeric x) (oy : numeric y), numeri
199
209
end ⟩
200
210
using_well_founded { dec_tac := pgame_wf_tac }
201
211
202
- lemma numeric.sub {x y : pgame} (ox : numeric x) (oy : numeric y) : numeric (x - y) := ox.add oy.neg
212
+ lemma sub {x y : pgame} (ox : numeric x) (oy : numeric y) : numeric (x - y) := ox.add oy.neg
213
+
214
+ end numeric
203
215
204
216
/-- Pre-games defined by natural numbers are numeric. -/
205
217
theorem numeric_nat : Π (n : ℕ), numeric n
0 commit comments