@@ -38,7 +38,7 @@ the rational circle `add_circle (1 : ℚ)`, and so we set things up more general
38
38
39
39
noncomputable theory
40
40
41
- open set int (hiding mem_zmultiples_iff) add_subgroup
41
+ open set int (hiding mem_zmultiples_iff) add_subgroup topological_space
42
42
43
43
variables {𝕜 : Type *}
44
44
80
80
(equiv_add_circle p q hp hq).symm (x : 𝕜) = (x * (q⁻¹ * p) : 𝕜) :=
81
81
rfl
82
82
83
- variables [floor_ring 𝕜]
83
+ variables [floor_ring 𝕜] [hp : fact (0 < p)]
84
+ include hp
84
85
85
86
/-- The natural equivalence between `add_circle p` and the half-open interval `[0, p)`. -/
86
- def equiv_Ico (hp : 0 < p) : add_circle p ≃ Ico 0 p :=
87
+ def equiv_Ico : add_circle p ≃ Ico 0 p :=
87
88
{ inv_fun := quotient_add_group.mk' _ ∘ coe,
88
- to_fun := λ x, ⟨(to_Ico_mod_periodic 0 hp).lift x, quot.induction_on x $ to_Ico_mod_mem_Ico' hp⟩,
89
+ to_fun := λ x, ⟨(to_Ico_mod_periodic 0 hp.out).lift x,
90
+ quot.induction_on x $ to_Ico_mod_mem_Ico' hp.out⟩,
89
91
right_inv := by { rintros ⟨x, hx⟩, ext, simp [to_Ico_mod_eq_self, hx.1 , hx.2 ], },
90
92
left_inv :=
91
93
begin
92
94
rintros ⟨x⟩,
93
- change quotient_add_group.mk (to_Ico_mod 0 hp x) = quotient_add_group.mk x,
95
+ change quotient_add_group.mk (to_Ico_mod 0 hp.out x) = quotient_add_group.mk x,
94
96
rw [quotient_add_group.eq', neg_add_eq_sub, self_sub_to_Ico_mod, zsmul_eq_mul],
95
97
apply int_cast_mul_mem_zmultiples,
96
98
end }
97
99
98
- @[simp] lemma coe_equiv_Ico_mk_apply (hp : 0 < p) ( x : 𝕜) :
99
- (equiv_Ico p hp $ quotient_add_group.mk x : 𝕜) = fract (x / p) * p :=
100
- to_Ico_mod_eq_fract_mul hp x
100
+ @[simp] lemma coe_equiv_Ico_mk_apply (x : 𝕜) :
101
+ (equiv_Ico p $ quotient_add_group.mk x : 𝕜) = fract (x / p) * p :=
102
+ to_Ico_mod_eq_fract_mul _ x
101
103
102
- @[continuity] lemma continuous_equiv_Ico_symm (hp : 0 < p) : continuous (equiv_Ico p hp ).symm :=
104
+ @[continuity] lemma continuous_equiv_Ico_symm : continuous (equiv_Ico p).symm :=
103
105
continuous_coinduced_rng.comp continuous_induced_dom
104
106
105
107
/-- The image of the closed interval `[0, p]` under the quotient map `𝕜 → add_circle p` is the
106
108
entire space. -/
107
- @[simp] lemma coe_image_Icc_eq (hp : 0 < p) :
109
+ @[simp] lemma coe_image_Icc_eq :
108
110
(coe : 𝕜 → add_circle p) '' (Icc 0 p) = univ :=
109
111
begin
110
112
refine eq_univ_iff_forall.mpr (λ x, _),
111
- let y := equiv_Ico p hp x,
112
- exact ⟨y, ⟨y.2 .1 , y.2 .2 .le⟩, (equiv_Ico p hp ).symm_apply_apply x⟩,
113
+ let y := equiv_Ico p x,
114
+ exact ⟨y, ⟨y.2 .1 , y.2 .2 .le⟩, (equiv_Ico p).symm_apply_apply x⟩,
113
115
end
114
116
115
117
end linear_ordered_field
116
118
117
119
variables (p : ℝ)
118
120
119
- lemma compact_space (hp : 0 < p) : compact_space $ add_circle p :=
121
+ /-- The "additive circle" `ℝ ⧸ (ℤ ∙ p)` is compact. -/
122
+ instance compact_space [fact (0 < p)] : compact_space $ add_circle p :=
120
123
begin
121
- rw [← is_compact_univ_iff, ← coe_image_Icc_eq p hp ],
124
+ rw [← is_compact_univ_iff, ← coe_image_Icc_eq p],
122
125
exact is_compact_Icc.image (add_circle.continuous_mk' p),
123
126
end
124
127
125
- end add_circle
128
+ /-- The action on `ℝ` by right multiplication of its the subgroup `zmultiples p` (the multiples of
129
+ `p:ℝ`) is properly discontinuous. -/
130
+ instance : properly_discontinuous_vadd (zmultiples p).opposite ℝ :=
131
+ (zmultiples p).properly_discontinuous_vadd_opposite_of_tendsto_cofinite
132
+ (add_subgroup.tendsto_zmultiples_subtype_cofinite p)
126
133
127
- /-- The unit circle `ℝ ⧸ ℤ`. -/
128
- abbreviation unit_add_circle := add_circle (1 : ℝ)
134
+ /-- The "additive circle" `ℝ ⧸ (ℤ ∙ p)` is Hausdorff. -/
135
+ instance : t2_space (add_circle p) := t2_space_of_properly_discontinuous_vadd_of_t2_space
136
+
137
+ /-- The "additive circle" `ℝ ⧸ (ℤ ∙ p)` is normal. -/
138
+ instance [fact (0 < p)] : normal_space (add_circle p) := normal_of_compact_t2
129
139
130
- namespace unit_add_circle
140
+ /-- The "additive circle" `ℝ ⧸ (ℤ ∙ p)` is second-countable. -/
141
+ instance : second_countable_topology (add_circle p) := quotient_add_group.second_countable_topology
131
142
132
- instance : compact_space unit_add_circle := add_circle.compact_space _ zero_lt_one
143
+ end add_circle
144
+
145
+ private lemma fact_zero_lt_one : fact ((0 :ℝ) < 1 ) := ⟨zero_lt_one⟩
146
+ local attribute [instance] fact_zero_lt_one
133
147
134
- end unit_add_circle
148
+ /-- The unit circle `ℝ ⧸ ℤ`. -/
149
+ @[derive [compact_space, normal_space, second_countable_topology] ]
150
+ abbreviation unit_add_circle := add_circle (1 : ℝ)
0 commit comments