Skip to content
New issue

Have a question about this project? Sign up for a free GitHub account to open an issue and contact its maintainers and the community.

By clicking “Sign up for GitHub”, you agree to our terms of service and privacy statement. We’ll occasionally send you account related emails.

Already on GitHub? Sign in to your account

chore(SetTheory/Game/Ordinal): make Ordinal.toGame an order embedding #19417

Open
wants to merge 2 commits into
base: master
Choose a base branch
from
Open
Show file tree
Hide file tree
Changes from all commits
Commits
File filter

Filter by extension

Filter by extension

Conversations
Failed to load comments.
Loading
Jump to
Jump to file
Failed to load files.
Loading
Diff view
Diff view
2 changes: 1 addition & 1 deletion Mathlib/SetTheory/Game/Birthday.lean
Original file line number Diff line number Diff line change
Expand Up @@ -207,7 +207,7 @@ theorem birthday_ordinalToGame (o : Ordinal) : birthday o.toGame = o := by
apply birthday_quot_le_pGameBirthday
· let ⟨x, hx₁, hx₂⟩ := birthday_eq_pGameBirthday o.toGame
rw [← hx₂, ← toPGame_le_iff]
rw [← PGame.equiv_iff_game_eq] at hx₁
rw [← mk_toPGame, ← PGame.equiv_iff_game_eq] at hx₁
exact hx₁.2.trans (PGame.le_birthday x)

@[simp, norm_cast]
Expand Down
37 changes: 19 additions & 18 deletions Mathlib/SetTheory/Game/Ordinal.lean
Original file line number Diff line number Diff line change
Expand Up @@ -155,15 +155,18 @@ noncomputable def toPGameEmbedding : Ordinal.{u} ↪o PGame.{u} where
map_rel_iff' := @toPGame_le_iff

/-- Converts an ordinal into the corresponding game. -/
noncomputable abbrev toGame (o : Ordinal) : Game := ⟦o.toPGame⟧

/-- The order embedding version of `toGame`. -/
@[simps]
noncomputable def toGameEmbedding : Ordinal.{u} ↪o Game.{u} where
toFun := Ordinal.toGame
noncomputable def toGame : Ordinal.{u} ↪o Game.{u} where
toFun o := ⟦o.toPGame⟧
inj' a b := by simpa [AntisymmRel] using le_antisymm
map_rel_iff' := toPGame_le_iff

@[simp]
theorem mk_toPGame (o : Ordinal) : ⟦o.toPGame⟧ = o.toGame :=
rfl

@[deprecated toGame (since := "2024-11-23")]
alias toGameEmbedding := toGame

@[simp]
theorem toGame_zero : toGame 0 = 0 :=
game_eq toPGame_zero
Expand All @@ -173,22 +176,20 @@ theorem toGame_one : toGame 1 = 1 :=
game_eq toPGame_one

theorem toGame_injective : Function.Injective toGame :=
toGameEmbedding.injective
toGame.injective

@[simp]
theorem toGame_lf_iff {a b : Ordinal} : Game.LF (toGame a) (toGame b) ↔ a < b :=
theorem toGame_lf_iff {a b : Ordinal} : Game.LF a.toGame b.toGame ↔ a < b :=
toPGame_lf_iff

@[simp]
theorem toGame_le_iff {a b : Ordinal} : toGame a ≤ toGame b ↔ a ≤ b :=
theorem toGame_le_iff {a b : Ordinal} : a.toGame ≤ b.toGame ↔ a ≤ b :=
toPGame_le_iff

@[simp]
theorem toGame_lt_iff {a b : Ordinal} : toGame a < toGame b ↔ a < b :=
theorem toGame_lt_iff {a b : Ordinal} : a.toGame < b.toGame ↔ a < b :=
toPGame_lt_iff

theorem toGame_eq_iff {a b : Ordinal} : toGame a = toGame b ↔ a = b :=
toGameEmbedding.inj
theorem toGame_eq_iff {a b : Ordinal} : a.toGame = b.toGame ↔ a = b :=
toGame.inj

/-- The natural addition of ordinals corresponds to their sum as games. -/
theorem toPGame_nadd (a b : Ordinal) : (a ♯ b).toPGame ≈ a.toPGame + b.toPGame := by
Expand Down Expand Up @@ -219,7 +220,7 @@ theorem toPGame_nmul (a b : Ordinal) : (a ⨳ b).toPGame ≈ a.toPGame * b.toPGa
refine ⟨le_of_forall_lf (fun i => ?_) isEmptyElim, le_of_forall_lf (fun i => ?_) isEmptyElim⟩
· rw [toPGame_moveLeft']
rcases lt_nmul_iff.1 (toLeftMovesToPGame_symm_lt i) with ⟨c, hc, d, hd, h⟩
rw [← toPGame_le_iff, le_iff_game_le, ← toGame, ← toGame, toGame_nadd _ _, toGame_nadd _ _,
rw [← toPGame_le_iff, le_iff_game_le, mk_toPGame, mk_toPGame, toGame_nadd _ _, toGame_nadd _ _,
← le_sub_iff_add_le] at h
refine lf_of_le_of_lf h <| (lf_congr_left ?_).1 <| moveLeft_lf <| toLeftMovesMul <| Sum.inl
⟨toLeftMovesToPGame ⟨c, hc⟩, toLeftMovesToPGame ⟨d, hd⟩⟩
Expand All @@ -232,15 +233,15 @@ theorem toPGame_nmul (a b : Ordinal) : (a ⨳ b).toPGame ≈ a.toPGame * b.toPGa
rw [mul_moveLeft_inl, toPGame_moveLeft', toPGame_moveLeft', lf_iff_game_lf,
quot_sub, quot_add, ← Game.not_le, le_sub_iff_add_le]
repeat rw [← game_eq (toPGame_nmul _ _)]
repeat rw [← toGame_nadd]
simp_rw [mk_toPGame, ← toGame_nadd]
apply toPGame_lf (nmul_nadd_lt _ _) <;>
exact toLeftMovesToPGame_symm_lt _
termination_by (a, b)

theorem toGame_nmul (a b : Ordinal) : (a ⨳ b).toGame = ⟦a.toPGame * b.toPGame⟧ :=
game_eq (toPGame_nmul a b)

@[simp, norm_cast]
@[simp]
Comment on lines -243 to +244
Copy link
Contributor

Choose a reason for hiding this comment

The reason will be displayed to describe this comment to others. Learn more.

Why this change?

Copy link
Collaborator Author

Choose a reason for hiding this comment

The reason will be displayed to describe this comment to others. Learn more.

"badly shaped simp lemma, can't start with coe"

theorem toGame_natCast : ∀ n : ℕ, toGame n = n
| 0 => Quot.sound (zeroToPGameRelabelling).equiv
| n + 1 => by
Expand All @@ -249,6 +250,6 @@ theorem toGame_natCast : ∀ n : ℕ, toGame n = n
rfl

theorem toPGame_natCast (n : ℕ) : toPGame n ≈ n := by
rw [PGame.equiv_iff_game_eq, ← toGame, toGame_natCast, quot_natCast]
rw [PGame.equiv_iff_game_eq, mk_toPGame, toGame_natCast, quot_natCast]

end Ordinal