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

Commit 8f73b07

Browse files
committed
feat(set_theory/surreal/dyadic): define add_monoid_hom structure on dyadic map (#11052)
The proof is mechanical and mostly requires unraveling definitions. The above map cannot be extended to ring morphism as so far there's not multiplication structure on surreal numbers.
1 parent 32cd278 commit 8f73b07

File tree

3 files changed

+59
-22
lines changed

3 files changed

+59
-22
lines changed

src/group_theory/submonoid/membership.lean

Lines changed: 7 additions & 2 deletions
Original file line numberDiff line numberDiff line change
@@ -223,6 +223,8 @@ lemma powers_subset {n : M} {P : submonoid M} (h : n ∈ P) : powers n ≤ P :=
223223
@[simps] def pow (n : M) (m : ℕ) : powers n :=
224224
(powers_hom M n).mrange_restrict (multiplicative.of_add m)
225225

226+
lemma pow_apply (n : M) (m : ℕ) : submonoid.pow n m = ⟨n ^ m, m, rfl⟩ := rfl
227+
226228
/-- Logarithms from powers to natural numbers. -/
227229
def log [decidable_eq M] {n : M} (p : powers n) : ℕ :=
228230
nat.find $ (mem_powers_iff p.val n).mp p.prop
@@ -234,8 +236,8 @@ lemma pow_right_injective_iff_pow_injective {n : M} :
234236
function.injective (λ m : ℕ, n ^ m) ↔ function.injective (pow n) :=
235237
subtype.coe_injective.of_comp_iff (pow n)
236238

237-
theorem log_pow_eq_self [decidable_eq M] {n : M} (h : function.injective (λ m : ℕ, n ^ m)) (m : ℕ) :
238-
log (pow n m) = m :=
239+
@[simp] theorem log_pow_eq_self [decidable_eq M] {n : M} (h : function.injective (λ m : ℕ, n ^ m))
240+
(m : ℕ) : log (pow n m) = m :=
239241
pow_right_injective_iff_pow_injective.mp h $ pow_log_eq_self _
240242

241243
/-- The exponentiation map is an isomorphism from the additive monoid on natural numbers to powers
@@ -249,6 +251,9 @@ def pow_log_equiv [decidable_eq M] {n : M} (h : function.injective (λ m : ℕ,
249251
right_inv := pow_log_eq_self,
250252
map_mul' := λ _ _, by { simp only [pow, map_mul, of_add_add, to_add_mul] } }
251253

254+
lemma log_mul [decidable_eq M] {n : M} (h : function.injective (λ m : ℕ, n ^ m))
255+
(x y : powers (n : M)) : log (x * y) = log x + log y := (pow_log_equiv h).symm.map_mul x y
256+
252257
theorem log_pow_int_eq_self {x : ℤ} (h : 1 < x.nat_abs) (m : ℕ) : log (pow x m) = m :=
253258
(pow_log_equiv (int.pow_right_injective h)).symm_apply_apply _
254259

src/ring_theory/localization.lean

Lines changed: 3 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -985,6 +985,9 @@ lemma mk_zero (b) : (mk 0 b : localization M) = 0 :=
985985
calc mk 0 b = mk 0 1 : mk_eq_mk_iff.mpr (r_of_eq (by simp))
986986
... = 0 : by unfold has_zero.zero localization.zero
987987

988+
lemma lift_on_zero {p : Type*} (f : ∀ (a : R) (b : M), p) (H) : lift_on 0 f H = f 0 1 :=
989+
by rw [← mk_zero 1, lift_on_mk]
990+
988991
private meta def tac := `[
989992
{ intros,
990993
simp only [add_mk, localization.mk_mul, neg_mk, ← mk_zero 1],

src/set_theory/surreal/dyadic.lean

Lines changed: 49 additions & 20 deletions
Original file line numberDiff line numberDiff line change
@@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE.
44
Authors: Apurva Nakade
55
-/
66
import algebra.algebra.basic
7-
import group_theory.monoid_localization
7+
import ring_theory.localization
88
import set_theory.surreal.basic
99

1010
/-!
@@ -178,33 +178,62 @@ begin
178178
linarith },
179179
end
180180

181-
/-- The map `dyadic_map` sends ⟦⟨m, 2^n⟩⟧ to m • half ^ n. -/
182-
def dyadic_map (x : localization (submonoid.powers (2 : ℤ))) : surreal :=
183-
localization.lift_on x (λ x y, x • pow_half (submonoid.log y)) $
181+
/-- The additive monoid morphism `dyadic_map` sends ⟦⟨m, 2^n⟩⟧ to m • half ^ n. -/
182+
def dyadic_map : localization.away (2 : ℤ) →+ surreal :=
183+
{ to_fun :=
184+
λ x, localization.lift_on x (λ x y, x • pow_half (submonoid.log y)) $
185+
begin
186+
intros m₁ m₂ n₁ n₂ h₁,
187+
obtain ⟨⟨n₃, y₃, hn₃⟩, h₂⟩ := localization.r_iff_exists.mp h₁,
188+
simp only [subtype.coe_mk, mul_eq_mul_right_iff] at h₂,
189+
cases h₂,
190+
{ simp only,
191+
obtain ⟨a₁, ha₁⟩ := n₁.prop,
192+
obtain ⟨a₂, ha₂⟩ := n₂.prop,
193+
have hn₁ : n₁ = submonoid.pow 2 a₁ := subtype.ext ha₁.symm,
194+
have hn₂ : n₂ = submonoid.pow 2 a₂ := subtype.ext ha₂.symm,
195+
have h₂ : 1 < (2 : ℤ).nat_abs, from one_lt_two,
196+
rw [hn₁, hn₂, submonoid.log_pow_int_eq_self h₂, submonoid.log_pow_int_eq_self h₂],
197+
apply dyadic_aux,
198+
rwa [ha₁, ha₂] },
199+
{ have := nat.one_le_pow y₃ 2 nat.succ_pos',
200+
linarith }
201+
end,
202+
map_zero' := localization.lift_on_zero _ _,
203+
map_add' := λ x y, localization.induction_on₂ x y $
204+
begin
205+
rintro ⟨a, ⟨b, ⟨b', rfl⟩⟩⟩ ⟨c, ⟨d, ⟨d', rfl⟩⟩⟩,
206+
have h₂ : 1 < (2 : ℤ).nat_abs, from one_lt_two,
207+
have hpow₂ := submonoid.log_pow_int_eq_self h₂,
208+
simp_rw submonoid.pow_apply at hpow₂,
209+
simp_rw [localization.add_mk, localization.lift_on_mk, subtype.coe_mk,
210+
submonoid.log_mul (int.pow_right_injective h₂), hpow₂],
211+
calc (2 ^ b' * c + 2 ^ d' * a) • pow_half (b' + d')
212+
= (c * 2 ^ b') • pow_half (b' + d') + (a * 2 ^ d') • pow_half (d' + b')
213+
: by simp only [add_smul, mul_comm,add_comm]
214+
... = c • pow_half d' + a • pow_half b' : by simp only [zsmul_pow_two_pow_half]
215+
... = a • pow_half b' + c • pow_half d' : add_comm _ _,
216+
end }
217+
218+
@[simp] lemma dyadic_map_apply (m : ℤ) (p : submonoid.powers (2 : ℤ)) :
219+
dyadic_map (is_localization.mk' (localization (submonoid.powers 2)) m p) =
220+
m • pow_half (submonoid.log p) :=
184221
begin
185-
intros m₁ m₂ n₁ n₂ h₁,
186-
obtain ⟨⟨n₃, y₃, hn₃⟩, h₂⟩ := localization.r_iff_exists.mp h₁,
187-
simp only [subtype.coe_mk, mul_eq_mul_right_iff] at h₂,
188-
cases h₂,
189-
{ simp only,
190-
obtain ⟨a₁, ha₁⟩ := n₁.prop,
191-
obtain ⟨a₂, ha₂⟩ := n₂.prop,
192-
have hn₁ : n₁ = submonoid.pow 2 a₁ := subtype.ext ha₁.symm,
193-
have hn₂ : n₂ = submonoid.pow 2 a₂ := subtype.ext ha₂.symm,
194-
have h₂ : 1 < (2 : ℤ).nat_abs, from dec_trivial,
195-
rw [hn₁, hn₂, submonoid.log_pow_int_eq_self h₂, submonoid.log_pow_int_eq_self h₂],
196-
apply dyadic_aux,
197-
rwa [ha₁, ha₂] },
198-
{ have := nat.one_le_pow y₃ 2 nat.succ_pos',
199-
linarith }
222+
rw ← localization.mk_eq_mk',
223+
refl,
200224
end
201225

226+
@[simp] lemma dyadic_map_apply_pow (m : ℤ) (n : ℕ) :
227+
dyadic_map (is_localization.mk' (localization (submonoid.powers 2)) m (submonoid.pow 2 n)) =
228+
m • pow_half n :=
229+
by rw [dyadic_map_apply, @submonoid.log_pow_int_eq_self 2 one_lt_two]
230+
202231
/-- We define dyadic surreals as the range of the map `dyadic_map`. -/
203232
def dyadic : set surreal := set.range dyadic_map
204233

205234
-- We conclude with some ideas for further work on surreals; these would make fun projects.
206235

207-
-- TODO show that the map from dyadic rationals to surreals is a group homomorphism, and injective
236+
-- TODO show that the map from dyadic rationals to surreals is injective
208237

209238
-- TODO map the reals into the surreals, using dyadic Dedekind cuts
210239
-- TODO show this is a group homomorphism, and injective

0 commit comments

Comments
 (0)