@@ -99,24 +99,36 @@ begin
99
99
coe_power_series],
100
100
end
101
101
102
+ @[simp] lemma of_power_series_C (r : R) :
103
+ of_power_series ℤ R (power_series.C R r) = hahn_series.C r :=
104
+ begin
105
+ ext n,
106
+ simp only [C, single_coeff, of_power_series_apply, ring_hom.coe_mk],
107
+ split_ifs with hn hn,
108
+ { rw [hn, ←int.coe_nat_zero],
109
+ convert @emb_domain_coeff _ _ _ _ _ _ _ _ 0 ,
110
+ simp },
111
+ { rw emb_domain_notin_image_support,
112
+ simp only [not_exists, set.mem_image, to_power_series_symm_apply_coeff, mem_support,
113
+ power_series.coeff_C],
114
+ intro,
115
+ simp [ne.symm hn] {contextual := tt} }
116
+ end
117
+
102
118
@[simp] lemma of_power_series_X :
103
119
of_power_series ℤ R power_series.X = single 1 1 :=
104
120
begin
105
121
ext n,
106
- cases n,
107
- { rw [int.of_nat_eq_coe, ← int.nat_cast_eq_coe_nat, of_power_series_apply_coeff],
108
- by_cases h1 : n = 1 ,
109
- { simp [h1] },
110
- { rw [power_series.coeff_X, single_coeff, if_neg h1, if_neg],
111
- contrapose! h1,
112
- rw [← nat.cast_one] at h1,
113
- exact nat.cast_injective h1 } },
114
- { rw [of_power_series_apply, emb_domain_notin_range, single_coeff_of_ne],
115
- { dec_trivial },
116
- rw [set.mem_range, not_exists],
117
- intro m,
118
- simp only [rel_embedding.coe_fn_mk, function.embedding.coe_fn_mk, int.nat_cast_eq_coe_nat],
119
- dec_trivial }
122
+ simp only [single_coeff, of_power_series_apply, ring_hom.coe_mk],
123
+ split_ifs with hn hn,
124
+ { rw [hn, ←int.coe_nat_one],
125
+ convert @emb_domain_coeff _ _ _ _ _ _ _ _ 1 ,
126
+ simp },
127
+ { rw emb_domain_notin_image_support,
128
+ simp only [not_exists, set.mem_image, to_power_series_symm_apply_coeff, mem_support,
129
+ power_series.coeff_X],
130
+ intro,
131
+ simp [ne.symm hn] {contextual := tt} }
120
132
end
121
133
122
134
end semiring
@@ -181,3 +193,58 @@ is_localization.of_le (submonoid.powers (power_series.X : power_series K)) _
181
193
hahn_series.of_power_series_injective hf)
182
194
183
195
end laurent_series
196
+
197
+ namespace polynomial
198
+
199
+ section laurent_series
200
+
201
+ variables [comm_semiring R] (p q : polynomial R)
202
+
203
+ open polynomial laurent_series hahn_series
204
+
205
+ lemma coe_laurent : (p : laurent_series R) = of_power_series ℤ R p := rfl
206
+
207
+ @[norm_cast] lemma coe_coe : ((p : power_series R) : laurent_series R) = p := rfl
208
+
209
+ @[simp] lemma coe_laurent_zero : ((0 : polynomial R) : laurent_series R) = 0 :=
210
+ by rw [coe_laurent, coe_zero, _root_.map_zero]
211
+
212
+ @[simp] lemma coe_laurent_one : ((1 : polynomial R) : laurent_series R) = 1 :=
213
+ by rw [coe_laurent, coe_one, _root_.map_one]
214
+
215
+ @[norm_cast] lemma coe_laurent_add : ((p + q : polynomial R) : laurent_series R) = p + q :=
216
+ by rw [coe_laurent, coe_add, _root_.map_add, ←coe_laurent, ←coe_laurent]
217
+
218
+ @[norm_cast] lemma coe_laurent_mul : ((p * q : polynomial R) : laurent_series R) = p * q :=
219
+ by rw [coe_laurent, coe_mul, _root_.map_mul, ←coe_laurent, ←coe_laurent]
220
+
221
+ @[norm_cast] lemma coeff_coe_laurent_coe (i : ℕ) :
222
+ ((p : polynomial R) : laurent_series R).coeff i = p.coeff i :=
223
+ by rw [←coe_coe, coeff_coe_power_series, coeff_coe]
224
+
225
+ lemma coeff_coe_laurent (i : ℤ) :
226
+ ((p : polynomial R) : laurent_series R).coeff i = if i < 0 then 0 else p.coeff i.nat_abs :=
227
+ begin
228
+ cases i,
229
+ { rw [int.nat_abs_of_nat_core, int.of_nat_eq_coe, coeff_coe_laurent_coe,
230
+ if_neg (int.coe_nat_nonneg _).not_lt] },
231
+ { rw [coe_laurent, of_power_series_apply, emb_domain_notin_image_support,
232
+ if_pos (int.neg_succ_lt_zero _)],
233
+ simp only [not_exists, rel_embedding.coe_fn_mk, set.mem_image, not_and, coeff_coe,
234
+ function.embedding.coe_fn_mk, ne.def, to_power_series_symm_apply_coeff, mem_support,
235
+ int.nat_cast_eq_coe_nat, int.coe_nat_eq, implies_true_iff, not_false_iff] }
236
+ end
237
+
238
+ @[simp] lemma coe_laurent_C (r : R) : ((C r : polynomial R) : laurent_series R) = hahn_series.C r :=
239
+ by rw [coe_laurent, coe_C, of_power_series_C]
240
+
241
+ @[simp] lemma coe_laurent_X : ((X : polynomial R) : laurent_series R) = single 1 1 :=
242
+ by rw [coe_laurent, coe_X, of_power_series_X]
243
+
244
+ @[norm_cast] lemma coe_laurent_smul (r : R) :
245
+ ((r • p : polynomial R) : laurent_series R) = r • p :=
246
+ by rw [smul_eq_C_mul, coe_laurent_mul, coe_laurent_C, C_mul_eq_smul]
247
+
248
+ end laurent_series
249
+
250
+ end polynomial
0 commit comments