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

Commit c37e3e7

Browse files
committed
refactor(field_theory/intermediate_field): remove old_structure_cmd (#9620)
Co-authored-by: Scott Morrison <scott.morrison@gmail.com>
1 parent b39feab commit c37e3e7

File tree

1 file changed

+8
-11
lines changed

1 file changed

+8
-11
lines changed

src/field_theory/intermediate_field.lean

Lines changed: 8 additions & 11 deletions
Original file line numberDiff line numberDiff line change
@@ -42,27 +42,24 @@ open_locale big_operators
4242

4343
variables (K L : Type*) [field K] [field L] [algebra K L]
4444

45-
section
46-
set_option old_structure_cmd true
47-
4845
/-- `S : intermediate_field K L` is a subset of `L` such that there is a field
4946
tower `L / S / K`. -/
50-
structure intermediate_field extends subalgebra K L, subfield L
47+
structure intermediate_field extends subalgebra K L :=
48+
(neg_mem' : ∀ x ∈ carrier, -x ∈ carrier)
49+
(inv_mem' : ∀ x ∈ carrier, x⁻¹ ∈ carrier)
5150

5251
/-- Reinterpret an `intermediate_field` as a `subalgebra`. -/
5352
add_decl_doc intermediate_field.to_subalgebra
5453

55-
/-- Reinterpret an `intermediate_field` as a `subfield`. -/
56-
add_decl_doc intermediate_field.to_subfield
57-
58-
end
59-
6054
variables {K L} (S : intermediate_field K L)
6155

6256
namespace intermediate_field
6357

58+
/-- Reinterpret an `intermediate_field` as a `subfield`. -/
59+
def to_subfield : subfield L := { ..S.to_subalgebra, ..S }
60+
6461
instance : set_like (intermediate_field K L) L :=
65-
intermediate_field.carrier, λ p q h, by cases p; cases q; congr'
62+
λ S, S.to_subalgebra.carrier, by { rintros ⟨⟨⟩⟩ ⟨⟨⟩⟩ ⟨h⟩, congr, }
6663

6764
@[simp]
6865
lemma mem_carrier {s : intermediate_field K L} {x : L} : x ∈ s.carrier ↔ x ∈ s := iff.rfl
@@ -77,7 +74,7 @@ set_like.ext h
7774

7875
@[simp] lemma mem_mk (s : set L) (hK : ∀ x, algebra_map K L x ∈ s)
7976
(ho hm hz ha hn hi) (x : L) :
80-
x ∈ intermediate_field.mk s ho hm hz ha hK hn hi ↔ x ∈ s := iff.rfl
77+
x ∈ intermediate_field.mk (subalgebra.mk s ho hm hz ha hK) hn hi ↔ x ∈ s := iff.rfl
8178

8279
@[simp] lemma mem_to_subalgebra (s : intermediate_field K L) (x : L) :
8380
x ∈ s.to_subalgebra ↔ x ∈ s := iff.rfl

0 commit comments

Comments
 (0)