Skip to content

Commit

Permalink
feat(linear_algebra/affine_space/barycentric_coords): affine bases ex…
Browse files Browse the repository at this point in the history
…ist over fields (#10333)

Formalized as part of the Sphere Eversion project.
  • Loading branch information
ocfnash committed Nov 23, 2021
1 parent 7a6e6d8 commit a1338d6
Showing 1 changed file with 34 additions and 0 deletions.
34 changes: 34 additions & 0 deletions src/linear_algebra/affine_space/barycentric_coords.lean
Expand Up @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE.
Authors: Oliver Nash
-/
import linear_algebra.affine_space.independent
import linear_algebra.affine_space.finite_dimensional
import linear_algebra.determinant

/-!
Expand Down Expand Up @@ -265,4 +266,37 @@ end

end comm_ring

section field

-- TODO Relax `field` to `division_ring` (results are still true)
variables [field k] [module k V]
include V

variables (k V P)

lemma exists_affine_basis : ∃ (s : set P), nonempty (affine_basis ↥s k P) :=
begin
obtain ⟨s, -, h_tot, h_ind⟩ := exists_affine_independent k V (set.univ : set P),
refine ⟨s, ⟨⟨(coe : s → P), h_ind, _⟩⟩⟩,
rw [subtype.range_coe, h_tot, affine_subspace.span_univ],
end

variables {k V P}

lemma exists_affine_basis_of_finite_dimensional {ι : Type*} [fintype ι] [finite_dimensional k V]
(h : fintype.card ι = finite_dimensional.finrank k V + 1) :
nonempty (affine_basis ι k P) :=
begin
obtain ⟨s, ⟨⟨incl, h_ind, h_tot⟩⟩⟩ := affine_basis.exists_affine_basis k V P,
haveI : fintype s := fintype_of_fin_dim_affine_independent k h_ind,
have hs : fintype.card ι = fintype.card s,
{ rw h, exact (h_ind.affine_span_eq_top_iff_card_eq_finrank_add_one.mp h_tot).symm, },
rw ← affine_independent_equiv (fintype.equiv_of_card_eq hs) at h_ind,
refine ⟨⟨_, h_ind, _⟩⟩,
rw range_comp,
simp [h_tot],
end

end field

end affine_basis

0 comments on commit a1338d6

Please sign in to comment.