Skip to content
This repository was archived by the owner on Jul 24, 2024. It is now read-only.

Commit 91288e3

Browse files
feat(data/fin/tuple): rename fin.append to matrix.vec_append, introduce a new fin.append and fin.repeat. (#10346)
We already had `fin.append v w h`, which combines the append operation with casting. This commit removes the `h` argument, allowing it to be defeq to `fin.add_cases`, and moves the previous definition to the name `matrix.vec_append` matching `matrix.vec_cons` and similar. Another advantage of dropping `h` is that it is clearer how to state lemmas like `append_assoc`, as we only have one proof argument to keep track of instead of four. As it turns out, to implement a `gmonoid` structure on tuples (under concatenation), `fin.append` without the `h` argument is all that's needed. We implement `matrix.vec_append` in terms of `fin.append`, but provide a lemma that unfolds it to the old definition so as to avoid having to rewrite all the other proofs. Removing `matrix.vec_append` entirely is left to investigate in some future PR. Co-authored-by: ChrisHughes24 <chrishughes24@gmail.com>
1 parent f4ecb59 commit 91288e3

File tree

6 files changed

+196
-73
lines changed

6 files changed

+196
-73
lines changed

src/data/fin/tuple/basic.lean

Lines changed: 106 additions & 11 deletions
Original file line numberDiff line numberDiff line change
@@ -25,6 +25,8 @@ We define the following operations:
2525
* `fin.insert_nth` : insert an element to a tuple at a given position.
2626
* `fin.find p` : returns the first index `n` where `p n` is satisfied, and `none` if it is never
2727
satisfied.
28+
* `fin.append a b` : append two tuples.
29+
* `fin.repeat n a` : repeat a tuple `n` times.
2830
2931
-/
3032
universes u v
@@ -246,17 +248,110 @@ set.ext $ λ y, exists_fin_succ.trans $ eq_comm.or iff.rfl
246248
set.range (fin.cons x b : fin n.succ → α) = insert x (set.range b) :=
247249
by rw [range_fin_succ, cons_zero, tail_cons]
248250

249-
/-- `fin.append ho u v` appends two vectors of lengths `m` and `n` to produce
250-
one of length `o = m + n`. `ho` provides control of definitional equality
251-
for the vector length. -/
252-
def append {α : Type*} {o : ℕ} (ho : o = m + n) (u : fin m → α) (v : fin n → α) : fin o → α :=
253-
λ i, if h : (i : ℕ) < m
254-
then u ⟨i, h⟩
255-
else v ⟨(i : ℕ) - m, (tsub_lt_iff_left (le_of_not_lt h)).2 (ho ▸ i.property)⟩
256-
257-
@[simp] lemma fin_append_apply_zero {α : Type*} {o : ℕ} (ho : (o + 1) = (m + 1) + n)
258-
(u : fin (m + 1) → α) (v : fin n → α) :
259-
fin.append ho u v 0 = u 0 := rfl
251+
section append
252+
253+
/-- Append a tuple of length `m` to a tuple of length `n` to get a tuple of length `m + n`.
254+
This is a non-dependent version of `fin.add_cases`. -/
255+
def append {α : Type*} (a : fin m → α) (b : fin n → α) : fin (m + n) → α :=
256+
@fin.add_cases _ _ (λ _, α) a b
257+
258+
@[simp] lemma append_left {α : Type*} (u : fin m → α) (v : fin n → α) (i : fin m) :
259+
append u v (fin.cast_add n i) = u i :=
260+
add_cases_left _ _ _
261+
262+
@[simp] lemma append_right {α : Type*} (u : fin m → α) (v : fin n → α) (i : fin n) :
263+
append u v (nat_add m i) = v i :=
264+
add_cases_right _ _ _
265+
266+
lemma append_right_nil {α : Type*} (u : fin m → α) (v : fin n → α) (hv : n = 0) :
267+
append u v = u ∘ fin.cast (by rw [hv, add_zero]) :=
268+
begin
269+
refine funext (fin.add_cases (λ l, _) (λ r, _)),
270+
{ rw [append_left, function.comp_apply],
271+
refine congr_arg u (fin.ext _),
272+
simp },
273+
{ exact (fin.cast hv r).elim0' }
274+
end
275+
276+
@[simp] lemma append_elim0' {α : Type*} (u : fin m → α) :
277+
append u fin.elim0' = u ∘ fin.cast (add_zero _) :=
278+
append_right_nil _ _ rfl
279+
280+
lemma append_left_nil {α : Type*} (u : fin m → α) (v : fin n → α) (hu : m = 0) :
281+
append u v = v ∘ fin.cast (by rw [hu, zero_add]) :=
282+
begin
283+
refine funext (fin.add_cases (λ l, _) (λ r, _)),
284+
{ exact (fin.cast hu l).elim0' },
285+
{ rw [append_right, function.comp_apply],
286+
refine congr_arg v (fin.ext _),
287+
simp [hu] },
288+
end
289+
290+
@[simp] lemma elim0'_append {α : Type*} (v : fin n → α) :
291+
append fin.elim0' v = v ∘ fin.cast (zero_add _) :=
292+
append_left_nil _ _ rfl
293+
294+
lemma append_assoc {p : ℕ} {α : Type*} (a : fin m → α) (b : fin n → α) (c : fin p → α) :
295+
append (append a b) c = append a (append b c) ∘ fin.cast (add_assoc _ _ _) :=
296+
begin
297+
ext i,
298+
rw function.comp_apply,
299+
refine fin.add_cases (λ l, _) (λ r, _) i,
300+
{ rw append_left,
301+
refine fin.add_cases (λ ll, _) (λ lr, _) l,
302+
{ rw append_left,
303+
simp [cast_add_cast_add] },
304+
{ rw append_right,
305+
simp [cast_add_nat_add], }, },
306+
{ rw append_right,
307+
simp [←nat_add_nat_add] },
308+
end
309+
310+
end append
311+
312+
section repeat
313+
314+
/-- Repeat `a` `m` times. For example `fin.repeat 2 ![0, 3, 7] = ![0, 3, 7, 0, 3, 7]`. -/
315+
@[simp] def repeat {α : Type*} (m : ℕ) (a : fin n → α) : fin (m * n) → α
316+
| i := a i.mod_nat
317+
318+
@[simp] lemma repeat_zero {α : Type*} (a : fin n → α) :
319+
repeat 0 a = fin.elim0' ∘ cast (zero_mul _) :=
320+
funext $ λ x, (cast (zero_mul _) x).elim0'
321+
322+
@[simp] lemma repeat_one {α : Type*} (a : fin n → α) :
323+
repeat 1 a = a ∘ cast (one_mul _) :=
324+
begin
325+
generalize_proofs h,
326+
apply funext,
327+
rw (fin.cast h.symm).surjective.forall,
328+
intro i,
329+
simp [mod_nat, nat.mod_eq_of_lt i.is_lt],
330+
end
331+
332+
lemma repeat_succ {α : Type*} (a : fin n → α) (m : ℕ) :
333+
repeat m.succ a = append a (repeat m a) ∘ cast ((nat.succ_mul _ _).trans (add_comm _ _)) :=
334+
begin
335+
generalize_proofs h,
336+
apply funext,
337+
rw (fin.cast h.symm).surjective.forall,
338+
refine fin.add_cases (λ l, _) (λ r, _),
339+
{ simp [mod_nat, nat.mod_eq_of_lt l.is_lt], },
340+
{ simp [mod_nat] }
341+
end
342+
343+
@[simp] lemma repeat_add {α : Type*} (a : fin n → α) (m₁ m₂ : ℕ) :
344+
repeat (m₁ + m₂) a = append (repeat m₁ a) (repeat m₂ a) ∘ cast (add_mul _ _ _) :=
345+
begin
346+
generalize_proofs h,
347+
apply funext,
348+
rw (fin.cast h.symm).surjective.forall,
349+
refine fin.add_cases (λ l, _) (λ r, _),
350+
{ simp [mod_nat, nat.mod_eq_of_lt l.is_lt], },
351+
{ simp [mod_nat, nat.add_mod] }
352+
end
353+
354+
end repeat
260355

261356
end tuple
262357

src/data/fin/vec_notation.lean

Lines changed: 42 additions & 16 deletions
Original file line numberDiff line numberDiff line change
@@ -45,8 +45,7 @@ variables {α : Type u}
4545
section matrix_notation
4646

4747
/-- `![]` is the vector with no entries. -/
48-
def vec_empty : fin 0 → α :=
49-
fin_zero_elim
48+
def vec_empty : fin 0 → α := fin.elim0'
5049

5150
/-- `vec_cons h t` prepends an entry `h` to a vector `t`.
5251
@@ -174,16 +173,43 @@ of elements by virtue of the semantics of `bit0` and `bit1` and of
174173
addition on `fin n`).
175174
-/
176175

177-
@[simp] lemma empty_append (v : fin n → α) : fin.append (zero_add _).symm ![] v = v :=
178-
by { ext, simp [fin.append] }
176+
/-- `vec_append ho u v` appends two vectors of lengths `m` and `n` to produce
177+
one of length `o = m + n`. This is a variant of `fin.append` with an additional `ho` argument,
178+
which provides control of definitional equality for the vector length.
179+
180+
This turns out to be helpful when providing simp lemmas to reduce `![a, b, c] n`, and also means
181+
that `vec_append ho u v 0` is valid. `fin.append u v 0` is not valid in this case because there is
182+
no `has_zero (fin (m + n))` instance. -/
183+
def vec_append {α : Type*} {o : ℕ} (ho : o = m + n) (u : fin m → α) (v : fin n → α) : fin o → α :=
184+
fin.append u v ∘ fin.cast ho
185+
186+
lemma vec_append_eq_ite {α : Type*} {o : ℕ} (ho : o = m + n) (u : fin m → α) (v : fin n → α) :
187+
vec_append ho u v = λ i,
188+
if h : (i : ℕ) < m
189+
then u ⟨i, h⟩
190+
else v ⟨(i : ℕ) - m, (tsub_lt_iff_left (le_of_not_lt h)).2 (ho ▸ i.property)⟩ :=
191+
begin
192+
ext i,
193+
rw [vec_append, fin.append, function.comp_apply, fin.add_cases],
194+
congr' with hi,
195+
simp only [eq_rec_constant],
196+
refl,
197+
end
198+
199+
@[simp] lemma vec_append_apply_zero {α : Type*} {o : ℕ} (ho : (o + 1) = (m + 1) + n)
200+
(u : fin (m + 1) → α) (v : fin n → α) :
201+
vec_append ho u v 0 = u 0 := rfl
202+
203+
@[simp] lemma empty_vec_append (v : fin n → α) : vec_append (zero_add _).symm ![] v = v :=
204+
by { ext, simp [vec_append_eq_ite] }
179205

180-
@[simp] lemma cons_append (ho : o + 1 = m + 1 + n) (x : α) (u : fin m → α) (v : fin n → α) :
181-
fin.append ho (vec_cons x u) v =
182-
vec_cons x (fin.append (by rwa [add_assoc, add_comm 1, ←add_assoc,
206+
@[simp] lemma cons_vec_append (ho : o + 1 = m + 1 + n) (x : α) (u : fin m → α) (v : fin n → α) :
207+
vec_append ho (vec_cons x u) v =
208+
vec_cons x (vec_append (by rwa [add_assoc, add_comm 1, ←add_assoc,
183209
add_right_cancel_iff] at ho) u v) :=
184210
begin
185211
ext i,
186-
simp_rw [fin.append],
212+
simp_rw [vec_append_eq_ite],
187213
split_ifs with h,
188214
{ rcases i with ⟨⟨⟩ | i, hi⟩,
189215
{ simp },
@@ -205,10 +231,10 @@ only alternate elements (odd-numbered). -/
205231
def vec_alt1 (hm : m = n + n) (v : fin m → α) (k : fin n) : α :=
206232
v ⟨(k : ℕ) + k + 1, hm.symm ▸ nat.add_succ_lt_add k.property k.property⟩
207233

208-
lemma vec_alt0_append (v : fin n → α) : vec_alt0 rfl (fin.append rfl v v) = v ∘ bit0 :=
234+
lemma vec_alt0_vec_append (v : fin n → α) : vec_alt0 rfl (vec_append rfl v v) = v ∘ bit0 :=
209235
begin
210236
ext i,
211-
simp_rw [function.comp, bit0, vec_alt0, fin.append],
237+
simp_rw [function.comp, bit0, vec_alt0, vec_append_eq_ite],
212238
split_ifs with h; congr,
213239
{ rw fin.coe_mk at h,
214240
simp only [fin.ext_iff, fin.coe_add, fin.coe_mk],
@@ -220,10 +246,10 @@ begin
220246
exact add_lt_add i.property i.property }
221247
end
222248

223-
lemma vec_alt1_append (v : fin (n + 1) → α) : vec_alt1 rfl (fin.append rfl v v) = v ∘ bit1 :=
249+
lemma vec_alt1_vec_append (v : fin (n + 1) → α) : vec_alt1 rfl (vec_append rfl v v) = v ∘ bit1 :=
224250
begin
225251
ext i,
226-
simp_rw [function.comp, vec_alt1, fin.append],
252+
simp_rw [function.comp, vec_alt1, vec_append_eq_ite],
227253
cases n,
228254
{ simp, congr },
229255
{ split_ifs with h; simp_rw [bit1, bit0]; congr,
@@ -248,12 +274,12 @@ end
248274
by simp [vec_head, vec_alt1]
249275

250276
@[simp] lemma cons_vec_bit0_eq_alt0 (x : α) (u : fin n → α) (i : fin (n + 1)) :
251-
vec_cons x u (bit0 i) = vec_alt0 rfl (fin.append rfl (vec_cons x u) (vec_cons x u)) i :=
252-
by rw vec_alt0_append
277+
vec_cons x u (bit0 i) = vec_alt0 rfl (vec_append rfl (vec_cons x u) (vec_cons x u)) i :=
278+
by rw vec_alt0_vec_append
253279

254280
@[simp] lemma cons_vec_bit1_eq_alt1 (x : α) (u : fin n → α) (i : fin (n + 1)) :
255-
vec_cons x u (bit1 i) = vec_alt1 rfl (fin.append rfl (vec_cons x u) (vec_cons x u)) i :=
256-
by rw vec_alt1_append
281+
vec_cons x u (bit1 i) = vec_alt1 rfl (vec_append rfl (vec_cons x u) (vec_cons x u)) i :=
282+
by rw vec_alt1_vec_append
257283

258284
@[simp] lemma cons_vec_alt0 (h : m + 1 + 1 = (n + 1) + (n + 1)) (x y : α) (u : fin m → α) :
259285
vec_alt0 h (vec_cons x (vec_cons y u)) = vec_cons x (vec_alt0

src/linear_algebra/cross_product.lean

Lines changed: 6 additions & 6 deletions
Original file line numberDiff line numberDiff line change
@@ -96,8 +96,8 @@ lemma triple_product_permutation (u v w : fin 3 → R) :
9696
u ⬝ᵥ (v ×₃ w) = v ⬝ᵥ (w ×₃ u) :=
9797
begin
9898
simp only [cross_apply, vec3_dot_product,
99-
matrix.head_cons, matrix.cons_vec_bit0_eq_alt0, matrix.empty_append, matrix.cons_val_one,
100-
matrix.cons_vec_alt0, matrix.cons_append, matrix.cons_val_zero],
99+
matrix.head_cons, matrix.cons_vec_bit0_eq_alt0, matrix.empty_vec_append, matrix.cons_val_one,
100+
matrix.cons_vec_alt0, matrix.cons_vec_append, matrix.cons_val_zero],
101101
ring,
102102
end
103103

@@ -108,17 +108,17 @@ theorem triple_product_eq_det (u v w : fin 3 → R) :
108108
begin
109109
simp only [vec3_dot_product, cross_apply, matrix.det_fin_three,
110110
matrix.head_cons, matrix.cons_vec_bit0_eq_alt0, matrix.empty_vec_alt0, matrix.cons_vec_alt0,
111-
matrix.vec_head_vec_alt0, fin.fin_append_apply_zero, matrix.empty_append, matrix.cons_append,
112-
matrix.cons_val', matrix.cons_val_one, matrix.cons_val_zero],
111+
matrix.vec_head_vec_alt0, matrix.vec_append_apply_zero, matrix.empty_vec_append,
112+
matrix.cons_vec_append, matrix.cons_val', matrix.cons_val_one, matrix.cons_val_zero],
113113
ring,
114114
end
115115

116116
/-- The scalar quadruple product identity, related to the Binet-Cauchy identity. -/
117117
theorem cross_dot_cross (u v w x : fin 3 → R) :
118118
(u ×₃ v) ⬝ᵥ (w ×₃ x) = (u ⬝ᵥ w) * (v ⬝ᵥ x) - (u ⬝ᵥ x) * (v ⬝ᵥ w) :=
119119
begin
120-
simp only [vec3_dot_product, cross_apply, cons_append, cons_vec_bit0_eq_alt0,
121-
cons_val_one, cons_vec_alt0, linear_map.mk₂_apply, cons_val_zero, head_cons, empty_append],
120+
simp only [vec3_dot_product, cross_apply, cons_vec_append, cons_vec_bit0_eq_alt0,
121+
cons_val_one, cons_vec_alt0, linear_map.mk₂_apply, cons_val_zero, head_cons, empty_vec_append],
122122
ring_nf,
123123
end
124124

src/model_theory/encoding.lean

Lines changed: 2 additions & 2 deletions
Original file line numberDiff line numberDiff line change
@@ -74,7 +74,7 @@ begin
7474
(fin_range n).map (option.some ∘ ts) ++ list_decode l,
7575
{ induction (fin_range n) with i l' l'ih,
7676
{ refl },
77-
{ rw [cons_bind, append_assoc, ih, map_cons, l'ih, cons_append] } },
77+
{ rw [cons_bind, list.append_assoc, ih, map_cons, l'ih, cons_append] } },
7878
have h' : ∀ i, (list_decode ((fin_range n).bind (λ (i : fin n), (ts i).list_encode) ++ l)).nth
7979
↑i = some (some (ts i)),
8080
{ intro i,
@@ -268,7 +268,7 @@ begin
268268
rw [list.drop_append_eq_append_drop, length_map, length_fin_range, nat.sub_self, drop,
269269
drop_eq_nil_of_le, nil_append],
270270
rw [length_map, length_fin_range], }, },
271-
{ rw [list_encode, append_assoc, cons_append, list_decode],
271+
{ rw [list_encode, list.append_assoc, cons_append, list_decode],
272272
simp only [subtype.val_eq_coe] at *,
273273
rw [(ih1 _).1, (ih1 _).2, (ih2 _).1, (ih2 _).2, sigma_imp, dif_pos rfl],
274274
exact ⟨rfl, rfl⟩, },

src/number_theory/legendre_symbol/gauss_sum.lean

Lines changed: 1 addition & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -290,7 +290,7 @@ begin
290290
{ ext, congr, apply pow_one },
291291
convert_to (0 + 1 * τ ^ 1 + 0 + (-1) * τ ^ 3 + 0 + (-1) * τ ^ 5 + 0 + 1 * τ ^ 7) ^ 2 = _,
292292
{ simp only [χ₈_apply, matrix.cons_val_zero, matrix.cons_val_one, matrix.head_cons,
293-
matrix.cons_vec_bit0_eq_alt0, matrix.cons_vec_bit1_eq_alt1, matrix.cons_append,
293+
matrix.cons_vec_bit0_eq_alt0, matrix.cons_vec_bit1_eq_alt1, matrix.cons_vec_append,
294294
matrix.cons_vec_alt0, matrix.cons_vec_alt1, int.cast_zero, int.cast_one, int.cast_neg,
295295
zero_mul], refl },
296296
convert_to 8 + (τ ^ 4 + 1) * (τ ^ 10 - 2 * τ ^ 8 - 2 * τ ^ 6 + 6 * τ ^ 4 + τ ^ 2 - 8) = _,

0 commit comments

Comments
 (0)