From f807f7b108cb420dab4d97d93c483e2c98ff73c4 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Xavier-Fran=C3=A7ois=20Roblot?= <46200072+xroblot@users.noreply.github.com> Date: Fri, 28 Apr 2023 17:46:02 +0000 Subject: [PATCH] feat LinearAlgebra.Basis: add basis.restrictScalars (#3707) Mathlib4 version of https://github.com/leanprover-community/mathlib/pull/18814 [`linear_algebra.basis`@`2f4cdce0c2f2f3b8cd58f05d556d03b468e1eb2e`..`04cdee31e196e30f507e8e9eb2d06e02c9ff6310`](https://leanprover-community.github.io/mathlib-port-status/file/linear_algebra/basis?range=2f4cdce0c2f2f3b8cd58f05d556d03b468e1eb2e..04cdee31e196e30f507e8e9eb2d06e02c9ff6310)) --- Mathlib/LinearAlgebra/Basis.lean | 57 +++++++++++++++++++++++++++++++- 1 file changed, 56 insertions(+), 1 deletion(-) diff --git a/Mathlib/LinearAlgebra/Basis.lean b/Mathlib/LinearAlgebra/Basis.lean index 40404b0ef8ce78..2ae5e7ffcf99fd 100644 --- a/Mathlib/LinearAlgebra/Basis.lean +++ b/Mathlib/LinearAlgebra/Basis.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Johannes Hölzl, Mario Carneiro, Alexander Bentkamp ! This file was ported from Lean 3 source module linear_algebra.basis -! leanprover-community/mathlib commit 2f4cdce0c2f2f3b8cd58f05d556d03b468e1eb2e +! leanprover-community/mathlib commit 04cdee31e196e30f507e8e9eb2d06e02c9ff6310 ! Please do not edit these lines, except to modify the commit id ! if you have ported upstream changes. -/ @@ -1649,3 +1649,58 @@ theorem quotient_prod_linearEquiv (p : Submodule K V) : Nonempty (((V ⧸ p) × #align quotient_prod_linear_equiv quotient_prod_linearEquiv end DivisionRing + +section RestrictScalars + +variable {S : Type _} [CommRing R] [Ring S] [Nontrivial S] [AddCommGroup M] + +variable [Algebra R S] [Module S M] [Module R M] + +variable [IsScalarTower R S M] [NoZeroSMulDivisors R S] (b : Basis ι S M) + +variable (R) + +open Submodule + +set_option synthInstance.etaExperiment true in +/-- Let `b` be a `S`-basis of `M`. Let `R` be a CommRing such that `Algebra R S` has no zero smul +divisors, then the submodule of `M` spanned by `b` over `R` admits `b` as a `R`-basis. -/ +noncomputable def Basis.restrictScalars : Basis ι R (span R (Set.range b)) := + Basis.span (b.linearIndependent.restrict_scalars (smul_left_injective R one_ne_zero)) +#align basis.restrict_scalars Basis.restrictScalars + +@[simp] +theorem Basis.restrictScalars_apply (i : ι) : (b.restrictScalars R i : M) = b i := by + simp only [Basis.restrictScalars, Basis.span_apply] +#align basis.restrict_scalars_apply Basis.restrictScalars_apply + +set_option synthInstance.etaExperiment true in +@[simp] +theorem Basis.restrictScalars_repr_apply (m : span R (Set.range b)) (i : ι) : + algebraMap R S ((b.restrictScalars R).repr m i) = b.repr m i := by + suffices + Finsupp.mapRange.linearMap (Algebra.linearMap R S) ∘ₗ (b.restrictScalars R).repr.toLinearMap = + ((b.repr : M →ₗ[S] ι →₀ S).restrictScalars R).domRestrict _ + by exact FunLike.congr_fun (LinearMap.congr_fun this m) i + refine Basis.ext (b.restrictScalars R) fun _ => ?_ + simp only [LinearMap.coe_comp, LinearEquiv.coe_toLinearMap, Function.comp_apply, map_one, + Basis.repr_self, Finsupp.mapRange.linearMap_apply, Finsupp.mapRange_single, + Algebra.linearMap_apply, LinearMap.domRestrict_apply, LinearEquiv.coe_coe, + Basis.restrictScalars_apply, LinearMap.coe_restrictScalars] +#align basis.restrict_scalars_repr_apply Basis.restrictScalars_repr_apply + +/-- Let `b` be a `S`-basis of `M`. Then `m : M` lies in the `R`-module spanned by `b` iff all the +coordinates of `m` on the basis `b` are in `R` (see `Basis.mem_span` for the case `R = S`). -/ +theorem Basis.mem_span_iff_repr_mem (m : M) : + m ∈ span R (Set.range b) ↔ ∀ i, b.repr m i ∈ Set.range (algebraMap R S) := by + refine + ⟨fun hm i => ⟨(b.restrictScalars R).repr ⟨m, hm⟩ i, b.restrictScalars_repr_apply R ⟨m, hm⟩ i⟩, + fun h => ?_⟩ + rw [← b.total_repr m, Finsupp.total_apply S _] + refine sum_mem fun i _ => ?_ + obtain ⟨_, h⟩ := h i + simp_rw [← h, algebraMap_smul] + exact smul_mem _ _ (subset_span (Set.mem_range_self i)) +#align basis.mem_span_iff_repr_mem Basis.mem_span_iff_repr_mem + +end RestrictScalars