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

Commit 019ead1

Browse files
committed
chore(linear_algebra): remove finite_dimensional import from finsupp_vector_space (#17474)
This PR removes `linear_algebra/finite_dimensional.lean` from the imports of `linear_algebra/finsupp_vector_space.lean`. It turns out the only results that needed this import can already be done in `linear_algebra/finite_dimensional.lean` so I just moved it over. The purpose of this change is to insert `linear_algebra/free_module/finite.lean` in between `linear_algebra/finrank.lean` and `linear_algebra/finite_dimensional.lean`, as discussed in the module docs for `linear_algebra/free_module/finite.lean` and in #17401.
1 parent 2efabeb commit 019ead1

File tree

6 files changed

+36
-31
lines changed

6 files changed

+36
-31
lines changed

src/field_theory/mv_polynomial.lean

Lines changed: 1 addition & 0 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: Johannes Hölzl
55
-/
66

7+
import data.mv_polynomial.comm_ring
78
import ring_theory.mv_polynomial.basic
89

910
/-!

src/linear_algebra/finite_dimensional.lean

Lines changed: 30 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -1497,3 +1497,33 @@ end
14971497

14981498
end End
14991499
end module
1500+
1501+
section module
1502+
1503+
open module
1504+
1505+
open_locale cardinal
1506+
1507+
lemma cardinal_mk_eq_cardinal_mk_field_pow_dim
1508+
(K V : Type u) [field K] [add_comm_group V] [module K V] [finite_dimensional K V] :
1509+
#V = #K ^ module.rank K V :=
1510+
begin
1511+
let s := basis.of_vector_space_index K V,
1512+
let hs := basis.of_vector_space K V,
1513+
calc #V = #(s →₀ K) : quotient.sound ⟨hs.repr.to_equiv⟩
1514+
... = #(s → K) : quotient.sound ⟨finsupp.equiv_fun_on_fintype⟩
1515+
... = _ : by rw [← cardinal.lift_inj.1 hs.mk_eq_dim, cardinal.power_def]
1516+
end
1517+
1518+
lemma cardinal_lt_aleph_0_of_finite_dimensional
1519+
(K V : Type u) [field K] [add_comm_group V] [module K V]
1520+
[_root_.finite K] [finite_dimensional K V] :
1521+
#V < ℵ₀ :=
1522+
begin
1523+
letI : is_noetherian K V := is_noetherian.iff_fg.2 infer_instance,
1524+
rw cardinal_mk_eq_cardinal_mk_field_pow_dim K V,
1525+
exact cardinal.power_lt_aleph_0 (cardinal.lt_aleph_0_of_finite K)
1526+
(is_noetherian.dim_lt_aleph_0 K V),
1527+
end
1528+
1529+
end module

src/linear_algebra/finsupp_vector_space.lean

Lines changed: 1 addition & 30 deletions
Original file line numberDiff line numberDiff line change
@@ -5,7 +5,6 @@ Authors: Johannes Hölzl
55
-/
66

77
import linear_algebra.dimension
8-
import linear_algebra.finite_dimensional
98
import linear_algebra.std_basis
109

1110
/-!
@@ -14,8 +13,7 @@ import linear_algebra.std_basis
1413
This file contains results on the `R`-module structure on functions of finite support from a type
1514
`ι` to an `R`-module `M`, in particular in the case that `R` is a field.
1615
17-
Furthermore, it contains some facts about isomorphisms of vector spaces from equality of dimension
18-
as well as the cardinality of finite dimensional vector spaces.
16+
Furthermore, it contains some facts about isomorphisms of vector spaces from equality of dimension.
1917
2018
## TODO
2119
@@ -184,33 +182,6 @@ end
184182

185183
end module
186184

187-
section module
188-
189-
open module
190-
191-
variables (K V : Type u) [field K] [add_comm_group V] [module K V]
192-
193-
lemma cardinal_mk_eq_cardinal_mk_field_pow_dim [finite_dimensional K V] :
194-
#V = #K ^ module.rank K V :=
195-
begin
196-
let s := basis.of_vector_space_index K V,
197-
let hs := basis.of_vector_space K V,
198-
calc #V = #(s →₀ K) : quotient.sound ⟨hs.repr.to_equiv⟩
199-
... = #(s → K) : quotient.sound ⟨finsupp.equiv_fun_on_fintype⟩
200-
... = _ : by rw [← cardinal.lift_inj.1 hs.mk_eq_dim, cardinal.power_def]
201-
end
202-
203-
lemma cardinal_lt_aleph_0_of_finite_dimensional [_root_.finite K] [finite_dimensional K V] :
204-
#V < ℵ₀ :=
205-
begin
206-
letI : is_noetherian K V := is_noetherian.iff_fg.2 infer_instance,
207-
rw cardinal_mk_eq_cardinal_mk_field_pow_dim K V,
208-
exact cardinal.power_lt_aleph_0 (cardinal.lt_aleph_0_of_finite K)
209-
(is_noetherian.dim_lt_aleph_0 K V),
210-
end
211-
212-
end module
213-
214185
namespace basis
215186

216187
variables {R M n : Type*}

src/ring_theory/localization/module.lean

Lines changed: 1 addition & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -3,7 +3,7 @@ Copyright (c) 2022 Anne Baanen. All rights reserved.
33
Released under Apache 2.0 license as described in the file LICENSE.
44
Authors: Junyan Xu, Anne Baanen
55
-/
6-
import linear_algebra.linear_independent
6+
import linear_algebra.basis
77
import ring_theory.localization.fraction_ring
88
import ring_theory.localization.integer
99

src/ring_theory/mv_polynomial/basic.lean

Lines changed: 2 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -5,6 +5,8 @@ Authors: Johannes Hölzl
55
-/
66

77
import algebra.char_p.basic
8+
import data.polynomial.algebra_map
9+
import data.mv_polynomial.variables
810
import linear_algebra.finsupp_vector_space
911

1012
/-!

src/ring_theory/tensor_product.lean

Lines changed: 1 addition & 0 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: Scott Morrison, Johan Commelin
55
-/
66

7+
import linear_algebra.finite_dimensional
78
import linear_algebra.tensor_product_basis
89
import ring_theory.adjoin.basic
910

0 commit comments

Comments
 (0)