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

Commit 28a6f0a

Browse files
committed
feat(set_theory/surreal/basic): add numeric.mk lemma, golf (#14962)
1 parent 54352be commit 28a6f0a

File tree

1 file changed

+32
-20
lines changed

1 file changed

+32
-20
lines changed

src/set_theory/surreal/basic.lean

Lines changed: 32 additions & 20 deletions
Original file line numberDiff line numberDiff line change
@@ -63,21 +63,29 @@ namespace pgame
6363
and all the elements of L and R are also numeric. -/
6464
def numeric : pgame → Prop
6565
| ⟨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))
6767

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)) :=
7070
by { cases x, refl }
7171

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) :
7379
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) :
7682
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) :
7985
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
8189

8290
@[elab_as_eliminator]
8391
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 <
151159

152160
theorem numeric_of_is_empty (x : pgame) [is_empty x.left_moves] [is_empty x.right_moves] :
153161
numeric x :=
154-
(numeric_def x).2is_empty_elim, is_empty_elim, is_empty_elim
162+
numeric.mk is_empty_elim is_empty_elim is_empty_elim
155163

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).2is_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
159167

160168
theorem numeric_of_is_empty_right_moves (x : pgame) [is_empty x.right_moves]
161169
(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
163171

164172
theorem numeric_zero : numeric 0 := numeric_of_is_empty 0
165173
theorem numeric_one : numeric 1 := numeric_of_is_empty_right_moves 1 $ λ _, numeric_zero
166174

167175
theorem numeric.neg : Π {x : pgame} (o : numeric x), numeric (-x)
168176
| ⟨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⟩
169177

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 :=
171181
(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 :=
173183
(o.move_left_lt i).le
174184

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 :=
176186
(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 :=
178188
(o.lt_move_right j).le
179189

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)
181191
| ⟨xl, xr, xL, xR⟩ ⟨yl, yr, yL, yR⟩ ox oy :=
182192
begin
183193
rintros (ix|iy) (jx|jy),
@@ -199,7 +209,9 @@ theorem numeric.add : Π {x y : pgame} (ox : numeric x) (oy : numeric y), numeri
199209
end
200210
using_well_founded { dec_tac := pgame_wf_tac }
201211

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
203215

204216
/-- Pre-games defined by natural numbers are numeric. -/
205217
theorem numeric_nat : Π (n : ℕ), numeric n

0 commit comments

Comments
 (0)