@@ -3,8 +3,8 @@ Copyright (c) 2021 Heather Macbeth. All rights reserved.
3
3
Released under Apache 2.0 license as described in the file LICENSE.
4
4
Authors: Heather Macbeth
5
5
-/
6
- import analysis.complex.basic
7
- import data.complex.exponential
6
+ import analysis.special_functions.exp
7
+ import topology.continuous_function.basic
8
8
9
9
/-!
10
10
# The circle
@@ -86,16 +86,23 @@ instance : topological_group circle :=
86
86
complex.conj_cle.continuous.comp continuous_subtype_coe }
87
87
88
88
/-- The map `λ t, exp (t * I)` from `ℝ` to the unit circle in `ℂ`. -/
89
- def exp_map_circle (t : ℝ) : circle :=
90
- ⟨exp (t * I), by simp [exp_mul_I, abs_cos_add_sin_mul_I]⟩
89
+ def exp_map_circle : C(ℝ, circle) :=
90
+ { to_fun := λ t, ⟨exp (t * I), by simp [exp_mul_I, abs_cos_add_sin_mul_I]⟩ }
91
91
92
92
@[simp] lemma exp_map_circle_apply (t : ℝ) : ↑(exp_map_circle t) = complex.exp (t * complex.I) :=
93
93
rfl
94
94
95
+ @[simp] lemma exp_map_circle_zero : exp_map_circle 0 = 1 :=
96
+ subtype.ext $ by rw [exp_map_circle_apply, of_real_zero, zero_mul, exp_zero, submonoid.coe_one]
97
+
98
+ @[simp] lemma exp_map_circle_add (x y : ℝ) :
99
+ exp_map_circle (x + y) = exp_map_circle x * exp_map_circle y :=
100
+ subtype.ext $ by simp only [exp_map_circle_apply, submonoid.coe_mul, of_real_add, add_mul,
101
+ complex.exp_add]
102
+
95
103
/-- The map `λ t, exp (t * I)` from `ℝ` to the unit circle in `ℂ`, considered as a homomorphism of
96
104
groups. -/
97
105
def exp_map_circle_hom : ℝ →+ (additive circle) :=
98
- { to_fun := exp_map_circle,
99
- map_zero' := by { rw exp_map_circle, convert of_mul_one, simp },
100
- map_add' := λ x y, show exp_map_circle (x + y) = (exp_map_circle x) * (exp_map_circle y),
101
- from subtype.ext $ by simp [exp_map_circle, exp_add, add_mul] }
106
+ { to_fun := additive.of_mul ∘ exp_map_circle,
107
+ map_zero' := exp_map_circle_zero,
108
+ map_add' := exp_map_circle_add }
0 commit comments