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

Commit 3aac028

Browse files
committed
feat(field_theory/intermediate_field): equalities from inclusions and dimension bounds (#4828)
Co-authored-by: tb65536 <tb65536@users.noreply.github.com>
1 parent 6ec7aec commit 3aac028

File tree

2 files changed

+46
-5
lines changed

2 files changed

+46
-5
lines changed

src/field_theory/intermediate_field.lean

Lines changed: 37 additions & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -5,7 +5,7 @@ Authors: Anne Baanen
55
-/
66

77
import field_theory.subfield
8-
import ring_theory.algebra_tower
8+
import field_theory.tower
99

1010
/-!
1111
# Intermediate fields
@@ -42,6 +42,7 @@ A `subalgebra` is closed under all operations except `⁻¹`,
4242
intermediate field, field extension
4343
-/
4444

45+
open finite_dimensional
4546
open_locale big_operators
4647

4748
variables (K L : Type*) [field K] [field L] [algebra K L]
@@ -286,4 +287,39 @@ instance has_lift2 {F : intermediate_field K L} :
286287

287288
end tower
288289

290+
section finite_dimensional
291+
292+
instance finite_dimensional_left [finite_dimensional K L] (F : intermediate_field K L) :
293+
finite_dimensional K F :=
294+
finite_dimensional.finite_dimensional_submodule F.to_subalgebra.to_submodule
295+
296+
instance finite_dimensional_right [finite_dimensional K L] (F : intermediate_field K L) :
297+
finite_dimensional F L :=
298+
right K F L
299+
300+
lemma eq_of_le_of_findim_le [finite_dimensional K L] {F E : intermediate_field K L} (h_le : F ≤ E)
301+
(h_findim : findim K E ≤ findim K F) : F = E :=
302+
intermediate_field.ext'_iff.mpr (submodule.ext'_iff.mp (eq_of_le_of_findim_le
303+
(show F.to_subalgebra.to_submodule ≤ E.to_subalgebra.to_submodule, by exact h_le) h_findim))
304+
305+
lemma eq_of_le_of_findim_eq [finite_dimensional K L] {F E : intermediate_field K L} (h_le : F ≤ E)
306+
(h_findim : findim K F = findim K E) : F = E :=
307+
eq_of_le_of_findim_le h_le h_findim.ge
308+
309+
lemma eq_of_le_of_findim_le' [finite_dimensional K L] {F E : intermediate_field K L} (h_le : F ≤ E)
310+
(h_findim : findim F L ≤ findim E L) : F = E :=
311+
begin
312+
apply eq_of_le_of_findim_le h_le,
313+
have h1 := findim_mul_findim K F L,
314+
have h2 := findim_mul_findim K E L,
315+
have h3 : 0 < findim E L := findim_pos,
316+
nlinarith,
317+
end
318+
319+
lemma eq_of_le_of_findim_eq' [finite_dimensional K L] {F E : intermediate_field K L} (h_le : F ≤ E)
320+
(h_findim : findim F L = findim E L) : F = E :=
321+
eq_of_le_of_findim_le' h_le h_findim.le
322+
323+
end finite_dimensional
324+
289325
end intermediate_field

src/linear_algebra/finite_dimensional.lean

Lines changed: 9 additions & 4 deletions
Original file line numberDiff line numberDiff line change
@@ -656,14 +656,19 @@ end linear_equiv
656656

657657
namespace finite_dimensional
658658

659+
lemma eq_of_le_of_findim_le {S₁ S₂ : submodule K V} [finite_dimensional K S₂] (hle : S₁ ≤ S₂)
660+
(hd : findim K S₂ ≤ findim K S₁) : S₁ = S₂ :=
661+
begin
662+
rw ←linear_equiv.findim_eq (submodule.comap_subtype_equiv_of_le hle) at hd,
663+
exact le_antisymm hle (submodule.comap_subtype_eq_top.1 (eq_top_of_findim_eq
664+
(le_antisymm (comap (submodule.subtype S₂) S₁).findim_le hd))),
665+
end
666+
659667
/-- If a submodule is less than or equal to a finite-dimensional
660668
submodule with the same dimension, they are equal. -/
661669
lemma eq_of_le_of_findim_eq {S₁ S₂ : submodule K V} [finite_dimensional K S₂] (hle : S₁ ≤ S₂)
662670
(hd : findim K S₁ = findim K S₂) : S₁ = S₂ :=
663-
begin
664-
rw ←linear_equiv.findim_eq (submodule.comap_subtype_equiv_of_le hle) at hd,
665-
exact le_antisymm hle (submodule.comap_subtype_eq_top.1 (eq_top_of_findim_eq hd))
666-
end
671+
eq_of_le_of_findim_le hle hd.ge
667672

668673
end finite_dimensional
669674

0 commit comments

Comments
 (0)