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

Commit 3d1f16a

Browse files
committed
feat(analysis/normed_space/multilinear): define mk_pi_algebra (#4316)
I'm going to use this definition for converting `(mv_)power_series` to `formal_multilinear_series`.
1 parent 1362701 commit 3d1f16a

File tree

3 files changed

+212
-26
lines changed

3 files changed

+212
-26
lines changed

src/analysis/normed_space/multilinear.lean

Lines changed: 109 additions & 8 deletions
Original file line numberDiff line numberDiff line change
@@ -539,6 +539,112 @@ begin
539539
exact mul_nonneg (norm_nonneg _) (pow_nonneg (norm_nonneg _) _)
540540
end
541541

542+
section
543+
544+
variables (𝕜 ι) (A : Type*) [normed_comm_ring A] [normed_algebra 𝕜 A]
545+
546+
/-- The continuous multilinear map on `A^ι`, where `A` is a normed commutative algebra
547+
over `𝕜`, associating to `m` the product of all the `m i`.
548+
549+
See also `continuous_multilinear_map.mk_pi_algebra_fin`. -/
550+
protected def mk_pi_algebra : continuous_multilinear_map 𝕜 (λ i : ι, A) A :=
551+
@multilinear_map.mk_continuous 𝕜 ι (λ i : ι, A) A _ _ _ _ _ _ _
552+
(multilinear_map.mk_pi_algebra 𝕜 ι A) (if nonempty ι then 1 else ∥(1 : A)∥) $
553+
begin
554+
intro m,
555+
by_cases hι : nonempty ι,
556+
{ resetI, simp [hι, norm_prod_le' univ univ_nonempty] },
557+
{ simp [eq_empty_of_not_nonempty hι univ, hι] }
558+
end
559+
560+
variables {A 𝕜 ι}
561+
562+
@[simp] lemma mk_pi_algebra_apply (m : ι → A) :
563+
continuous_multilinear_map.mk_pi_algebra 𝕜 ι A m = ∏ i, m i :=
564+
rfl
565+
566+
lemma norm_mk_pi_algebra_le [nonempty ι] :
567+
∥continuous_multilinear_map.mk_pi_algebra 𝕜 ι A∥ ≤ 1 :=
568+
calc ∥continuous_multilinear_map.mk_pi_algebra 𝕜 ι A∥ ≤ if nonempty ι then 1 else ∥(1 : A)∥ :
569+
multilinear_map.mk_continuous_norm_le _ (by split_ifs; simp [zero_le_one]) _
570+
... = _ : if_pos ‹_›
571+
572+
lemma norm_mk_pi_algebra_of_empty (h : ¬nonempty ι) :
573+
∥continuous_multilinear_map.mk_pi_algebra 𝕜 ι A∥ = ∥(1 : A)∥ :=
574+
begin
575+
apply le_antisymm,
576+
calc ∥continuous_multilinear_map.mk_pi_algebra 𝕜 ι A∥ ≤ if nonempty ι then 1 else ∥(1 : A)∥ :
577+
multilinear_map.mk_continuous_norm_le _ (by split_ifs; simp [zero_le_one]) _
578+
... = ∥(1 : A)∥ : if_neg ‹_›,
579+
convert ratio_le_op_norm _ (λ _, 1); [skip, apply_instance],
580+
simp [eq_empty_of_not_nonempty h univ]
581+
end
582+
583+
@[simp] lemma norm_mk_pi_algebra [norm_one_class A] :
584+
∥continuous_multilinear_map.mk_pi_algebra 𝕜 ι A∥ = 1 :=
585+
begin
586+
by_cases hι : nonempty ι,
587+
{ resetI,
588+
refine le_antisymm norm_mk_pi_algebra_le _,
589+
convert ratio_le_op_norm _ (λ _, 1); [skip, apply_instance],
590+
simp },
591+
{ simp [norm_mk_pi_algebra_of_empty hι] }
592+
end
593+
594+
end
595+
596+
section
597+
598+
variables (𝕜 n) (A : Type*) [normed_ring A] [normed_algebra 𝕜 A]
599+
600+
/-- The continuous multilinear map on `A^n`, where `A` is a normed algebra over `𝕜`, associating to
601+
`m` the product of all the `m i`.
602+
603+
See also: `multilinear_map.mk_pi_algebra`. -/
604+
protected def mk_pi_algebra_fin : continuous_multilinear_map 𝕜 (λ i : fin n, A) A :=
605+
@multilinear_map.mk_continuous 𝕜 (fin n) (λ i : fin n, A) A _ _ _ _ _ _ _
606+
(multilinear_map.mk_pi_algebra_fin 𝕜 n A) (nat.cases_on n ∥(1 : A)∥ (λ _, 1)) $
607+
begin
608+
intro m,
609+
cases n,
610+
{ simp },
611+
{ have : @list.of_fn A n.succ m ≠ [] := by simp,
612+
simpa [← fin.prod_of_fn] using list.norm_prod_le' this }
613+
end
614+
615+
variables {A 𝕜 n}
616+
617+
@[simp] lemma mk_pi_algebra_fin_apply (m : fin n → A) :
618+
continuous_multilinear_map.mk_pi_algebra_fin 𝕜 n A m = (list.of_fn m).prod :=
619+
rfl
620+
621+
lemma norm_mk_pi_algebra_fin_succ_le :
622+
∥continuous_multilinear_map.mk_pi_algebra_fin 𝕜 n.succ A∥ ≤ 1 :=
623+
multilinear_map.mk_continuous_norm_le _ zero_le_one _
624+
625+
lemma norm_mk_pi_algebra_fin_le_of_pos (hn : 0 < n) :
626+
∥continuous_multilinear_map.mk_pi_algebra_fin 𝕜 n A∥ ≤ 1 :=
627+
by cases n; [exact hn.false.elim, exact norm_mk_pi_algebra_fin_succ_le]
628+
629+
lemma norm_mk_pi_algebra_fin_zero :
630+
∥continuous_multilinear_map.mk_pi_algebra_fin 𝕜 0 A∥ = ∥(1 : A)∥ :=
631+
begin
632+
refine le_antisymm (multilinear_map.mk_continuous_norm_le _ (norm_nonneg _) _) _,
633+
convert ratio_le_op_norm _ (λ _, 1); [simp, apply_instance]
634+
end
635+
636+
lemma norm_mk_pi_algebra_fin [norm_one_class A] :
637+
∥continuous_multilinear_map.mk_pi_algebra_fin 𝕜 n A∥ = 1 :=
638+
begin
639+
cases n,
640+
{ simp [norm_mk_pi_algebra_fin_zero] },
641+
{ refine le_antisymm norm_mk_pi_algebra_fin_succ_le _,
642+
convert ratio_le_op_norm _ (λ _, 1); [skip, apply_instance],
643+
simp }
644+
end
645+
646+
end
647+
542648
variables (𝕜 ι)
543649

544650
/-- The canonical continuous multilinear map on `𝕜^ι`, associating to `m` the product of all the
@@ -553,14 +659,9 @@ variables {𝕜 ι}
553659
@[simp] lemma mk_pi_field_apply (z : E₂) (m : ι → 𝕜) :
554660
(continuous_multilinear_map.mk_pi_field 𝕜 ι z : (ι → 𝕜) → E₂) m = (∏ i, m i) • z := rfl
555661

556-
lemma mk_pi_ring_apply_one_eq_self (f : continuous_multilinear_map 𝕜 (λ(i : ι), 𝕜) E₂) :
662+
lemma mk_pi_field_apply_one_eq_self (f : continuous_multilinear_map 𝕜 (λ(i : ι), 𝕜) E₂) :
557663
continuous_multilinear_map.mk_pi_field 𝕜 ι (f (λi, 1)) = f :=
558-
begin
559-
ext m,
560-
have : m = (λi, m i • 1), by { ext j, simp },
561-
conv_rhs { rw [this, f.map_smul_univ] },
562-
refl
563-
end
664+
to_multilinear_map_inj f.to_multilinear_map.mk_pi_ring_apply_one_eq_self
564665

565666
variables (𝕜 ι E₂)
566667

@@ -575,7 +676,7 @@ protected def pi_field_equiv_aux : E₂ ≃ₗ[𝕜] (continuous_multilinear_map
575676
map_add' := λ z z', by { ext m, simp [smul_add] },
576677
map_smul' := λ c z, by { ext m, simp [smul_smul, mul_comm] },
577678
left_inv := λ z, by simp,
578-
right_inv := λ f, f.mk_pi_ring_apply_one_eq_self }
679+
right_inv := λ f, f.mk_pi_field_apply_one_eq_self }
579680

580681
/-- Continuous multilinear maps on `𝕜^n` with values in `E₂` are in bijection with `E₂`, as such a
581682
continuous multilinear map is completely determined by its value on the constant vector made of

src/linear_algebra/multilinear.lean

Lines changed: 97 additions & 17 deletions
Original file line numberDiff line numberDiff line change
@@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE.
44
Authors: Sébastien Gouëzel
55
-/
66
import linear_algebra.basic
7+
import algebra.algebra.basic
78
import tactic.omega
89
import data.fintype.sort
910

@@ -414,6 +415,29 @@ end apply_sum
414415

415416
end semiring
416417

418+
end multilinear_map
419+
420+
namespace linear_map
421+
variables [semiring R] [Πi, add_comm_monoid (M₁ i)] [add_comm_monoid M₂] [add_comm_monoid M₃]
422+
[∀i, semimodule R (M₁ i)] [semimodule R M₂] [semimodule R M₃]
423+
424+
/-- Composing a multilinear map with a linear map gives again a multilinear map. -/
425+
def comp_multilinear_map (g : M₂ →ₗ[R] M₃) (f : multilinear_map R M₁ M₂) :
426+
multilinear_map R M₁ M₃ :=
427+
{ to_fun := g ∘ f,
428+
map_add' := λ m i x y, by simp,
429+
map_smul' := λ m i c x, by simp }
430+
431+
@[simp] lemma coe_comp_multilinear_map (g : M₂ →ₗ[R] M₃) (f : multilinear_map R M₁ M₂) :
432+
⇑(g.comp_multilinear_map f) = g ∘ f := rfl
433+
434+
lemma comp_multilinear_map_apply (g : M₂ →ₗ[R] M₃) (f : multilinear_map R M₁ M₂) (m : Π i, M₁ i) :
435+
g.comp_multilinear_map f m = g (f m) := rfl
436+
437+
end linear_map
438+
439+
namespace multilinear_map
440+
417441
section comm_semiring
418442

419443
variables [comm_semiring R] [∀i, add_comm_monoid (M₁ i)] [∀i, add_comm_monoid (M i)] [add_comm_monoid M₂]
@@ -450,15 +474,83 @@ instance : has_scalar R (multilinear_map R M₁ M₂) := ⟨λ c f,
450474

451475
@[simp] lemma smul_apply (c : R) (m : Πi, M₁ i) : (c • f) m = c • f m := rfl
452476

477+
section
478+
479+
variables (R ι) (A : Type*) [comm_semiring A] [algebra R A] [fintype ι]
480+
481+
/-- Given an `R`-algebra `A`, `mk_pi_algebra` is the multilinear map on `A^ι` associating
482+
to `m` the product of all the `m i`.
483+
484+
See also `multilinear_map.mk_pi_algebra_fin` for a version that works with a non-commutative
485+
algebra `A` but requires `ι = fin n`. -/
486+
protected def mk_pi_algebra : multilinear_map R (λ i : ι, A) A :=
487+
{ to_fun := λ m, ∏ i, m i,
488+
map_add' := λ m i x y, by simp [finset.prod_update_of_mem, add_mul],
489+
map_smul' := λ m i c x, by simp [finset.prod_update_of_mem] }
490+
491+
variables {R A ι}
492+
493+
@[simp] lemma mk_pi_algebra_apply (m : ι → A) :
494+
multilinear_map.mk_pi_algebra R ι A m = ∏ i, m i :=
495+
rfl
496+
497+
end
498+
499+
section
500+
501+
variables (R n) (A : Type*) [semiring A] [algebra R A]
502+
503+
/-- Given an `R`-algebra `A`, `mk_pi_algebra_fin` is the multilinear map on `A^n` associating
504+
to `m` the product of all the `m i`.
505+
506+
See also `multilinear_map.mk_pi_algebra` for a version that assumes `[comm_semiring A]` but works
507+
for `A^ι` with any finite type `ι`. -/
508+
protected def mk_pi_algebra_fin : multilinear_map R (λ i : fin n, A) A :=
509+
{ to_fun := λ m, (list.of_fn m).prod,
510+
map_add' :=
511+
begin
512+
intros m i x y,
513+
have : (list.fin_range n).index_of i < n,
514+
by simpa using list.index_of_lt_length.2 (list.mem_fin_range i),
515+
simp [list.of_fn_eq_map, (list.nodup_fin_range n).map_update, list.prod_update_nth, add_mul,
516+
this, mul_add, add_mul]
517+
end,
518+
map_smul' :=
519+
begin
520+
intros m i c x,
521+
have : (list.fin_range n).index_of i < n,
522+
by simpa using list.index_of_lt_length.2 (list.mem_fin_range i),
523+
simp [list.of_fn_eq_map, (list.nodup_fin_range n).map_update, list.prod_update_nth, this]
524+
end }
525+
526+
variables {R A n}
527+
528+
@[simp] lemma mk_pi_algebra_fin_apply (m : fin n → A) :
529+
multilinear_map.mk_pi_algebra_fin R n A m = (list.of_fn m).prod :=
530+
rfl
531+
532+
lemma mk_pi_algebra_fin_apply_const (a : A) :
533+
multilinear_map.mk_pi_algebra_fin R n A (λ _, a) = a ^ n :=
534+
by simp
535+
536+
end
537+
538+
/-- Given an `R`-multilinear map `f` taking values in `R`, `f.smul_right z` is the map
539+
sending `m` to `f m • z`. -/
540+
def smul_right (f : multilinear_map R M₁ R) (z : M₂) : multilinear_map R M₁ M₂ :=
541+
(linear_map.smul_right linear_map.id z).comp_multilinear_map f
542+
543+
@[simp] lemma smul_right_apply (f : multilinear_map R M₁ R) (z : M₂) (m : Π i, M₁ i) :
544+
f.smul_right z m = f m • z :=
545+
rfl
546+
453547
variables (R ι)
454548

455549
/-- The canonical multilinear map on `R^ι` when `ι` is finite, associating to `m` the product of
456-
all the `m i` (multiplied by a fixed reference element `z` in the target module) -/
550+
all the `m i` (multiplied by a fixed reference element `z` in the target module). See also
551+
`mk_pi_algebra` for a more general version. -/
457552
protected def mk_pi_ring [fintype ι] (z : M₂) : multilinear_map R (λ(i : ι), R) M₂ :=
458-
{ to_fun := λm, (∏ i, m i) • z,
459-
map_add' := λ m i x y, by simp [finset.prod_update_of_mem, add_mul, add_smul],
460-
map_smul' := λ m i c x, by { rw [smul_eq_mul],
461-
simp [finset.prod_update_of_mem, smul_smul, mul_assoc] } }
553+
(multilinear_map.mk_pi_algebra R ι R).smul_right z
462554

463555
variables {R ι}
464556

@@ -529,18 +621,6 @@ end comm_ring
529621

530622
end multilinear_map
531623

532-
namespace linear_map
533-
variables [ring R] [∀i, add_comm_group (M₁ i)] [add_comm_group M₂] [add_comm_group M₃]
534-
[∀i, module R (M₁ i)] [module R M₂] [module R M₃]
535-
536-
/-- Composing a multilinear map with a linear map gives again a multilinear map. -/
537-
def comp_multilinear_map (g : M₂ →ₗ[R] M₃) (f : multilinear_map R M₁ M₂) : multilinear_map R M₁ M₃ :=
538-
{ to_fun := λ m, g (f m),
539-
map_add' := λ m i x y, by simp,
540-
map_smul' := λ m i c x, by simp }
541-
542-
end linear_map
543-
544624
section currying
545625
/-!
546626
### Currying

src/topology/algebra/multilinear.lean

Lines changed: 6 additions & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -70,8 +70,13 @@ instance : has_coe_to_fun (continuous_multilinear_map R M₁ M₂) :=
7070

7171
@[simp] lemma coe_coe : (f.to_multilinear_map : (Π i, M₁ i) → M₂) = f := rfl
7272

73+
theorem to_multilinear_map_inj :
74+
function.injective (continuous_multilinear_map.to_multilinear_map :
75+
continuous_multilinear_map R M₁ M₂ → multilinear_map R M₁ M₂)
76+
| ⟨f, hf⟩ ⟨g, hg⟩ rfl := rfl
77+
7378
@[ext] theorem ext {f f' : continuous_multilinear_map R M₁ M₂} (H : ∀ x, f x = f' x) : f = f' :=
74-
by { cases f, cases f', congr' with x, exact H x }
79+
to_multilinear_map_inj $ multilinear_map.ext H
7580

7681
@[simp] lemma map_add (m : Πi, M₁ i) (i : ι) (x y : M₁ i) :
7782
f (update m i (x + y)) = f (update m i x) + f (update m i y) :=

0 commit comments

Comments
 (0)