@@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE.
4
4
Authors: Apurva Nakade
5
5
-/
6
6
import algebra.algebra.basic
7
- import group_theory.monoid_localization
7
+ import ring_theory.localization
8
8
import set_theory.surreal.basic
9
9
10
10
/-!
@@ -178,33 +178,62 @@ begin
178
178
linarith },
179
179
end
180
180
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) :=
184
221
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,
200
224
end
201
225
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
+
202
231
/-- We define dyadic surreals as the range of the map `dyadic_map`. -/
203
232
def dyadic : set surreal := set.range dyadic_map
204
233
205
234
-- We conclude with some ideas for further work on surreals; these would make fun projects.
206
235
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
208
237
209
238
-- TODO map the reals into the surreals, using dyadic Dedekind cuts
210
239
-- TODO show this is a group homomorphism, and injective
0 commit comments