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

Commit f2edd79

Browse files
committed
fix(linear_algebra/basis): add missing decidable arguments in lemmas (#18252)
The resulting lemmas are syntactically more general. To ensure that this catches all of them, this also removes `open_locale classical` from the file, replacing it with manual construction of classical decidability instances within definitions, and `classical` within proofs.
1 parent d6fad0e commit f2edd79

File tree

3 files changed

+21
-15
lines changed

3 files changed

+21
-15
lines changed

src/analysis/normed_space/pi_Lp.lean

Lines changed: 1 addition & 2 deletions
Original file line numberDiff line numberDiff line change
@@ -698,8 +698,7 @@ basis.of_equiv_fun (pi_Lp.linear_equiv p 𝕜 (λ _ : ι, 𝕜))
698698

699699
@[simp] lemma basis_fun_apply [decidable_eq ι] (i) :
700700
basis_fun p 𝕜 ι i = (pi_Lp.equiv p _).symm (pi.single i 1) :=
701-
by { simp_rw [basis_fun, basis.coe_of_equiv_fun, pi_Lp.linear_equiv_symm_apply, pi.single],
702-
congr /- Get rid of a `decidable_eq` mismatch. -/ }
701+
by simp_rw [basis_fun, basis.coe_of_equiv_fun, pi_Lp.linear_equiv_symm_apply, pi.single]
703702

704703
@[simp] lemma basis_fun_repr (x : pi_Lp p (λ i : ι, 𝕜)) (i : ι) :
705704
(basis_fun p 𝕜 ι).repr x i = x i :=

src/linear_algebra/basis.lean

Lines changed: 19 additions & 10 deletions
Original file line numberDiff line numberDiff line change
@@ -66,7 +66,7 @@ noncomputable theory
6666
universe u
6767

6868
open function set submodule
69-
open_locale classical big_operators
69+
open_locale big_operators
7070

7171
variables {ι : Type*} {ι' : Type*} {R : Type*} {R₂ : Type*} {K : Type*}
7272
variables {M : Type*} {M' M'' : Type*} {V : Type u} {V' : Type*}
@@ -367,6 +367,7 @@ by rw [coe_reindex, range_reindex']
367367

368368
/-- `b.reindex_range` is a basis indexed by `range b`, the basis vectors themselves. -/
369369
def reindex_range : basis (range b) R M :=
370+
by haveI := classical.dec (nontrivial R); exact
370371
if h : nontrivial R then
371372
by letI := h; exact b.reindex (equiv.of_injective b (basis.injective b))
372373
else
@@ -419,7 +420,7 @@ b.reindex_range_repr' _ rfl
419420

420421
section fintype
421422

422-
variables [fintype ι]
423+
variables [fintype ι] [decidable_eq M]
423424

424425
/-- `b.reindex_finset_range` is a basis indexed by `finset.univ.image b`,
425426
the finite set of basis vectors themselves. -/
@@ -750,11 +751,13 @@ linear_equiv.trans b.repr
750751
..finsupp.equiv_fun_on_finite } : (ι →₀ R) ≃ₗ[R] (ι → R))
751752

752753
/-- A module over a finite ring that admits a finite basis is finite. -/
753-
def module.fintype_of_fintype [fintype R] : fintype M :=
754-
fintype.of_equiv _ b.equiv_fun.to_equiv.symm
754+
def module.fintype_of_fintype (b : basis ι R M) [fintype R] : fintype M :=
755+
by haveI := classical.dec_eq ι; exact
756+
fintype.of_equiv _ b.equiv_fun.to_equiv.symm
755757

756-
theorem module.card_fintype [fintype R] [fintype M] :
758+
theorem module.card_fintype (b : basis ι R M) [fintype R] [fintype M] :
757759
card M = (card R) ^ (card ι) :=
760+
by classical; exact
758761
calc card M = card (ι → R) : card_congr b.equiv_fun.to_equiv
759762
... = card R ^ card ι : card_fun
760763

@@ -781,7 +784,8 @@ lemma basis.sum_repr (u : M) : ∑ i, b.repr u i • b i = u :=
781784
b.sum_equiv_fun u
782785

783786
@[simp]
784-
lemma basis.equiv_fun_self (i j : ι) : b.equiv_fun (b i) j = if i = j then 1 else 0 :=
787+
lemma basis.equiv_fun_self [decidable_eq ι] (i j : ι) :
788+
b.equiv_fun (b i) j = if i = j then 1 else 0 :=
785789
by { rw [b.equiv_fun_apply, b.repr_self_apply] }
786790

787791
lemma basis.repr_sum_self (c : ι → R) : ⇑(b.repr (∑ i, c i • b i)) = c :=
@@ -802,14 +806,15 @@ basis.of_repr $ e.trans $ linear_equiv.symm $ finsupp.linear_equiv_fun_on_finite
802806
@[simp] lemma basis.of_equiv_fun_repr_apply (e : M ≃ₗ[R] (ι → R)) (x : M) (i : ι) :
803807
(basis.of_equiv_fun e).repr x i = e x i := rfl
804808

805-
@[simp] lemma basis.coe_of_equiv_fun (e : M ≃ₗ[R] (ι → R)) :
809+
@[simp] lemma basis.coe_of_equiv_fun [decidable_eq ι] (e : M ≃ₗ[R] (ι → R)) :
806810
(basis.of_equiv_fun e : ι → M) = λ i, e.symm (function.update 0 i 1) :=
807811
funext $ λ i, e.injective $ funext $ λ j,
808812
by simp [basis.of_equiv_fun, ←finsupp.single_eq_pi_single, finsupp.single_eq_update]
809813

810814
@[simp] lemma basis.of_equiv_fun_equiv_fun
811815
(v : basis ι R M) : basis.of_equiv_fun v.equiv_fun = v :=
812816
begin
817+
classical,
813818
ext j,
814819
simp only [basis.equiv_fun_symm_apply, basis.coe_of_equiv_fun],
815820
simp_rw [function.update_apply, ite_smul],
@@ -971,7 +976,7 @@ by simp [hli.repr_eq_single j, h]
971976

972977
/-- Given a basis, the `i`th element of the dual basis evaluates to the Kronecker delta on the
973978
`j`th element of the basis. -/
974-
lemma mk_coord_apply {i j : ι} :
979+
lemma mk_coord_apply [decidable_eq ι] {i j : ι} :
975980
(basis.mk hli hsp).coord i (v j) = if j = i then 1 else 0 :=
976981
begin
977982
cases eq_or_ne j i,
@@ -1058,6 +1063,7 @@ mk_apply
10581063
@[simp] lemma coord_units_smul (e : basis ι R₂ M) (w : ι → R₂ˣ) (i : ι) :
10591064
(e.units_smul w).coord i = (w i)⁻¹ • e.coord i :=
10601065
begin
1066+
classical,
10611067
apply e.ext,
10621068
intros j,
10631069
transitivity ((e.units_smul w).coord i) ((w j)⁻¹ • (e.units_smul w) j),
@@ -1227,7 +1233,9 @@ let s := set.range v,
12271233
b := hs.to_subtype_range.extend (subset_univ (set.range v)) in
12281234
(basis.extend hs.to_subtype_range).reindex $ equiv.symm $
12291235
calc ι ⊕ (b \ s : set V) ≃ s ⊕ (b \ s : set V) : equiv.sum_congr e (equiv.refl _)
1230-
... ≃ b : equiv.set.sum_diff_subset (hs.to_subtype_range.subset_extend _)
1236+
... ≃ b :
1237+
by haveI := classical.dec_pred (∈ s); exact
1238+
equiv.set.sum_diff_subset (hs.to_subtype_range.subset_extend _)
12311239

12321240
lemma subset_extend {s : set V} (hs : linear_independent K (coe : s → V)) :
12331241
s ⊆ hs.extend (set.subset_univ _) :=
@@ -1275,7 +1283,8 @@ variables (K V)
12751283

12761284
theorem vector_space.card_fintype [fintype K] [fintype V] :
12771285
∃ n : ℕ, card V = (card K) ^ n :=
1278-
⟨card (basis.of_vector_space_index K V), module.card_fintype (basis.of_vector_space K V)⟩
1286+
by classical; exact
1287+
⟨card (basis.of_vector_space_index K V), module.card_fintype (basis.of_vector_space K V)⟩
12791288

12801289
section atoms_of_submodule_lattice
12811290

src/linear_algebra/std_basis.lean

Lines changed: 1 addition & 3 deletions
Original file line numberDiff line numberDiff line change
@@ -34,7 +34,6 @@ this is a basis over `fin 3 → R`.
3434

3535
open function submodule
3636
open_locale big_operators
37-
open_locale big_operators
3837

3938
namespace linear_map
4039

@@ -247,8 +246,7 @@ basis.of_equiv_fun (linear_equiv.refl _ _)
247246
@[simp] lemma basis_fun_apply [decidable_eq η] (i) :
248247
basis_fun R η i = std_basis R (λ (i : η), R) i 1 :=
249248
by { simp only [basis_fun, basis.coe_of_equiv_fun, linear_equiv.refl_symm,
250-
linear_equiv.refl_apply, std_basis_apply],
251-
congr /- Get rid of a `decidable_eq` mismatch. -/ }
249+
linear_equiv.refl_apply, std_basis_apply] }
252250

253251
@[simp] lemma basis_fun_repr (x : η → R) (i : η) :
254252
(pi.basis_fun R η).repr x i = x i :=

0 commit comments

Comments
 (0)