@@ -19,14 +19,19 @@ We define the additive circle `add_circle p` as the quotient `𝕜 ⧸ (ℤ ∙
19
19
See also `circle` and `real.angle`. For the normed group structure on `add_circle`, see
20
20
`add_circle.normed_add_comm_group` in a later file.
21
21
22
- ## Main definitions:
22
+ ## Main definitions and results :
23
23
24
24
* `add_circle`: the additive circle `𝕜 ⧸ (ℤ ∙ p)` for some period `p : 𝕜`
25
25
* `unit_add_circle`: the special case `ℝ ⧸ ℤ`
26
26
* `add_circle.equiv_add_circle`: the rescaling equivalence `add_circle p ≃+ add_circle q`
27
- * `add_circle.equiv_Ico`: the natural equivalence `add_circle p ≃ Ico 0 p `
27
+ * `add_circle.equiv_Ico`: the natural equivalence `add_circle p ≃ Ico a (a + p) `
28
28
* `add_circle.add_order_of_div_of_gcd_eq_one`: rational points have finite order
29
29
* `add_circle.exists_gcd_eq_one_of_is_of_fin_add_order`: finite-order points are rational
30
+ * `add_circle.homeo_Icc_quot`: the natural topological equivalence between `add_circle p` and
31
+ `Icc a (a + p)` with its endpoints identified.
32
+ * `add_circle.lift_Ico_continuous`: if `f : ℝ → B` is continuous, and `f a = f (a + p)` for
33
+ some `a`, then there is a continuous function `add_circle p → B` which agrees with `f` on
34
+ `Icc a (a + p)`.
30
35
31
36
## Implementation notes:
32
37
@@ -43,9 +48,9 @@ the rational circle `add_circle (1 : ℚ)`, and so we set things up more general
43
48
44
49
noncomputable theory
45
50
46
- open set add_subgroup topological_space
51
+ open set function add_subgroup topological_space
47
52
48
- variables {𝕜 : Type *}
53
+ variables {𝕜 : Type *} {B : Type *}
49
54
50
55
/-- The "additive circle": `𝕜 ⧸ (ℤ ∙ p)`. See also `circle` and `real.angle`. -/
51
56
@[derive [add_comm_group, topological_space, topological_add_group, inhabited, has_coe_t 𝕜] ,
@@ -80,31 +85,62 @@ begin
80
85
{ exact ⟨(n : ℤ), by simp⟩, },
81
86
end
82
87
88
+ @[simp] lemma coe_add_period (x : 𝕜) : (((x + p) : 𝕜) : add_circle p) = x :=
89
+ begin
90
+ rw [quotient_add_group.coe_add, ←eq_sub_iff_add_eq', sub_self, quotient_add_group.eq_zero_iff],
91
+ exact mem_zmultiples p,
92
+ end
93
+
83
94
@[continuity, nolint unused_arguments] protected lemma continuous_mk' :
84
95
continuous (quotient_add_group.mk' (zmultiples p) : 𝕜 → add_circle p) :=
85
96
continuous_coinduced_rng
86
97
87
98
variables [hp : fact (0 < p)]
88
99
include hp
89
100
90
- variables [archimedean 𝕜]
101
+ variables (a : 𝕜) [archimedean 𝕜]
102
+
103
+ /-- The natural equivalence between `add_circle p` and the half-open interval `[a, a + p)`. -/
104
+ def equiv_Ico : add_circle p ≃ Ico a (a + p) := quotient_add_group.equiv_Ico_mod a hp.out
105
+
106
+ /-- Given a function on `[a, a + p)`, lift it to `add_circle p`. -/
107
+ def lift_Ico (f : 𝕜 → B) : add_circle p → B := restrict _ f ∘ add_circle.equiv_Ico p a
108
+
109
+ variables {p a}
110
+
111
+ lemma coe_eq_coe_iff_of_mem_Ico {x y : 𝕜}
112
+ (hx : x ∈ Ico a (a + p)) (hy : y ∈ Ico a (a + p)) : (x : add_circle p) = y ↔ x = y :=
113
+ begin
114
+ refine ⟨λ h, _, by tauto⟩,
115
+ suffices : (⟨x, hx⟩ : Ico a (a + p)) = ⟨y, hy⟩, by exact subtype.mk.inj this ,
116
+ apply_fun equiv_Ico p a at h,
117
+ rw [←(equiv_Ico p a).right_inv ⟨x, hx⟩, ←(equiv_Ico p a).right_inv ⟨y, hy⟩],
118
+ exact h
119
+ end
120
+
121
+ lemma lift_Ico_coe_apply {f : 𝕜 → B} {x : 𝕜} (hx : x ∈ Ico a (a + p)) : lift_Ico p a f ↑x = f x :=
122
+ begin
123
+ have : (equiv_Ico p a) x = ⟨x, hx⟩,
124
+ { rw equiv.apply_eq_iff_eq_symm_apply,
125
+ refl, },
126
+ rw [lift_Ico, comp_apply, this ],
127
+ refl,
128
+ end
91
129
92
- /-- The natural equivalence between `add_circle p` and the half-open interval `[0, p)`. -/
93
- def equiv_Ico : add_circle p ≃ Ico 0 p :=
94
- (quotient_add_group.equiv_Ico_mod 0 hp.out).trans $ equiv.set.of_eq $ by rw zero_add
130
+ variables (p a)
95
131
96
- @[continuity] lemma continuous_equiv_Ico_symm : continuous (equiv_Ico p).symm :=
132
+ @[continuity] lemma continuous_equiv_Ico_symm : continuous (equiv_Ico p a ).symm :=
97
133
continuous_quotient_mk.comp continuous_subtype_coe
98
134
99
- /-- The image of the closed-open interval `[0, p)` under the quotient map `𝕜 → add_circle p` is the
100
- entire space. -/
101
- @[simp] lemma coe_image_Ico_eq : (coe : 𝕜 → add_circle p) '' Ico 0 p = univ :=
102
- by { rw image_eq_range, exact (equiv_Ico p).symm.range_eq_univ }
135
+ /-- The image of the closed-open interval `[a, a + p)` under the quotient map `𝕜 → add_circle p` is
136
+ the entire space. -/
137
+ @[simp] lemma coe_image_Ico_eq : (coe : 𝕜 → add_circle p) '' Ico a (a + p) = univ :=
138
+ by { rw image_eq_range, exact (equiv_Ico p a ).symm.range_eq_univ }
103
139
104
140
/-- The image of the closed interval `[0, p]` under the quotient map `𝕜 → add_circle p` is the
105
141
entire space. -/
106
- @[simp] lemma coe_image_Icc_eq : (coe : 𝕜 → add_circle p) '' Icc 0 p = univ :=
107
- eq_top_mono (image_subset _ Ico_subset_Icc_self) $ coe_image_Ico_eq _
142
+ @[simp] lemma coe_image_Icc_eq : (coe : 𝕜 → add_circle p) '' Icc a (a + p) = univ :=
143
+ eq_top_mono (image_subset _ Ico_subset_Icc_self) $ coe_image_Ico_eq _ _
108
144
109
145
end linear_ordered_add_comm_group
110
146
@@ -133,19 +169,19 @@ section floor_ring
133
169
variables [floor_ring 𝕜]
134
170
135
171
@[simp] lemma coe_equiv_Ico_mk_apply (x : 𝕜) :
136
- (equiv_Ico p $ quotient_add_group.mk x : 𝕜) = int.fract (x / p) * p :=
172
+ (equiv_Ico p 0 $ quotient_add_group.mk x : 𝕜) = int.fract (x / p) * p :=
137
173
to_Ico_mod_eq_fract_mul _ x
138
174
139
175
instance : divisible_by (add_circle p) ℤ :=
140
- { div := λ x n, (↑(((n : 𝕜)⁻¹) * (equiv_Ico p x : 𝕜)) : add_circle p),
176
+ { div := λ x n, (↑(((n : 𝕜)⁻¹) * (equiv_Ico p 0 x : 𝕜)) : add_circle p),
141
177
div_zero := λ x,
142
178
by simp only [algebra_map.coe_zero, quotient_add_group.coe_zero, inv_zero, zero_mul],
143
179
div_cancel := λ n x hn,
144
180
begin
145
181
replace hn : (n : 𝕜) ≠ 0 , { norm_cast, assumption, },
146
- change n • quotient_add_group.mk' _ ((n : 𝕜)⁻¹ * ↑(equiv_Ico p x)) = x,
182
+ change n • quotient_add_group.mk' _ ((n : 𝕜)⁻¹ * ↑(equiv_Ico p 0 x)) = x,
147
183
rw [← map_zsmul, ← smul_mul_assoc, zsmul_eq_mul, mul_inv_cancel hn, one_mul],
148
- exact (equiv_Ico p).symm_apply_apply x,
184
+ exact (equiv_Ico p 0 ).symm_apply_apply x,
149
185
end , }
150
186
151
187
end floor_ring
@@ -224,11 +260,11 @@ begin
224
260
change ∃ m, gcd m n = 1 ∧ m < n ∧ ↑((↑m / ↑n) * p) = u,
225
261
have hn : 0 < n := add_order_of_pos' h,
226
262
have hn₀ : (n : 𝕜) ≠ 0 , { norm_cast, exact ne_of_gt hn, },
227
- let x := (equiv_Ico p u : 𝕜),
228
- have hxu : (x : add_circle p) = u := (equiv_Ico p).symm_apply_apply u,
263
+ let x := (equiv_Ico p 0 u : 𝕜),
264
+ have hxu : (x : add_circle p) = u := (equiv_Ico p 0 ).symm_apply_apply u,
229
265
have hx₀ : 0 < (add_order_of (x : add_circle p)), { rw ← hxu at h, exact add_order_of_pos' h, },
230
266
have hx₁ : 0 < x,
231
- { refine lt_of_le_of_ne (equiv_Ico p u).2 .1 _,
267
+ { refine lt_of_le_of_ne (equiv_Ico p 0 u).2 .1 _,
232
268
contrapose! hu,
233
269
rw [← hxu, ← hu, quotient_add_group.coe_zero], },
234
270
obtain ⟨m, hm : m • p = add_order_of ↑x • x⟩ := (coe_eq_zero_of_pos_iff p hp.out
@@ -239,8 +275,9 @@ begin
239
275
refine ⟨m, (_ : gcd m n = 1 ), (_ : m < n), hux⟩,
240
276
{ have := gcd_mul_add_order_of_div_eq p m hn,
241
277
rwa [hux, nat.mul_left_eq_self_iff hn] at this , },
242
- { have : n • x < n • p := smul_lt_smul_of_pos (equiv_Ico p u).2 .2 hn,
243
- rwa [nsmul_eq_mul, nsmul_eq_mul, ← hm, mul_lt_mul_right hp.out, nat.cast_lt] at this , },
278
+ { have : n • x < n • p := smul_lt_smul_of_pos _ hn,
279
+ rwa [nsmul_eq_mul, nsmul_eq_mul, ← hm, mul_lt_mul_right hp.out, nat.cast_lt] at this ,
280
+ simpa [zero_add] using (equiv_Ico p 0 u).2 .2 , },
244
281
end
245
282
246
283
end finite_order_points
@@ -252,7 +289,7 @@ variables (p : ℝ)
252
289
/-- The "additive circle" `ℝ ⧸ (ℤ ∙ p)` is compact. -/
253
290
instance compact_space [fact (0 < p)] : compact_space $ add_circle p :=
254
291
begin
255
- rw [← is_compact_univ_iff, ← coe_image_Icc_eq p],
292
+ rw [← is_compact_univ_iff, ← coe_image_Icc_eq p 0 ],
256
293
exact is_compact_Icc.image (add_circle.continuous_mk' p),
257
294
end
258
295
@@ -279,3 +316,100 @@ local attribute [instance] fact_zero_lt_one
279
316
/-- The unit circle `ℝ ⧸ ℤ`. -/
280
317
@[derive [compact_space, normal_space, second_countable_topology] ]
281
318
abbreviation unit_add_circle := add_circle (1 : ℝ)
319
+
320
+ section identify_Icc_ends
321
+ /-! This section proves that for any `a`, the natural map from `[a, a + p] ⊂ ℝ` to `add_circle p`
322
+ gives an identification of `add_circle p`, as a topological space, with the quotient of `[a, a + p]`
323
+ by the equivalence relation identifying the endpoints. -/
324
+
325
+ namespace add_circle
326
+
327
+ section linear_ordered_add_comm_group
328
+
329
+ variables [linear_ordered_add_comm_group 𝕜] [topological_space 𝕜] [order_topology 𝕜]
330
+ (p a : 𝕜) [hp : fact (0 < p)]
331
+
332
+ include hp
333
+
334
+ local notation `𝕋 ` := add_circle p
335
+
336
+ /-- The relation identifying the endpoints of `Icc a (a + p)`. -/
337
+ inductive endpoint_ident : Icc a (a + p) → Icc a (a + p) → Prop
338
+ | mk : endpoint_ident
339
+ ⟨a, left_mem_Icc.mpr $ le_add_of_nonneg_right hp.out.le⟩
340
+ ⟨a + p, right_mem_Icc.mpr $ le_add_of_nonneg_right hp.out.le⟩
341
+
342
+ variables [archimedean 𝕜]
343
+
344
+ /-- The equivalence between `add_circle p` and the quotient of `[a, a + p]` by the relation
345
+ identifying the endpoints. -/
346
+ def equiv_Icc_quot : 𝕋 ≃ quot (endpoint_ident p a) :=
347
+ { to_fun := λ x, quot.mk _ $ subtype.map id Ico_subset_Icc_self (equiv_Ico _ _ x),
348
+ inv_fun := λ x, quot.lift_on x coe $ by { rintro _ _ ⟨_⟩, exact (coe_add_period p a).symm },
349
+ left_inv := (equiv_Ico p a).symm_apply_apply,
350
+ right_inv := quot.ind $ by
351
+ { rintro ⟨x, hx⟩,
352
+ have := _,
353
+ rcases ne_or_eq x (a + p) with h | rfl,
354
+ { revert x, exact this },
355
+ { rw ← quot.sound endpoint_ident.mk, exact this _ _ (lt_add_of_pos_right a hp.out).ne },
356
+ intros x hx h,
357
+ congr, ext1,
358
+ apply congr_arg subtype.val ((equiv_Ico p a).right_inv ⟨x, hx.1 , hx.2 .lt_of_ne h⟩) } }
359
+
360
+ end linear_ordered_add_comm_group
361
+
362
+ section real
363
+
364
+ variables {p a : ℝ} [hp : fact (0 < p)]
365
+ include hp
366
+
367
+ local notation `𝕋 ` := add_circle p
368
+
369
+ /- doesn't work if inlined in `homeo_of_equiv_compact_to_t2` -- why? -/
370
+ private lemma continuous_equiv_Icc_quot_symm : continuous (equiv_Icc_quot p a).symm :=
371
+ continuous_quot_lift _ $ (add_circle.continuous_mk' p).comp continuous_subtype_coe
372
+
373
+ /-- The natural map from `[a, a + p] ⊂ ℝ` with endpoints identified to `ℝ / ℤ • p`, as a
374
+ homeomorphism of topological spaces. -/
375
+ def homeo_Icc_quot : 𝕋 ≃ₜ quot (endpoint_ident p a) :=
376
+ (continuous.homeo_of_equiv_compact_to_t2 continuous_equiv_Icc_quot_symm).symm
377
+
378
+ /-! We now show that a continuous function on `[a, a + p]` satisfying `f a = f (a + p)` is
379
+ the pullback of a continuous function on `unit_add_circle`. -/
380
+
381
+ lemma eq_of_end_ident {f : ℝ → B} (hf : f a = f (a + p)) (x y : Icc a (a + p)) :
382
+ endpoint_ident p a x y → f x = f y := by { rintro ⟨_⟩, exact hf }
383
+
384
+ lemma lift_Ico_eq_lift_Icc {f : ℝ → B} (h : f a = f (a + p)) :
385
+ lift_Ico p a f = (quot.lift (restrict (Icc a $ a + p) f) $ eq_of_end_ident h)
386
+ ∘ equiv_Icc_quot p a :=
387
+ funext (λ x, by refl)
388
+
389
+ lemma lift_Ico_continuous [topological_space B] {f : ℝ → B} (hf : f a = f (a + p))
390
+ (hc : continuous_on f $ Icc a (a + p)) : continuous (lift_Ico p a f) :=
391
+ begin
392
+ rw lift_Ico_eq_lift_Icc hf,
393
+ refine continuous.comp _ homeo_Icc_quot.continuous_to_fun,
394
+ exact continuous_coinduced_dom.mpr (continuous_on_iff_continuous_restrict.mp hc),
395
+ end
396
+
397
+ end real
398
+
399
+ section zero_based
400
+
401
+ variables {p : ℝ} [hp : fact (0 < p)]
402
+ include hp
403
+
404
+ lemma lift_Ico_zero_coe_apply {f : ℝ → B} {x : ℝ} (hx : x ∈ Ico 0 p) :
405
+ lift_Ico p 0 f ↑x = f x := lift_Ico_coe_apply (by rwa zero_add)
406
+
407
+ lemma lift_Ico_zero_continuous [topological_space B] {f : ℝ → B}
408
+ (hf : f 0 = f p) (hc : continuous_on f $ Icc 0 p) : continuous (lift_Ico p 0 f) :=
409
+ lift_Ico_continuous (by rwa zero_add : f 0 = f (0 + p)) (by rwa zero_add)
410
+
411
+ end zero_based
412
+
413
+ end add_circle
414
+
415
+ end identify_Icc_ends
0 commit comments