@@ -66,7 +66,7 @@ noncomputable theory
66
66
universe u
67
67
68
68
open function set submodule
69
- open_locale classical big_operators
69
+ open_locale big_operators
70
70
71
71
variables {ι : Type *} {ι' : Type *} {R : Type *} {R₂ : Type *} {K : Type *}
72
72
variables {M : Type *} {M' M'' : Type *} {V : Type u} {V' : Type *}
@@ -367,6 +367,7 @@ by rw [coe_reindex, range_reindex']
367
367
368
368
/-- `b.reindex_range` is a basis indexed by `range b`, the basis vectors themselves. -/
369
369
def reindex_range : basis (range b) R M :=
370
+ by haveI := classical.dec (nontrivial R); exact
370
371
if h : nontrivial R then
371
372
by letI := h; exact b.reindex (equiv.of_injective b (basis.injective b))
372
373
else
@@ -419,7 +420,7 @@ b.reindex_range_repr' _ rfl
419
420
420
421
section fintype
421
422
422
- variables [fintype ι]
423
+ variables [fintype ι] [decidable_eq M]
423
424
424
425
/-- `b.reindex_finset_range` is a basis indexed by `finset.univ.image b`,
425
426
the finite set of basis vectors themselves. -/
@@ -750,11 +751,13 @@ linear_equiv.trans b.repr
750
751
..finsupp.equiv_fun_on_finite } : (ι →₀ R) ≃ₗ[R] (ι → R))
751
752
752
753
/-- 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
755
757
756
- theorem module.card_fintype [fintype R] [fintype M] :
758
+ theorem module.card_fintype (b : basis ι R M) [fintype R] [fintype M] :
757
759
card M = (card R) ^ (card ι) :=
760
+ by classical; exact
758
761
calc card M = card (ι → R) : card_congr b.equiv_fun.to_equiv
759
762
... = card R ^ card ι : card_fun
760
763
@@ -781,7 +784,8 @@ lemma basis.sum_repr (u : M) : ∑ i, b.repr u i • b i = u :=
781
784
b.sum_equiv_fun u
782
785
783
786
@[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 :=
785
789
by { rw [b.equiv_fun_apply, b.repr_self_apply] }
786
790
787
791
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
802
806
@[simp] lemma basis.of_equiv_fun_repr_apply (e : M ≃ₗ[R] (ι → R)) (x : M) (i : ι) :
803
807
(basis.of_equiv_fun e).repr x i = e x i := rfl
804
808
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)) :
806
810
(basis.of_equiv_fun e : ι → M) = λ i, e.symm (function.update 0 i 1 ) :=
807
811
funext $ λ i, e.injective $ funext $ λ j,
808
812
by simp [basis.of_equiv_fun, ←finsupp.single_eq_pi_single, finsupp.single_eq_update]
809
813
810
814
@[simp] lemma basis.of_equiv_fun_equiv_fun
811
815
(v : basis ι R M) : basis.of_equiv_fun v.equiv_fun = v :=
812
816
begin
817
+ classical,
813
818
ext j,
814
819
simp only [basis.equiv_fun_symm_apply, basis.coe_of_equiv_fun],
815
820
simp_rw [function.update_apply, ite_smul],
@@ -971,7 +976,7 @@ by simp [hli.repr_eq_single j, h]
971
976
972
977
/-- Given a basis, the `i`th element of the dual basis evaluates to the Kronecker delta on the
973
978
`j`th element of the basis. -/
974
- lemma mk_coord_apply {i j : ι} :
979
+ lemma mk_coord_apply [decidable_eq ι] {i j : ι} :
975
980
(basis.mk hli hsp).coord i (v j) = if j = i then 1 else 0 :=
976
981
begin
977
982
cases eq_or_ne j i,
@@ -1058,6 +1063,7 @@ mk_apply
1058
1063
@[simp] lemma coord_units_smul (e : basis ι R₂ M) (w : ι → R₂ˣ) (i : ι) :
1059
1064
(e.units_smul w).coord i = (w i)⁻¹ • e.coord i :=
1060
1065
begin
1066
+ classical,
1061
1067
apply e.ext,
1062
1068
intros j,
1063
1069
transitivity ((e.units_smul w).coord i) ((w j)⁻¹ • (e.units_smul w) j),
@@ -1227,7 +1233,9 @@ let s := set.range v,
1227
1233
b := hs.to_subtype_range.extend (subset_univ (set.range v)) in
1228
1234
(basis.extend hs.to_subtype_range).reindex $ equiv.symm $
1229
1235
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 _)
1231
1239
1232
1240
lemma subset_extend {s : set V} (hs : linear_independent K (coe : s → V)) :
1233
1241
s ⊆ hs.extend (set.subset_univ _) :=
@@ -1275,7 +1283,8 @@ variables (K V)
1275
1283
1276
1284
theorem vector_space.card_fintype [fintype K] [fintype V] :
1277
1285
∃ 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)⟩
1279
1288
1280
1289
section atoms_of_submodule_lattice
1281
1290
0 commit comments