From 40e74d71fffb24cf01d1073fd4ea8fa27d20f3c0 Mon Sep 17 00:00:00 2001 From: Chris Hughes <33847686+ChrisHughes24@users.noreply.github.com> Date: Sat, 3 Aug 2019 21:04:35 +0200 Subject: [PATCH 01/51] chore(ring_theory/algebra): make first type argument explicit in alg_hom MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Now this works, and it didn't work previously even with `@` ```lean structure alg_equiv (α β γ : Type*) [comm_ring α] [ring β] [ring γ] [algebra α β] [algebra α γ] extends alg_hom α β γ := ``` --- src/ring_theory/algebra.lean | 4 ++-- 1 file changed, 2 insertions(+), 2 deletions(-) diff --git a/src/ring_theory/algebra.lean b/src/ring_theory/algebra.lean index 358fbcde7c630..ef36075705122 100644 --- a/src/ring_theory/algebra.lean +++ b/src/ring_theory/algebra.lean @@ -144,12 +144,12 @@ variables {R A} end algebra /-- Defining the homomorphism in the category R-Alg. -/ -structure alg_hom {R : Type u} (A : Type v) (B : Type w) +structure alg_hom (R : Type u) (A : Type v) (B : Type w) [comm_ring R] [ring A] [ring B] [algebra R A] [algebra R B] := (to_fun : A → B) [hom : is_ring_hom to_fun] (commutes' : ∀ r : R, to_fun (algebra_map A r) = algebra_map B r) -infixr ` →ₐ `:25 := alg_hom +infixr ` →ₐ `:25 := alg_hom _ notation A ` →ₐ[`:25 R `] ` B := @alg_hom R A B _ _ _ _ _ namespace alg_hom From 790c892a6c822bd0784da13cc9cc610282a6eb45 Mon Sep 17 00:00:00 2001 From: Chris Hughes <33847686+ChrisHughes24@users.noreply.github.com> Date: Sat, 3 Aug 2019 21:09:53 +0200 Subject: [PATCH 02/51] Update algebra.lean --- src/ring_theory/algebra.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/src/ring_theory/algebra.lean b/src/ring_theory/algebra.lean index ef36075705122..63693d21384fd 100644 --- a/src/ring_theory/algebra.lean +++ b/src/ring_theory/algebra.lean @@ -150,7 +150,7 @@ structure alg_hom (R : Type u) (A : Type v) (B : Type w) (commutes' : ∀ r : R, to_fun (algebra_map A r) = algebra_map B r) infixr ` →ₐ `:25 := alg_hom _ -notation A ` →ₐ[`:25 R `] ` B := @alg_hom R A B _ _ _ _ _ +notation A ` →ₐ[`:25 R `] ` B := alg_hom R A B namespace alg_hom From c0633b5afabf2cfcc0dab3ccdcb6083842c85891 Mon Sep 17 00:00:00 2001 From: Chris Hughes Date: Sat, 3 Aug 2019 22:46:45 +0200 Subject: [PATCH 03/51] feat(field_theory/algebraic_closure) --- src/field_theory/algebraic_closure.lean | 557 ++++++++++++++++++++++++ 1 file changed, 557 insertions(+) create mode 100644 src/field_theory/algebraic_closure.lean diff --git a/src/field_theory/algebraic_closure.lean b/src/field_theory/algebraic_closure.lean new file mode 100644 index 0000000000000..0c2cf19f6296c --- /dev/null +++ b/src/field_theory/algebraic_closure.lean @@ -0,0 +1,557 @@ +/- +Copyright (c) 2019 Chris Hughes. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Chris Hughes +-/ + +import ring_theory.adjoin_root +import algebra.direct_limit +import set_theory.schroeder_bernstein +import ring_theory.integral_closure + +universes u v w +open polynomial zorn set function +variables {K : Type u} [discrete_field K] +noncomputable theory + +/- Turn down the instance priority for subtype.decidable_eq and use classical.dec_eq everywhere, + to avoid diamonds -/ +local attribute [instance, priority 0] subtype.decidable_eq + +lemma injective_eq {α : Sort*} : injective (eq : α → α → Prop) := +λ _ _ h, h.symm ▸ rfl + +section minimal_polynomial +/- To be moved -/ +variables {α : Type u} {β : Type v} [discrete_field α] [discrete_field β] [algebra α β] + +def minimal_polynomial {x : β} (hx : is_integral α x) : polynomial α := sorry + +lemma minimal_polynomial_irreducible {x : β} (hx : is_integral α x) : + irreducible (minimal_polynomial hx) := sorry + +lemma minimal_polynomial_monic {x : β} (hx : is_integral α x) : monic (minimal_polynomial hx) := sorry + +@[simp] lemma aeval_minimal_polynomial {x : β} (hx : is_integral α x) : + aeval α β x (minimal_polynomial hx) = 0 := sorry + +lemma root_minimal_polynomial {x : β} (hx : is_integral α x) {y : α} + (h : (minimal_polynomial hx).eval y = 0) : algebra_map β y = x := sorry + +end minimal_polynomial + +@[instance] lemma equiv.is_ring_hom {α β : Type*} [ring β] (e : α ≃ β) : + @is_ring_hom β α _ (equiv.ring e) e.symm := +by split; simp [equiv.mul_def, equiv.add_def, equiv.one_def] + +instance equiv.is_ring_hom.symm {α β : Type*} [ring β] (e : α ≃ β) : + @is_ring_hom α β (equiv.ring e) _ e := +by split; simp [equiv.mul_def, equiv.add_def, equiv.one_def] + +def equiv.ring_equiv {α β : Type*} [ring β] (e : α ≃ β) : + @ring_equiv α β (equiv.ring e) _ := +{ hom := by apply_instance, ..e } + +lemma exists_root_of_equiv {α β : Type*} [comm_ring α] [comm_ring β] [decidable_eq α] + [decidable_eq β] (e : α ≃r β) {f : polynomial α} {x : β} (hx : f.eval₂ e.to_equiv x = 0) : + f.eval (e.symm.to_equiv x) = 0 := +begin + letI : is_ring_hom e.to_equiv := e.hom, + rw [← e.to_equiv.injective.eq_iff, + ← eval₂_hom e.to_equiv, ring_equiv.to_equiv_symm, + equiv.apply_symm_apply, is_ring_hom.map_zero e.to_equiv, hx], +end + +set_option old_structure_cmd true + +structure alg_equiv (α β γ : Type*) [comm_ring α] [ring β] [ring γ] + [algebra α β] [algebra α γ] extends alg_hom α β γ, equiv β γ + +set_option old_structure_cmd false + +infix ` ≃ₐ `:25 := alg_equiv _ +notation A ` ≃ₐ[`:25 R `] ` B := alg_equiv R A B + +namespace alg_equiv +variables {α : Type u} {β : Type v} {γ : Type w} [comm_ring α] [ring β] [ring γ] + [algebra α β] [algebra α γ] + +protected def symm (e : β ≃ₐ[α] γ) : γ ≃ₐ[α] β := sorry + +end alg_equiv + +section thing + +local attribute [instance] classical.dec + +private lemma thing_aux {X : Type u} {Y : Type v} {Z : Type w} (fxy : X ↪ Y) (fxz : X ↪ Z) + (hYZ : (Z ↪ Y) → false) : ↥-range fxy.1 ↪ ↥-range fxz.1 := +classical.choice $ or.resolve_left embedding.total $ + λ ⟨f⟩, hYZ $ + calc Z ↪ range fxz ⊕ ↥-range fxz : + (equiv.set.sum_compl _).symm.to_embedding + ... ↪ range fxy ⊕ ↥-range fxy : + embedding.sum_congr + (((equiv.set.range _ fxz.2).symm.to_embedding).trans + (equiv.set.range _ fxy.2).to_embedding) + f + ... ↪ Y : (equiv.set.sum_compl _).to_embedding + +private def thing {X : Type u} {Y : Type v} {Z : Type w} (fxy : X ↪ Y) (fxz : X ↪ Z) + (hYZ : (Z ↪ Y) → false) : Y ↪ Z := +calc Y ↪ range fxy ⊕ ↥-range fxy : (equiv.set.sum_compl _).symm.to_embedding +... ↪ range fxz ⊕ ↥-range fxz : embedding.sum_congr + ((equiv.set.range _ fxy.2).symm.to_embedding.trans + (equiv.set.range _ fxz.2).to_embedding) + (thing_aux fxy fxz hYZ) +... ↪ Z : (equiv.set.sum_compl _).to_embedding + +private lemma thing_commutes {X : Type u} {Y : Type v} {Z : Type w} (fxy : X ↪ Y) (fxz : X ↪ Z) + (hYZ : (Z ↪ Y) → false) (x : X) : thing fxy fxz hYZ (fxy x) = fxz x := +have (⟨fxy x, mem_range_self _⟩ : range fxy) = equiv.set.range _ fxy.2 x, from rfl, +begin + dsimp only [thing, embedding.trans_apply, equiv.trans_apply, function.comp, + equiv.to_embedding_coe_fn], + simp only [equiv.set.sum_compl_symm_apply_of_mem (mem_range_self _), + embedding.sum_congr_apply_inl, equiv.set.sum_compl_apply_inl, + embedding.trans_apply, equiv.to_embedding_coe_fn, this, equiv.symm_apply_apply], + refl +end + +end thing + +class is_algebraically_closed (K : Type u) [nonzero_comm_ring K] [decidable_eq K] := +(exists_root : ∀ f : polynomial K, 0 < degree f → ∃ x, is_root f x) + +section is_algebraically_closed + +lemma is_algebraically_closed_of_irreducible_has_root + (h : ∀ f : polynomial K, irreducible f → ∃ x, is_root f x) : + is_algebraically_closed K := +⟨λ f hf0, let ⟨g, hg⟩ := is_noetherian_ring.exists_irreducible_factor + (show ¬ is_unit f, from λ h, by rw [is_unit_iff_degree_eq_zero] at h; + rw h at hf0; exact lt_irrefl _ hf0) + (λ h, by rw ← degree_eq_bot at h; + rw h at hf0; exact absurd hf0 dec_trivial) in + let ⟨x, hx⟩ := h g hg.1 in + let ⟨i, hi⟩ := hg.2 in + ⟨x, by rw [hi, is_root.def, eval_mul, show _ = _, from hx, zero_mul]⟩⟩ + +-- /- An algebraic extension of -/ +-- lemma equiv_of_algebraic + +end is_algebraically_closed + +namespace algebraic_closure + +section classical + +local attribute [instance, priority 1] classical.dec + +/-- The `big_type` with cardinality strictly larger than any algebraic extension -/ +private def big_type (K : Type u) [discrete_field K] := set (ℕ × polynomial K) + +private def algebraic_embedding_aux {L : Type*} [discrete_field L] [algebra K L] + (h : ∀ l : L, is_integral K l) (x : L) : ℕ × polynomial K := +let f := classical.some (h x) in +⟨list.index_of x (quotient.out ((f.map (algebra_map L)).roots.1)), f⟩ + +private lemma algebraic_embedding_aux_injective + {L : Type*} [discrete_field L] [algebra K L] + (h : ∀ l : L, is_integral K l) : injective (algebraic_embedding_aux h) := +λ x y hxy, +let f := classical.some (h x) in +let g := classical.some (h y) in +have hf : monic f ∧ aeval K L x f = 0, from classical.some_spec (h x), +have hg : monic g ∧ aeval K L y g = 0, from classical.some_spec (h y), +have hfg : f = g, from (prod.ext_iff.1 hxy).2, +have hfg' : list.index_of x (quotient.out ((f.map (algebra_map L)).roots.1)) = + list.index_of y (quotient.out ((f.map (algebra_map L)).roots.1)), + from (prod.ext_iff.1 hxy).1.trans (hfg.symm ▸ rfl), +have hx : x ∈ quotient.out ((f.map (algebra_map L)).roots.1), + from multiset.mem_coe.1 begin + show x ∈ quotient.mk _, + rw [quotient.out_eq, ← finset.mem_def, mem_roots (mt (map_eq_zero (algebra_map L)).1 + (ne_zero_of_monic hf.1)), is_root.def, eval_map, ← aeval_def, hf.2], + end, +have hy : y ∈ quotient.out ((g.map (algebra_map L)).roots.1), + from multiset.mem_coe.1 begin + show y ∈ quotient.mk _, + rw [quotient.out_eq, ← finset.mem_def, mem_roots (mt (map_eq_zero (algebra_map L)).1 + (ne_zero_of_monic hg.1)), is_root.def, eval_map, ← aeval_def, hg.2], + end, +(list.index_of_inj hx (by rwa hfg)).1 hfg' + +private def algebraic_embedding_big_type {L : Type*} [discrete_field L] [algebra K L] + (h : ∀ l : L, is_integral K l) : L ↪ big_type K := +⟨_, injective_comp injective_eq $ algebraic_embedding_aux_injective h⟩ + +private def algebraic_embedding {L : Type*} [discrete_field L] [algebra K L] + (h : ∀ l : L, is_integral K l) : L ↪ ℕ × polynomial K := +⟨_, algebraic_embedding_aux_injective h⟩ + +private def bembedding (K : Type u) [discrete_field K] : K ↪ big_type K := +⟨λ a, show set _, from {(0, X - C a)}, λ a b, by simp [C_inj]⟩ + +instance range_bembedding.discrete_field : discrete_field (set.range (bembedding K)) := +equiv.discrete_field (equiv.set.range _ (bembedding K).2).symm + +private structure extension (K : Type u) [discrete_field K] : Type u := +(carrier : set (big_type K)) +[field : discrete_field ↥carrier] +[algebra : algebra K ↥carrier] +(algebraic : ∀ x : carrier, is_integral K x) + +attribute [instance] extension.field extension.algebra + +private def base_extension (K : Type u) [discrete_field K] : extension K := +{ carrier := set.range (bembedding K), + algebra := algebra.of_ring_hom (equiv.set.range _ (bembedding K).2).symm.symm + (by apply_instance), + algebraic := sorry } --a field is algebraic over an isomorphic field + +/-- not used but might help woth sorries -/ +private def extension.of_algebraic {L : Type v} [discrete_field L] [algebra K L] + (hL : ∀ x : L, is_integral K x) : extension K := +{ carrier := set.range (algebraic_embedding_big_type hL), + field := equiv.discrete_field (equiv.set.range _ (algebraic_embedding_big_type hL).2).symm, + algebra := sorry, -- a field isomorphic to an algebra is an algebra + algebraic := sorry -- a field isomorphic to an algebraic extension is algebraic + } + +instance : preorder (extension K) := +{ le := λ L M, ∃ hLM : L.carrier ⊆ M.carrier, is_ring_hom (inclusion hLM), + le_refl := λ _, ⟨set.subset.refl _, by convert is_ring_hom.id; ext; simp⟩, + le_trans := λ L M N ⟨hLM₁, hLM₂⟩ ⟨hMN₁, hMN₂⟩, ⟨set.subset.trans hLM₁ hMN₁, + by resetI; convert is_ring_hom.comp (inclusion hLM₁) (inclusion hMN₁)⟩ } + +private structure chain' (c : set (extension K)) : Prop := +(chain : chain (≤) c) + +local attribute [class] chain' + +private lemma is_chain (c : set (extension K)) [chain' c]: chain (≤) c := +chain'.chain (by apply_instance) + +section chain + +variables (c : set (extension K)) [hcn : nonempty c] +include c hcn + +variable [hcn' : chain' c] +include hcn' + +instance chain_directed_order : directed_order c := +⟨λ ⟨i, hi⟩ ⟨j, hj⟩, let ⟨k, hkc, hk⟩ := chain.directed_on + (is_chain c) i hi j hj in ⟨⟨k, hkc⟩, hk⟩⟩ + +private def chain_map (i j : c) (hij : i ≤ j) : i.1.carrier → j.1.carrier := +inclusion (exists.elim hij (λ h _, h)) + +instance chain_field_hom (i j : c) (hij : i ≤ j) : is_field_hom (chain_map c i j hij) := +exists.elim hij (λ _, id) + +instance chain_directed_system : directed_system (λ i : c, i.1.carrier) (chain_map c) := +by split; intros; simp [chain_map] + +private def chain_limit : Type u := ring.direct_limit (λ i : c, i.1.carrier) (chain_map c) + +private lemma of_eq_of (x : big_type K) (i j : c) (hi : x ∈ i.1.carrier) (hj : x ∈ j.1.carrier) : + ring.direct_limit.of (λ i : c, i.1.carrier) (chain_map c) i ⟨x, hi⟩ = + ring.direct_limit.of (λ i : c, i.1.carrier) (chain_map c) j ⟨x, hj⟩ := +have hij : i ≤ j ∨ j ≤ i, + from show i.1 ≤ j.1 ∨ j.1 ≤ i.1, from chain.total (is_chain c) i.2 j.2, +hij.elim + (λ hij, begin + rw ← @ring.direct_limit.of_f c _ _ _ (λ i : c, i.1.carrier) _ _ (chain_map c) _ + _ _ _ hij, + simp [chain_map, inclusion] + end) + (λ hij, begin + rw ← @ring.direct_limit.of_f c _ _ _ (λ i : c, i.1.carrier) _ _ (chain_map c) _ + _ _ _ hij, + simp [chain_map, inclusion] + end) + +private lemma injective_aux (i j : c) + (x y : ⋃ i : c, i.1.carrier) (hx : x.1 ∈ i.1.carrier) (hy : y.1 ∈ j.1.carrier) : + ring.direct_limit.of (λ i : c, i.1.carrier) (chain_map c) i ⟨x, hx⟩ = + ring.direct_limit.of (λ i : c, i.1.carrier) (chain_map c) j ⟨y, hy⟩ → + x = y := +have hij : i ≤ j ∨ j ≤ i, + from show i.1 ≤ j.1 ∨ j.1 ≤ i.1, from chain.total (is_chain c) i.2 j.2, +have hinj : ∀ (i j : c) (hij : i ≤ j), injective (chain_map c i j hij), + from λ _ _ _, is_field_hom.injective _, +hij.elim + (λ hij h, begin + rw ← @ring.direct_limit.of_f c _ _ _ (λ i : c, i.1.carrier) _ _ (chain_map c) _ + _ _ _ hij at h, + simpa [chain_map, inclusion, subtype.coe_ext.symm] using ring.direct_limit.of_inj hinj j h, + end) + (λ hji h, begin + rw ← @ring.direct_limit.of_f c _ _ _ (λ i : c, i.1.carrier) _ _ (chain_map c) _ + _ _ _ hji at h, + simpa [chain_map, inclusion, subtype.coe_ext.symm] using ring.direct_limit.of_inj hinj i h, + end) + +private def equiv_direct_limit : (⋃ (i : c), i.1.carrier) ≃ + ring.direct_limit (λ i : c, i.1.carrier) (chain_map c) := +@equiv.of_bijective (⋃ i : c, i.1.carrier) + (ring.direct_limit (λ i : c, i.1.carrier) (chain_map c)) + (λ x, ring.direct_limit.of _ _ (classical.some (set.mem_Union.1 x.2)) + ⟨_, classical.some_spec (set.mem_Union.1 x.2)⟩) + ⟨λ x y, injective_aux _ _ _ _ _ _ _, + λ x, let ⟨i, ⟨y, hy⟩, hy'⟩ := ring.direct_limit.exists_of x in + ⟨⟨y, _, ⟨i, rfl⟩, hy⟩, begin + convert hy', + exact of_eq_of _ _ _ _ _ _ + end⟩⟩ + +instance Union_field : discrete_field (⋃ i : c, i.1.carrier) := +@equiv.discrete_field _ _ (equiv_direct_limit c) + (field.direct_limit.discrete_field _ _) + +set_option class.instance_max_depth 50 + +instance is_field_hom_Union (i : c) : is_field_hom + (inclusion (set.subset_Union (λ i : c, i.1.carrier) i)) := +suffices inclusion (set.subset_Union (λ i : c, i.1.carrier) i) = + ((equiv_direct_limit c).symm ∘ + ring.direct_limit.of (λ i : c, i.1.carrier) (chain_map c) i), + by rw this; exact is_ring_hom.comp _ _, +funext $ λ ⟨_, _⟩, + (equiv_direct_limit c).injective $ + by rw [function.comp_app, equiv.apply_symm_apply]; + exact of_eq_of _ _ _ _ _ _ + +end chain + +--def maximal_extension (c : set (extension K)) (hc : chain (≤) c) : extension K := + +private def maximal_extension_chain (c : set (extension K)) (hc : chain (≤) c) : + { ub : extension K // ∀ L, L ∈ c → L ≤ ub } := +if h : nonempty c + then by letI : chain' c := ⟨hc⟩; exact + ⟨{ carrier := ⋃ (i : c), i.1.carrier, + /- of_ring_hom probably works here. Field is isomorphic to direct limit of a bunch of extensions -/ + algebra := sorry, + algebraic := sorry }, -- Field is isomorphic to direct limit of some algebraic extensions + λ e he, ⟨by convert subset_Union _ (⟨e, he⟩ : c); refl, + algebraic_closure.is_field_hom_Union c ⟨e, he⟩⟩⟩ + else ⟨base_extension K, λ a ha, (h ⟨⟨a, ha⟩⟩).elim⟩ + +section adjoin_root +variables {L : extension K} (f : polynomial L.carrier) [hif : irreducible f] +include hif + +instance adjoin_root_algebraic_closure.field : + discrete_field (adjoin_root f) := adjoin_root.field + +instance adjoin_root_algebraic_closure.is_ring_hom : + is_ring_hom (@adjoin_root.of _ _ _ f) := adjoin_root.is_ring_hom + +--set_option class.instance_max_depth 100 +private def adjoin_root.of_embedding : L.carrier ↪ adjoin_root f := +⟨adjoin_root.of, @is_field_hom.injective _ _ _ _ _ $ adjoin_root_algebraic_closure.is_ring_hom f⟩ + +/- TODO: move -/ +instance adjoin_root.algebra : algebra K (adjoin_root f) := +algebra.of_ring_hom (adjoin_root.of ∘ algebra_map _) (is_ring_hom.comp _ _) + +variable (K) + +private def adjoin_root_extension_map : adjoin_root f ↪ big_type K := +thing (adjoin_root.of_embedding f) ⟨subtype.val, subtype.val_injective⟩ + (λ i, let e : big_type K ↪ ℕ × polynomial K := i.trans + (algebraic_embedding sorry) in --adjoining a root to an algebraic extension gives an algebraic extension + cantor_injective e.1 e.2) + +private lemma adjoin_root_extension_map_apply (x : L.carrier) : + (adjoin_root_extension_map K f) (@adjoin_root.of _ _ _ f x) = x.val := +thing_commutes _ _ _ _ + +instance range_adjoin_root_extension_map.discrete_field : + discrete_field (set.range (@adjoin_root_extension_map K _ _ f _)) := +equiv.discrete_field (equiv.set.range _ (embedding.inj _)).symm + +private def adjoin_root_extension : extension K := +{ carrier := set.range (@adjoin_root_extension_map K _ _ f _), + algebra := algebra.of_ring_hom + ((equiv.set.range _ (embedding.inj' (adjoin_root_extension_map K f))).symm.symm ∘ + algebra_map _) (is_ring_hom.comp _ _), + algebraic := sorry } -- adjoining a root to an algebraic extension gives an algebraic extension + +variable {L} +private lemma subset_adjoin_root_extension : L.carrier ⊆ (adjoin_root_extension K f).carrier := +λ x h, ⟨adjoin_root.of_embedding f ⟨x, h⟩, thing_commutes _ _ _ _⟩ + +private lemma adjoin_root_inclusion_eq : inclusion (subset_adjoin_root_extension K f) = + ((equiv.set.range _ (adjoin_root_extension_map K f).2).symm.symm ∘ adjoin_root.of_embedding f) := +funext $ λ ⟨_, _⟩, subtype.eq $ eq.symm $ adjoin_root_extension_map_apply _ _ _ + +private lemma le_adjoin_root_extension : L ≤ adjoin_root_extension K f := +⟨subset_adjoin_root_extension K f, + by rw [adjoin_root_inclusion_eq]; dsimp [adjoin_root.of_embedding]; exact is_ring_hom.comp _ _⟩ + +private def equiv_adjoin_root_of_le (h : adjoin_root_extension K f ≤ L) : + L.carrier ≃r adjoin_root f := +have left_inv : left_inverse (inclusion h.fst ∘ (equiv.set.range _ + (adjoin_root_extension_map K f).2)) adjoin_root.of, + from λ _, by simp [adjoin_root_extension_map_apply, inclusion], +{ to_fun := coe, + inv_fun := inclusion h.fst ∘ (equiv.set.range _ (adjoin_root_extension_map K f).2), + left_inv := left_inv, + right_inv := right_inverse_of_injective_of_left_inverse + (injective_comp (inclusion_injective _) (equiv.injective _)) + left_inv, + hom := by apply_instance } + +private def adjoin_root_equiv_adjoin_root_extension : adjoin_root f ≃r (adjoin_root_extension K f).carrier := +(equiv.set.range _ (adjoin_root_extension_map K f).2).symm.ring_equiv.symm + +end adjoin_root + +private lemma exists_algebraic_closure (K : Type u) [discrete_field K] : + ∃ m : extension K, ∀ a, m ≤ a → a ≤ m := +zorn (λ c hc, (maximal_extension_chain c hc).exists_of_subtype) (λ _ _ _, le_trans) + +private def closed_extension (K : Type u) [discrete_field K] := +classical.some (exists_algebraic_closure K) + +def algebraic_closure (K : Type u) [discrete_field K] : Type u := +((classical.some (exists_algebraic_closure K))).carrier + +end classical + +section is_algebraically_closed +/- In this section we prove the algebraic closure is algebraically closed -/ + +local attribute [reducible] algebraic_closure + +variables (f : polynomial (algebraic_closure K)) [hif : irreducible f] +include hif + +variable (K) + +def algebraic_closure_equiv_adjoin_root : algebraic_closure K ≃r adjoin_root f := +equiv_adjoin_root_of_le K f $ + classical.some_spec (exists_algebraic_closure K) _ (le_adjoin_root_extension _ _) + +instance ring_equiv.is_semiring_hom {α β : Type*} [ring α] [ring β] (e : α ≃r β) : + is_semiring_hom (e.to_equiv : α → β) := +is_ring_hom.is_semiring_hom _ + +omit hif + +private def is_algebraically_closed_aux : is_algebraically_closed (algebraic_closure K) := +is_algebraically_closed_of_irreducible_has_root $ +λ f hf, let e := by exactI algebraic_closure_equiv_adjoin_root K f in +⟨_, exists_root_of_equiv e (adjoin_root.eval₂_root f)⟩ + +end is_algebraically_closed + +/- To avoid diamonds, the `decidable_eq` instance is set to `classical.dec_eq`, + as opposed to the (noncomputable, but not def-eq to `classical.dec_eq`) instance given by + `(closed_extension K).field` -/ +instance : discrete_field (algebraic_closure K) := +{ has_decidable_eq := classical.dec_eq _, + ..(closed_extension K).field } + +instance : algebra K (algebraic_closure K) := (closed_extension K).algebra + +instance : is_algebraically_closed (algebraic_closure K) := +by convert is_algebraically_closed_aux K + +protected def is_integral : ∀ x : algebraic_closure K, is_integral K x := +(closed_extension K).algebraic + +attribute [irreducible] algebraic_closure closed_extension algebraic_closure.algebra + +section lift +/- In this section, the homomorphism from any algebraic extension into an algebraically + closed extension is proven to exist. The assumption that M is algebraically closed could probably + easily be switched to an assumption that M contains all the roots of polynomials in K -/ +variables {L : Type v} {M : Type w} [discrete_field L] [algebra K L] + [discrete_field M] [algebra K M] [is_algebraically_closed M] (hL : ∀ x : L, is_integral K x) + +/-- This structure is used to prove the existence of a homomorphism from any algebraic extension + into an algebraic closure -/ +variables (K L M) +include hL + +private structure subfield_and_hom extends extension K := +( to_algebraically_closed : carrier →ₐ[K] M ) +( to_field : carrier →ₐ[K] L ) + +variables {K L M} + +instance subfield_and_hom.preorder : preorder (subfield_and_hom K L M hL) := +preorder.lift subfield_and_hom.to_extension (by apply_instance) + +private def maximal_subfield_and_hom_chain (c : set (subfield_and_hom K L M hL)) (hc : chain (≤) c) : + ∃ ub : subfield_and_hom K L M hL, ∀ N, N ∈ c → N ≤ ub := +let ub := (maximal_extension_chain (subfield_and_hom.to_extension '' c) (chain.image (≤) _ _ (λ _ _, id) hc)) in +⟨{ to_algebraically_closed := sorry, --field in question is direct limit of a bunch of fields with + --algebra homs into M + to_field := sorry, -- direct limit of a bunch of subfields is also a subfield + ..ub.1 }, + λ n hN, ub.2 _ (mem_image_of_mem _ hN)⟩ + +private lemma exists_maximal_subfield_and_hom : ∃ N : subfield_and_hom K L M hL, + ∀ O, N ≤ O → O ≤ N := +zorn (maximal_subfield_and_hom_chain _) (λ _ _ _, le_trans) + +variable (M) + +private def maximal_subfield_and_hom : subfield_and_hom K L M hL := +classical.some (exists_maximal_subfield_and_hom hL) + +instance akgh : algebra (maximal_subfield_and_hom M hL).carrier L := +algebra.of_ring_hom (maximal_subfield_and_hom M hL).to_field (by apply_instance) + +-- Given K:L:M, if M is algebraic over K it is algebraic over L (names are different) +private lemma is_integral_over_maximal (x : L) : is_integral (maximal_subfield_and_hom M hL).carrier x := sorry + +variables (f : polynomial (maximal_subfield_and_hom M hL).carrier) [hif : irreducible f] + {x : L} (hxf : f.eval₂ (maximal_subfield_and_hom M hL).to_field x = 0) + +include hif hxf + +private def adjoin_root_subfield_and_hom : subfield_and_hom K L M hL := +{ to_algebraically_closed := sorry, -- should be adjoin_root.lift composed with an isomorphism + to_field := sorry, -- + ..adjoin_root_extension K f } + +private lemma le_adjoin_root_subfield_and_hom : maximal_subfield_and_hom M hL ≤ + adjoin_root_subfield_and_hom M hL f hxf := +le_adjoin_root_extension _ _ + +private def maximal_subfield_and_hom_equiv_adjoin_root := +equiv_adjoin_root_of_le K f $ + classical.some_spec (exists_maximal_subfield_and_hom hL) _ + (le_adjoin_root_subfield_and_hom M hL f hxf) + +omit hif hxf + +private lemma surjective_maximal_subfield_and_hom_to_field : + function.surjective (maximal_subfield_and_hom M hL).to_field := +λ x, let hx := is_integral_over_maximal M hL x in +by letI := minimal_polynomial_irreducible hx; exact +⟨_, root_minimal_polynomial hx + (exists_root_of_equiv (maximal_subfield_and_hom_equiv_adjoin_root M hL _ + (aeval_minimal_polynomial hx)) (adjoin_root.eval₂_root _))⟩ + +private def equiv_maximal_subfield_and_hom : + (maximal_subfield_and_hom M hL).carrier ≃ₐ[K] L := +{ ..(maximal_subfield_and_hom M hL).to_field, + ..equiv.of_bijective + ⟨is_field_hom.injective _, surjective_maximal_subfield_and_hom_to_field _ _⟩ } + +/-- The hom from an algebraic extension of K into an algebraic closure -/ +def lift : L →ₐ[K] M := +(maximal_subfield_and_hom M hL).to_algebraically_closed.comp +(equiv_maximal_subfield_and_hom M hL).symm.to_alg_hom + +end lift + +end algebraic_closure From e01a15e4344afdd15d25bda1712cf33201d1b64a Mon Sep 17 00:00:00 2001 From: Johan Commelin Date: Mon, 26 Aug 2019 14:53:26 +0200 Subject: [PATCH 04/51] Remove sorries about minimal polynomials --- src/field_theory/algebraic_closure.lean | 28 +--- src/field_theory/minimal_polynomial.lean | 190 +++++++++++++++++++++++ 2 files changed, 194 insertions(+), 24 deletions(-) create mode 100644 src/field_theory/minimal_polynomial.lean diff --git a/src/field_theory/algebraic_closure.lean b/src/field_theory/algebraic_closure.lean index 0c2cf19f6296c..960d195b45166 100644 --- a/src/field_theory/algebraic_closure.lean +++ b/src/field_theory/algebraic_closure.lean @@ -4,10 +4,9 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Chris Hughes -/ -import ring_theory.adjoin_root import algebra.direct_limit import set_theory.schroeder_bernstein -import ring_theory.integral_closure +import field_theory.minimal_polynomial universes u v w open polynomial zorn set function @@ -21,25 +20,6 @@ local attribute [instance, priority 0] subtype.decidable_eq lemma injective_eq {α : Sort*} : injective (eq : α → α → Prop) := λ _ _ h, h.symm ▸ rfl -section minimal_polynomial -/- To be moved -/ -variables {α : Type u} {β : Type v} [discrete_field α] [discrete_field β] [algebra α β] - -def minimal_polynomial {x : β} (hx : is_integral α x) : polynomial α := sorry - -lemma minimal_polynomial_irreducible {x : β} (hx : is_integral α x) : - irreducible (minimal_polynomial hx) := sorry - -lemma minimal_polynomial_monic {x : β} (hx : is_integral α x) : monic (minimal_polynomial hx) := sorry - -@[simp] lemma aeval_minimal_polynomial {x : β} (hx : is_integral α x) : - aeval α β x (minimal_polynomial hx) = 0 := sorry - -lemma root_minimal_polynomial {x : β} (hx : is_integral α x) {y : α} - (h : (minimal_polynomial hx).eval y = 0) : algebra_map β y = x := sorry - -end minimal_polynomial - @[instance] lemma equiv.is_ring_hom {α β : Type*} [ring β] (e : α ≃ β) : @is_ring_hom β α _ (equiv.ring e) e.symm := by split; simp [equiv.mul_def, equiv.add_def, equiv.one_def] @@ -536,10 +516,10 @@ omit hif hxf private lemma surjective_maximal_subfield_and_hom_to_field : function.surjective (maximal_subfield_and_hom M hL).to_field := λ x, let hx := is_integral_over_maximal M hL x in -by letI := minimal_polynomial_irreducible hx; exact -⟨_, root_minimal_polynomial hx +by letI := minimal_polynomial.irreducible hx; exact +⟨_, minimal_polynomial.root hx (exists_root_of_equiv (maximal_subfield_and_hom_equiv_adjoin_root M hL _ - (aeval_minimal_polynomial hx)) (adjoin_root.eval₂_root _))⟩ + (minimal_polynomial.aeval hx)) (adjoin_root.eval₂_root _))⟩ private def equiv_maximal_subfield_and_hom : (maximal_subfield_and_hom M hL).carrier ≃ₐ[K] L := diff --git a/src/field_theory/minimal_polynomial.lean b/src/field_theory/minimal_polynomial.lean new file mode 100644 index 0000000000000..68a929d0b7b0e --- /dev/null +++ b/src/field_theory/minimal_polynomial.lean @@ -0,0 +1,190 @@ +/- +Copyright (c) 2019 Chris Hughes. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Chris Hughes, Johan Commelin +-/ + +import ring_theory.adjoin_root +import ring_theory.integral_closure + +universes u v w +open polynomial set function + +namespace polynomial +variables {α : Type*} [decidable_eq α] [comm_semiring α] + +@[simp] lemma monic.leading_coeff {p : polynomial α} (hp : p.monic) : + leading_coeff p = 1 := hp + +lemma degree_eq_iff_nat_degree_eq {p : polynomial α} {n : ℕ} (hp : p ≠ 0) : + p.degree = n ↔ p.nat_degree = n := +by rw [degree_eq_nat_degree hp, with_bot.coe_eq_coe] + +lemma degree_eq_iff_nat_degree_eq_of_pos {p : polynomial α} {n : ℕ} (hn : n > 0) : + p.degree = n ↔ p.nat_degree = n := +begin + split, + { intro H, rwa ← degree_eq_iff_nat_degree_eq, rintro rfl, + rw degree_zero at H, exact option.no_confusion H }, + { intro H, rwa degree_eq_iff_nat_degree_eq, rintro rfl, + rw nat_degree_zero at H, linarith } +end + +instance coeff.is_add_group_hom {n : ℕ} : is_add_monoid_hom (λ p : polynomial α, p.coeff n) := +{ map_add := λ p q, coeff_add p q n, + map_zero := coeff_zero _ } + +lemma coeff_eq_zero_degree_lt {p : polynomial α} {n : ℕ} (h : p.degree < n) : + p.coeff n = 0 := +begin + suffices : n ∉ p.support, { rwa finsupp.not_mem_support_iff at this }, + contrapose! h, + apply finset.le_sup h +end + +lemma coeff_eq_zero_nat_degree_lt {p : polynomial α} {n : ℕ} (h : p.nat_degree < n) : + p.coeff n = 0 := +begin + apply coeff_eq_zero_degree_lt, + by_cases hp : p = 0, + { subst hp, exact with_bot.bot_lt_coe n }, + { rwa [degree_eq_nat_degree hp, with_bot.coe_lt_coe] } +end + +lemma as_sum (p : polynomial α) : + p = (finset.range (p.nat_degree + 1)).sum (λ i, C (p.coeff i) * X^i) := +begin + rw ext, intro n, symmetry, + calc _ = _ : (finset.sum_hom (λ q : polynomial α, q.coeff n)).symm + ... = _ : + begin + rw finset.sum_eq_single n; + try { simp only [mul_one, coeff_X_pow, if_pos rfl, coeff_C_mul], }, + { intros i hi hne, rw [if_neg hne.symm, mul_zero] }, + { intro h, rw [finset.mem_range, not_lt] at h, + exact coeff_eq_zero_nat_degree_lt h } + end +end + +-- lemma explicit_form_of_degree_eq_one {α : Type*} [decidable_eq α] [comm_semiring α] +-- {p : polynomial α} (hp : degree p = 1) : p = C p.leading_coeff * X + C (coeff p 0) := +-- begin +-- have ndeg : p.nat_degree = 1, +-- { }, +-- rw polynomial.ext, intro n, +-- by_cases hn : n ≤ 1, +-- { cases n, {simp}, +-- obtain rfl : n = 0, { rwa [nat.succ_le_succ_iff, nat.le_zero_iff] at hn }, +-- simp [hp, leading_coeff], } +-- end + +end polynomial + +noncomputable theory + +/- Turn down the instance priority for subtype.decidable_eq and use classical.dec_eq everywhere, + to avoid diamonds -/ +local attribute [instance, priority 0] subtype.decidable_eq + +variables {α : Type u} {β : Type v} [discrete_field α] [discrete_field β] [algebra α β] + +def minimal_polynomial {x : β} (hx : is_integral α x) : polynomial α := +well_founded.min polynomial.degree_lt_wf _ (ne_empty_iff_exists_mem.mpr hx) + +namespace minimal_polynomial +variables {x : β} (hx : is_integral α x) + +lemma monic : monic (minimal_polynomial hx) := +(well_founded.min_mem degree_lt_wf _ (ne_empty_iff_exists_mem.mpr hx)).1 + +@[simp] lemma aeval : aeval α β x (minimal_polynomial hx) = 0 := +(well_founded.min_mem degree_lt_wf _ (ne_empty_iff_exists_mem.mpr hx)).2 + +lemma ne_zero : (minimal_polynomial hx) ≠ 0 := ne_zero_of_monic (monic hx) + +lemma min {p : polynomial α} (pmonic : p.monic) (hp : polynomial.aeval α β x p = 0) : + degree (minimal_polynomial hx) ≤ degree p := +le_of_not_lt $ well_founded.not_lt_min degree_lt_wf _ (ne_empty_iff_exists_mem.mpr hx) ⟨pmonic, hp⟩ + +lemma degree_le_of_ne_zero {p : polynomial α} (pnz : p ≠ 0) (hp : polynomial.aeval α β x p = 0) : + degree (minimal_polynomial hx) ≤ degree p := +begin + rw ← degree_mul_leading_coeff_inv p pnz, + apply min _ (monic_mul_leading_coeff_inv pnz), + simp [hp] +end + +lemma unique {p : polynomial α} (pmonic : p.monic) (hp : polynomial.aeval α β x p = 0) + (pmin : ∀ q : polynomial α, q.monic → polynomial.aeval α β x q = 0 → degree p ≤ degree q) : + p = minimal_polynomial hx := +begin + symmetry, apply eq_of_sub_eq_zero, + by_contra hnz, + have := degree_le_of_ne_zero hx hnz (by simp [hp]), + contrapose! this, + apply degree_sub_lt _ (ne_zero hx), + { rw [(monic hx).leading_coeff, pmonic.leading_coeff] }, + { exact le_antisymm (min hx pmonic hp) (pmin (minimal_polynomial hx) (monic hx) (aeval hx)) }, +end + +lemma dvd {p : polynomial α} (hp : polynomial.aeval α β x p = 0) : + minimal_polynomial hx ∣ p := +begin + rw ← dvd_iff_mod_by_monic_eq_zero (monic hx), + by_contra hnz, + have := degree_le_of_ne_zero hx hnz _, + { contrapose! this, + exact degree_mod_by_monic_lt _ (monic hx) (ne_zero hx) }, + { rw ← mod_by_monic_add_div p (monic hx) at hp, + simpa using hp } +end + +lemma degree_ne_zero : degree (minimal_polynomial hx) ≠ 0 := +begin + assume deg_eq_zero, + have ndeg_eq_zero : nat_degree (minimal_polynomial hx) = 0, + { simpa using congr_arg nat_degree (eq_C_of_degree_eq_zero deg_eq_zero) }, + have eq_one : minimal_polynomial hx = 1, + { rw eq_C_of_degree_eq_zero deg_eq_zero, congr, + simpa [ndeg_eq_zero.symm] using (monic hx).leading_coeff }, + simpa [eq_one, aeval_def] using aeval hx +end + +lemma not_is_unit : ¬ is_unit (minimal_polynomial hx) := +begin + intro H, apply degree_ne_zero hx, + exact degree_eq_zero_of_is_unit H +end + +lemma prime : prime (minimal_polynomial hx) := +begin + refine ⟨ne_zero hx, not_is_unit hx, _⟩, + rintros p q ⟨d, h⟩, + have : polynomial.aeval α β x (p*q) = 0 := by simp [h, aeval hx], + replace : polynomial.aeval α β x p = 0 ∨ polynomial.aeval α β x q = 0 := by simpa, + cases this; [left, right]; apply dvd; assumption +end + +lemma irreducible : irreducible (minimal_polynomial hx) := +irreducible_of_prime (prime hx) + +lemma root {x : β} (hx : is_integral α x) {y : α} + (h : (minimal_polynomial hx).eval y = 0) : algebra_map β y = x := +begin + have ndeg_one : nat_degree (minimal_polynomial hx) = 1, + { rw ← polynomial.degree_eq_iff_nat_degree_eq_of_pos (nat.zero_lt_one), + exact degree_eq_one_of_irreducible_of_root (irreducible hx) h }, + have coeff_one : (minimal_polynomial hx).coeff 1 = 1, + { simpa [ndeg_one, leading_coeff] using (monic hx).leading_coeff }, + have hy : y = - coeff (minimal_polynomial hx) 0, + { rw (minimal_polynomial hx).as_sum at h, + apply eq_neg_of_add_eq_zero, + simpa [ndeg_one, finset.sum_range_succ, coeff_one] using h }, + subst y, + rw [algebra.map_neg, neg_eq_iff_add_eq_zero], + have H := aeval hx, + rw (minimal_polynomial hx).as_sum at H, + simpa [ndeg_one, finset.sum_range_succ, coeff_one, aeval_def] using H +end + +end minimal_polynomial From 2aff46a98242ac2a986c58302fe7bae3ea29eb3c Mon Sep 17 00:00:00 2001 From: Johan Commelin Date: Mon, 26 Aug 2019 15:22:41 +0200 Subject: [PATCH 05/51] Define alg_equiv.symm --- src/field_theory/algebraic_closure.lean | 24 +++++++++++++++++++++++- 1 file changed, 23 insertions(+), 1 deletion(-) diff --git a/src/field_theory/algebraic_closure.lean b/src/field_theory/algebraic_closure.lean index 960d195b45166..e89b0e23feaf1 100644 --- a/src/field_theory/algebraic_closure.lean +++ b/src/field_theory/algebraic_closure.lean @@ -42,6 +42,21 @@ begin equiv.apply_symm_apply, is_ring_hom.map_zero e.to_equiv, hx], end +namespace alg_hom +variables {α : Type u} {β : Type v} {γ : Type w} [comm_ring α] [ring β] [ring γ] + [algebra α β] [algebra α γ] (f : β →ₐ[α] γ) + +def inverse (g : γ → β) (h₁ : left_inverse g f) (h₂ : right_inverse g f) : γ →ₐ[α] β := +by dsimp [left_inverse, function.right_inverse] at h₁ h₂; exact +{ to_fun := g, + hom := show is_ring_hom g, from + { map_add := λ x y, by rw [← h₁ (g x + g y), f.map_add, h₂, h₂], + map_mul := λ x y, by rw [← h₁ (g x * g y), f.map_mul, h₂, h₂], + map_one := by rw [← h₁ 1, f.map_one] }, + commutes' := λ a, by rw [← h₁ (algebra_map β a), f.commutes] } + +end alg_hom + set_option old_structure_cmd true structure alg_equiv (α β γ : Type*) [comm_ring α] [ring β] [ring γ] @@ -56,7 +71,14 @@ namespace alg_equiv variables {α : Type u} {β : Type v} {γ : Type w} [comm_ring α] [ring β] [ring γ] [algebra α β] [algebra α γ] -protected def symm (e : β ≃ₐ[α] γ) : γ ≃ₐ[α] β := sorry +protected def refl : β ≃ₐ[α] β := +{ hom := is_ring_hom.id, + commutes' := λ b, rfl, + .. equiv.refl β } + +protected def symm (e : β ≃ₐ[α] γ) : γ ≃ₐ[α] β := +{ .. e.to_alg_hom.inverse e.inv_fun e.left_inv e.right_inv, + .. e.to_equiv.symm } end alg_equiv From d5c489766964a07601e4159a4e85a6cfc1eb71af Mon Sep 17 00:00:00 2001 From: Johan Commelin Date: Mon, 26 Aug 2019 15:23:33 +0200 Subject: [PATCH 06/51] typo --- src/field_theory/algebraic_closure.lean | 4 +++- 1 file changed, 3 insertions(+), 1 deletion(-) diff --git a/src/field_theory/algebraic_closure.lean b/src/field_theory/algebraic_closure.lean index e89b0e23feaf1..5df7fa183fc3b 100644 --- a/src/field_theory/algebraic_closure.lean +++ b/src/field_theory/algebraic_closure.lean @@ -80,6 +80,8 @@ protected def symm (e : β ≃ₐ[α] γ) : γ ≃ₐ[α] β := { .. e.to_alg_hom.inverse e.inv_fun e.left_inv e.right_inv, .. e.to_equiv.symm } +-- TODO: trans + end alg_equiv section thing @@ -212,7 +214,7 @@ private def base_extension (K : Type u) [discrete_field K] : extension K := (by apply_instance), algebraic := sorry } --a field is algebraic over an isomorphic field -/-- not used but might help woth sorries -/ +/-- not used but might help with sorries -/ private def extension.of_algebraic {L : Type v} [discrete_field L] [algebra K L] (hL : ∀ x : L, is_integral K x) : extension K := { carrier := set.range (algebraic_embedding_big_type hL), From 71ed4b556318f88b804e65b955218d585255bf67 Mon Sep 17 00:00:00 2001 From: Johan Commelin Date: Mon, 26 Aug 2019 15:44:42 +0200 Subject: [PATCH 07/51] Remove another sorry, in base_extension --- src/field_theory/algebraic_closure.lean | 22 +++++++++++++++++++++- 1 file changed, 21 insertions(+), 1 deletion(-) diff --git a/src/field_theory/algebraic_closure.lean b/src/field_theory/algebraic_closure.lean index 5df7fa183fc3b..18f68de192718 100644 --- a/src/field_theory/algebraic_closure.lean +++ b/src/field_theory/algebraic_closure.lean @@ -146,6 +146,18 @@ lemma is_algebraically_closed_of_irreducible_has_root end is_algebraically_closed +--move this +namespace polynomial +variables (R : Type u) (A : Type v) +variables [comm_ring R] [comm_ring A] [algebra R A] +variables [decidable_eq R] (x : A) + +@[simp] lemma aeval_X : aeval R A x X = x := eval₂_X _ x + +@[simp] lemma aeval_C (r : R) : aeval R A x (C r) = algebra_map A r := eval₂_C _ x + +end polynomial + namespace algebraic_closure section classical @@ -212,7 +224,15 @@ private def base_extension (K : Type u) [discrete_field K] : extension K := { carrier := set.range (bembedding K), algebra := algebra.of_ring_hom (equiv.set.range _ (bembedding K).2).symm.symm (by apply_instance), - algebraic := sorry } --a field is algebraic over an isomorphic field + algebraic := + begin + rintro ⟨_, x, rfl⟩, + refine ⟨X + C (-x), monic_X_add_C (-x), _⟩, + rw [alg_hom.map_add, C_neg, alg_hom.map_neg, polynomial.aeval_X, polynomial.aeval_C], + exact add_neg_self _ + end } + +#exit /-- not used but might help with sorries -/ private def extension.of_algebraic {L : Type v} [discrete_field L] [algebra K L] From 067984e4ee6c2ca21f8fcb0838369b351f111b19 Mon Sep 17 00:00:00 2001 From: Johan Commelin Date: Mon, 26 Aug 2019 18:02:38 +0200 Subject: [PATCH 08/51] Work in progress --- src/field_theory/algebraic_closure.lean | 75 +++++++++++++++++++++++-- 1 file changed, 69 insertions(+), 6 deletions(-) diff --git a/src/field_theory/algebraic_closure.lean b/src/field_theory/algebraic_closure.lean index 18f68de192718..2437e89c1a624 100644 --- a/src/field_theory/algebraic_closure.lean +++ b/src/field_theory/algebraic_closure.lean @@ -55,6 +55,15 @@ by dsimp [left_inverse, function.right_inverse] at h₁ h₂; exact map_one := by rw [← h₁ 1, f.map_one] }, commutes' := λ a, by rw [← h₁ (algebra_map β a), f.commutes] } +instance quotient.algebra (I : ideal α) : algebra α I.quotient := +algebra.of_ring_hom (ideal.quotient.mk I) (ideal.quotient.is_ring_hom_mk I) + +def induced_quotient_map (I J : ideal α) (h : I ≤ J) : + I.quotient →ₐ[α] J.quotient := +{ to_fun := ideal.quotient.lift I (ideal.quotient.mk J) $ + by { intros i hi, erw submodule.quotient.mk_eq_zero, exact h hi }, + commutes' := λ a, by { erw ideal.quotient.lift_mk, refl } } + end alg_hom set_option old_structure_cmd true @@ -84,6 +93,61 @@ protected def symm (e : β ≃ₐ[α] γ) : γ ≃ₐ[α] β := end alg_equiv +namespace adjoin_root +variables {α : Type*} [comm_ring α] [decidable_eq α] (f : polynomial α) + +instance : algebra α (adjoin_root f) := +algebra.of_ring_hom coe $ by apply_instance + +lemma fg_of_monic (hf : f.monic) : submodule.fg (⊤ : submodule α (adjoin_root f)) := +begin + refine ⟨(finset.range (f.nat_degree + 1)).image (λ i, adjoin_root.mk (X^i)), _⟩, + rw submodule.eq_top_iff', + intro a, + apply quotient.induction_on' a, + intro p, show mk p ∈ _, + rw [← mod_by_monic_add_div p hf, is_ring_hom.map_add mk, is_ring_hom.map_mul mk], + { refine submodule.add_mem _ _ _, + { apply_instance }, + { rw [(p %ₘ f).as_sum, ← finset.sum_hom mk], + { refine submodule.sum_mem _ _, + intros i hi, + rw [is_ring_hom.map_mul mk], + { show algebra_map _ (coeff (p %ₘ f) i) * _ ∈ _, + rw ← algebra.smul_def, + refine submodule.smul_mem _ _ (submodule.subset_span _), + rw [finset.mem_coe, finset.mem_image], + refine ⟨i, _, rfl⟩, + rw finset.mem_range at hi ⊢, + refine lt_of_lt_of_le hi (nat.succ_le_succ _), + by_cases (p %ₘ f) = 0, { simp [h] }, + rw ← with_bot.coe_le_coe, + rw ← degree_eq_nat_degree h, + apply le_trans (degree_mod_by_monic_le p hf), + exact degree_le_nat_degree }, + { apply_instance } }, + { apply_instance } }, + { convert submodule.zero_mem _, + convert zero_mul _ using 2, + erw [submodule.quotient.mk_eq_zero _], + apply submodule.subset_span, + exact mem_singleton _ } }, + all_goals { apply_instance } +end + +lemma fg {f : polynomial K} (hf : f ≠ 0) : submodule.fg (⊤ : submodule K (adjoin_root f)) := +begin + have := fg_of_monic _ (monic_mul_leading_coeff_inv hf), + convert submodule.fg_map this, swap, + { apply alg_hom.to_linear_map, + sorry + -- I would like to use alg_hom.induced_quotient_map here + }, + sorry +end + +end adjoin_root + section thing local attribute [instance] classical.dec @@ -232,8 +296,6 @@ private def base_extension (K : Type u) [discrete_field K] : extension K := exact add_neg_self _ end } -#exit - /-- not used but might help with sorries -/ private def extension.of_algebraic {L : Type v} [discrete_field L] [algebra K L] (hL : ∀ x : L, is_integral K x) : extension K := @@ -378,18 +440,19 @@ instance adjoin_root_algebraic_closure.is_ring_hom : private def adjoin_root.of_embedding : L.carrier ↪ adjoin_root f := ⟨adjoin_root.of, @is_field_hom.injective _ _ _ _ _ $ adjoin_root_algebraic_closure.is_ring_hom f⟩ -/- TODO: move -/ -instance adjoin_root.algebra : algebra K (adjoin_root f) := -algebra.of_ring_hom (adjoin_root.of ∘ algebra_map _) (is_ring_hom.comp _ _) - variable (K) +-- set_option class.instance_max_depth 80 +#exit + private def adjoin_root_extension_map : adjoin_root f ↪ big_type K := thing (adjoin_root.of_embedding f) ⟨subtype.val, subtype.val_injective⟩ (λ i, let e : big_type K ↪ ℕ × polynomial K := i.trans (algebraic_embedding sorry) in --adjoining a root to an algebraic extension gives an algebraic extension cantor_injective e.1 e.2) +-- GRRRR. The instance is there!! Why can't Lean find adjoin_root.algebra ??? + private lemma adjoin_root_extension_map_apply (x : L.carrier) : (adjoin_root_extension_map K f) (@adjoin_root.of _ _ _ f x) = x.val := thing_commutes _ _ _ _ From b64985a122c57cb72b3c44f74d2f6619d242aaeb Mon Sep 17 00:00:00 2001 From: Johan Commelin Date: Mon, 26 Aug 2019 22:39:05 +0200 Subject: [PATCH 09/51] Remove a sorry in maximal_extension_chain --- src/field_theory/algebraic_closure.lean | 20 ++++++++++++++++---- 1 file changed, 16 insertions(+), 4 deletions(-) diff --git a/src/field_theory/algebraic_closure.lean b/src/field_theory/algebraic_closure.lean index 2437e89c1a624..ab3be96f76b95 100644 --- a/src/field_theory/algebraic_closure.lean +++ b/src/field_theory/algebraic_closure.lean @@ -417,11 +417,23 @@ end chain private def maximal_extension_chain (c : set (extension K)) (hc : chain (≤) c) : { ub : extension K // ∀ L, L ∈ c → L ≤ ub } := if h : nonempty c - then by letI : chain' c := ⟨hc⟩; exact + then + let L := classical.some (classical.exists_true_of_nonempty h) in + by letI : chain' c := ⟨hc⟩; exact ⟨{ carrier := ⋃ (i : c), i.1.carrier, - /- of_ring_hom probably works here. Field is isomorphic to direct limit of a bunch of extensions -/ - algebra := sorry, - algebraic := sorry }, -- Field is isomorphic to direct limit of some algebraic extensions + algebra := + algebra.of_ring_hom ((inclusion (set.subset_Union (λ i : c, i.1.carrier) L)) ∘ algebra_map _) + (by { refine @is_ring_hom.comp _ _ _ _ _ _ _ _ _ _ }), + algebraic := + begin + rintro ⟨x, hx⟩, + rw mem_Union at hx, + cases hx with L' hx, + rcases (L'.val).algebraic ⟨x, hx⟩ with ⟨p, pmonic, hp⟩, + use [p, pmonic], + rw aeval_def at hp ⊢, + sorry, + end }, -- Field is isomorphic to direct limit of some algebraic extensions λ e he, ⟨by convert subset_Union _ (⟨e, he⟩ : c); refl, algebraic_closure.is_field_hom_Union c ⟨e, he⟩⟩⟩ else ⟨base_extension K, λ a ha, (h ⟨⟨a, ha⟩⟩).elim⟩ From 25080c99ce0d66e694230c2bcac7efc28169c515 Mon Sep 17 00:00:00 2001 From: Johan Commelin Date: Tue, 27 Aug 2019 12:19:11 +0200 Subject: [PATCH 10/51] Merge two sorries --- src/field_theory/algebraic_closure.lean | 49 +++++++++++++++++++++---- 1 file changed, 42 insertions(+), 7 deletions(-) diff --git a/src/field_theory/algebraic_closure.lean b/src/field_theory/algebraic_closure.lean index ab3be96f76b95..a9fcc491d8aab 100644 --- a/src/field_theory/algebraic_closure.lean +++ b/src/field_theory/algebraic_closure.lean @@ -42,6 +42,25 @@ begin equiv.apply_symm_apply, is_ring_hom.map_zero e.to_equiv, hx], end +namespace polynomial +variables {α : Type*} {β : Type*} {γ : Type*} +variables [decidable_eq α] [comm_semiring α] [comm_semiring β] [comm_semiring γ] +variables (f : α → β) (g : β → γ) [is_semiring_hom f] [is_semiring_hom g] (p : polynomial α) (x : β) + +lemma hom_eval₂ : g (p.eval₂ f x) = p.eval₂ (g ∘ f) (g x) := +begin + apply polynomial.induction_on p; clear p, + { intros a, rw [eval₂_C, eval₂_C] }, + { intros p q hp hq, simp only [hp, hq, eval₂_add, is_semiring_hom.map_add g] }, + { intros n a ih, + simp only [pow_succ', is_semiring_hom.map_mul g, (mul_assoc _ _ _).symm, + eval₂_C, eval₂_mul, eval₂_X] at ih ⊢, + rw ih } +end + + +end polynomial + namespace alg_hom variables {α : Type u} {β : Type v} {γ : Type w} [comm_ring α] [ring β] [ring γ] [algebra α β] [algebra α γ] (f : β →ₐ[α] γ) @@ -448,23 +467,27 @@ instance adjoin_root_algebraic_closure.field : instance adjoin_root_algebraic_closure.is_ring_hom : is_ring_hom (@adjoin_root.of _ _ _ f) := adjoin_root.is_ring_hom ---set_option class.instance_max_depth 100 private def adjoin_root.of_embedding : L.carrier ↪ adjoin_root f := ⟨adjoin_root.of, @is_field_hom.injective _ _ _ _ _ $ adjoin_root_algebraic_closure.is_ring_hom f⟩ variable (K) --- set_option class.instance_max_depth 80 -#exit +def aux_instance : algebra K (adjoin_root f) := +algebra.of_ring_hom (adjoin_root.of ∘ algebra_map _) (is_ring_hom.comp _ _) + +local attribute [instance] aux_instance + +lemma adjoin_root.algebraic (x : adjoin_root f) : is_integral K x := +begin + sorry +end private def adjoin_root_extension_map : adjoin_root f ↪ big_type K := thing (adjoin_root.of_embedding f) ⟨subtype.val, subtype.val_injective⟩ (λ i, let e : big_type K ↪ ℕ × polynomial K := i.trans - (algebraic_embedding sorry) in --adjoining a root to an algebraic extension gives an algebraic extension + (algebraic_embedding (adjoin_root.algebraic K f)) in cantor_injective e.1 e.2) --- GRRRR. The instance is there!! Why can't Lean find adjoin_root.algebra ??? - private lemma adjoin_root_extension_map_apply (x : L.carrier) : (adjoin_root_extension_map K f) (@adjoin_root.of _ _ _ f x) = x.val := thing_commutes _ _ _ _ @@ -478,7 +501,19 @@ private def adjoin_root_extension : extension K := algebra := algebra.of_ring_hom ((equiv.set.range _ (embedding.inj' (adjoin_root_extension_map K f))).symm.symm ∘ algebra_map _) (is_ring_hom.comp _ _), - algebraic := sorry } -- adjoining a root to an algebraic extension gives an algebraic extension + algebraic := + begin + rintro ⟨_, x, rfl⟩, + rcases adjoin_root.algebraic K f x with ⟨p, pmonic, hp⟩, + use [p, pmonic], + rw [aeval_def] at hp ⊢, + replace hp := congr_arg + ((equiv.set.range _ (embedding.inj' (adjoin_root_extension_map K f))).symm.symm) hp, + convert hp using 1, + symmetry, + convert polynomial.hom_eval₂ _ _ _ _, + all_goals {apply_instance} + end } variable {L} private lemma subset_adjoin_root_extension : L.carrier ⊆ (adjoin_root_extension K f).carrier := From ca7c85dcdb7a689da6a6a2c61d70b9a02f22f120 Mon Sep 17 00:00:00 2001 From: Johan Commelin Date: Tue, 27 Aug 2019 14:26:31 +0200 Subject: [PATCH 11/51] More sorries removed --- src/field_theory/algebraic_closure.lean | 117 +++++++++++++++++++----- 1 file changed, 94 insertions(+), 23 deletions(-) diff --git a/src/field_theory/algebraic_closure.lean b/src/field_theory/algebraic_closure.lean index a9fcc491d8aab..df42feee14122 100644 --- a/src/field_theory/algebraic_closure.lean +++ b/src/field_theory/algebraic_closure.lean @@ -315,20 +315,55 @@ private def base_extension (K : Type u) [discrete_field K] : extension K := exact add_neg_self _ end } -/-- not used but might help with sorries -/ -private def extension.of_algebraic {L : Type v} [discrete_field L] [algebra K L] - (hL : ∀ x : L, is_integral K x) : extension K := -{ carrier := set.range (algebraic_embedding_big_type hL), - field := equiv.discrete_field (equiv.set.range _ (algebraic_embedding_big_type hL).2).symm, - algebra := sorry, -- a field isomorphic to an algebra is an algebra - algebraic := sorry -- a field isomorphic to an algebraic extension is algebraic - } +-- /-- not used but might help with sorries -/ +-- private def extension.of_algebraic {L : Type v} [discrete_field L] [algebra K L] +-- (hL : ∀ x : L, is_integral K x) : extension K := +-- { carrier := set.range (algebraic_embedding_big_type hL), +-- field := equiv.discrete_field (equiv.set.range _ (algebraic_embedding_big_type hL).2).symm, +-- algebra := sorry, -- a field isomorphic to an algebra is an algebra +-- algebraic := sorry -- a field isomorphic to an algebraic extension is algebraic +-- } + +@[simp] lemma inclusion_refl {α : Type*} {s : set α} (h : s ⊆ s) : inclusion h = id := +funext $ λ x, by { cases x, refl } + +@[simp] lemma inclusion_trans {α : Type*} {s t u : set α} (hst : s ⊆ t) (htu : t ⊆ u) : + inclusion (set.subset.trans hst htu) = inclusion htu ∘ inclusion hst := +funext $ λ x, by { cases x, refl } instance : preorder (extension K) := -{ le := λ L M, ∃ hLM : L.carrier ⊆ M.carrier, is_ring_hom (inclusion hLM), - le_refl := λ _, ⟨set.subset.refl _, by convert is_ring_hom.id; ext; simp⟩, - le_trans := λ L M N ⟨hLM₁, hLM₂⟩ ⟨hMN₁, hMN₂⟩, ⟨set.subset.trans hLM₁ hMN₁, - by resetI; convert is_ring_hom.comp (inclusion hLM₁) (inclusion hMN₁)⟩ } +{ le := λ L M, ∃ hLM : L.carrier ⊆ M.carrier, (is_ring_hom (inclusion hLM) ∧ + (inclusion hLM ∘ (algebra_map L.carrier : K → L.carrier) = algebra_map M.carrier)), + le_refl := λ L, + begin + use set.subset.refl L.carrier, + rw inclusion_refl, + exact ⟨is_ring_hom.id, comp.left_id _⟩ + end, + le_trans := λ L M N ⟨hLM₁, hLM₂, hLM₃⟩ ⟨hMN₁, hMN₂, hMN₃⟩, + begin + use set.subset.trans hLM₁ hMN₁, + rw inclusion_trans, resetI, + refine ⟨is_ring_hom.comp _ _, _⟩, + rw [comp.assoc, hLM₃, hMN₃] + end } + +lemma le_def {L M : extension K} : + L ≤ M ↔ ∃ hLM : L.carrier ⊆ M.carrier, (is_ring_hom (inclusion hLM) ∧ + (inclusion hLM ∘ (algebra_map L.carrier : K → L.carrier) = algebra_map M.carrier)) := iff.rfl + +lemma subset_of_le {L M : extension K} (h : L ≤ M) : L.carrier ⊆ M.carrier := +by { rw le_def at h, rcases h with ⟨_,_,_⟩, assumption } + +lemma ring_hom_of_le {L M : extension K} (h : L ≤ M) : + is_ring_hom (inclusion $ subset_of_le h) := +by { rw le_def at h, rcases h with ⟨_,_,_⟩, assumption } + +lemma compat {L M : extension K} (h : L ≤ M) : + inclusion (subset_of_le h) ∘ (algebra_map L.carrier : K → L.carrier) = algebra_map M.carrier := +by { rw le_def at h, rcases h with ⟨_,_,_⟩, assumption } + +local attribute [instance] ring_hom_of_le private structure chain' (c : set (extension K)) : Prop := (chain : chain (≤) c) @@ -354,10 +389,14 @@ private def chain_map (i j : c) (hij : i ≤ j) : i.1.carrier → j.1.carrier := inclusion (exists.elim hij (λ h _, h)) instance chain_field_hom (i j : c) (hij : i ≤ j) : is_field_hom (chain_map c i j hij) := -exists.elim hij (λ _, id) +ring_hom_of_le hij instance chain_directed_system : directed_system (λ i : c, i.1.carrier) (chain_map c) := -by split; intros; simp [chain_map] +begin + split; intros; simp only [chain_map], + { exact congr_fun (inclusion_refl _) x }, + { exact (congr_fun (inclusion_trans _ _) x).symm } +end private def chain_limit : Type u := ring.direct_limit (λ i : c, i.1.carrier) (chain_map c) @@ -429,9 +468,33 @@ funext $ λ ⟨_, _⟩, by rw [function.comp_app, equiv.apply_symm_apply]; exact of_eq_of _ _ _ _ _ _ -end chain +def Union_algebra (L : c) : algebra K (⋃ i : c, i.1.carrier) := +algebra.of_ring_hom ((inclusion (set.subset_Union (λ i : c, i.1.carrier) L)) ∘ algebra_map _) +(by { refine @is_ring_hom.comp _ _ _ _ _ _ _ _ _ _ }) ---def maximal_extension (c : set (extension K)) (hc : chain (≤) c) : extension K := +lemma Union_compat (L : c) (M : c) : + (inclusion (set.subset_Union (λ i : c, i.1.carrier) M)) ∘ + (algebra_map (M.val.carrier) : K → M.val.carrier) = + by haveI := Union_algebra c L; exact algebra_map (↥⋃ (i : ↥c), (i.val).carrier) := +begin + rcases chain.directed_on (is_chain c) L.1 L.2 M.1 M.2 with ⟨N, hN, hLN, hMN⟩, + rw show (inclusion (set.subset_Union (λ i : c, i.1.carrier) M)) = + ((inclusion (set.subset_Union (λ i : c, i.1.carrier) ⟨N, hN⟩)) ∘ + inclusion (subset_of_le hMN)), + { funext x, refl }, + rw comp.assoc, + rw show inclusion (subset_of_le hMN) ∘ (algebra_map _ : K → (M.val).carrier) = + inclusion (subset_of_le hLN) ∘ algebra_map _, + { rw [compat, ← compat] }, + rw ← comp.assoc, + rw ← show (inclusion (set.subset_Union (λ i : c, i.1.carrier) L)) = + ((inclusion (set.subset_Union (λ i : c, i.1.carrier) ⟨N, hN⟩)) ∘ + inclusion (subset_of_le hLN)), + { funext x, refl }, + refl +end + +end chain private def maximal_extension_chain (c : set (extension K)) (hc : chain (≤) c) : { ub : extension K // ∀ L, L ∈ c → L ≤ ub } := @@ -440,9 +503,7 @@ if h : nonempty c let L := classical.some (classical.exists_true_of_nonempty h) in by letI : chain' c := ⟨hc⟩; exact ⟨{ carrier := ⋃ (i : c), i.1.carrier, - algebra := - algebra.of_ring_hom ((inclusion (set.subset_Union (λ i : c, i.1.carrier) L)) ∘ algebra_map _) - (by { refine @is_ring_hom.comp _ _ _ _ _ _ _ _ _ _ }), + algebra := Union_algebra c L, algebraic := begin rintro ⟨x, hx⟩, @@ -451,10 +512,17 @@ if h : nonempty c rcases (L'.val).algebraic ⟨x, hx⟩ with ⟨p, pmonic, hp⟩, use [p, pmonic], rw aeval_def at hp ⊢, - sorry, - end }, -- Field is isomorphic to direct limit of some algebraic extensions + replace hp := congr_arg (inclusion (set.subset_Union (λ i : c, i.1.carrier) L')) hp, + convert hp using 1; symmetry, + { rw polynomial.hom_eval₂ _ (inclusion _) p _, + congr' 1, + { exact Union_compat c L L' }, + { apply_instance, }, + { apply is_ring_hom.is_semiring_hom, } }, + { refine is_ring_hom.map_zero _ }, + end }, λ e he, ⟨by convert subset_Union _ (⟨e, he⟩ : c); refl, - algebraic_closure.is_field_hom_Union c ⟨e, he⟩⟩⟩ + algebraic_closure.is_field_hom_Union c ⟨e, he⟩, Union_compat c L ⟨e, he⟩⟩⟩ else ⟨base_extension K, λ a ha, (h ⟨⟨a, ha⟩⟩).elim⟩ section adjoin_root @@ -525,7 +593,10 @@ funext $ λ ⟨_, _⟩, subtype.eq $ eq.symm $ adjoin_root_extension_map_apply _ private lemma le_adjoin_root_extension : L ≤ adjoin_root_extension K f := ⟨subset_adjoin_root_extension K f, - by rw [adjoin_root_inclusion_eq]; dsimp [adjoin_root.of_embedding]; exact is_ring_hom.comp _ _⟩ + begin + rw [adjoin_root_inclusion_eq]; dsimp [adjoin_root.of_embedding], + exact ⟨is_ring_hom.comp _ _, rfl⟩ + end⟩ private def equiv_adjoin_root_of_le (h : adjoin_root_extension K f ≤ L) : L.carrier ≃r adjoin_root f := From 542fca8171e593b068748c878cccc81f3bd28249 Mon Sep 17 00:00:00 2001 From: Johan Commelin Date: Tue, 27 Aug 2019 16:58:17 +0200 Subject: [PATCH 12/51] More work on transitivity of algebraicity --- src/field_theory/algebraic_closure.lean | 58 ++++++++++++++++++++++--- 1 file changed, 51 insertions(+), 7 deletions(-) diff --git a/src/field_theory/algebraic_closure.lean b/src/field_theory/algebraic_closure.lean index df42feee14122..18142a403cb82 100644 --- a/src/field_theory/algebraic_closure.lean +++ b/src/field_theory/algebraic_closure.lean @@ -154,15 +154,35 @@ begin all_goals { apply_instance } end +open submodule + lemma fg {f : polynomial K} (hf : f ≠ 0) : submodule.fg (⊤ : submodule K (adjoin_root f)) := begin - have := fg_of_monic _ (monic_mul_leading_coeff_inv hf), - convert submodule.fg_map this, swap, - { apply alg_hom.to_linear_map, - sorry - -- I would like to use alg_hom.induced_quotient_map here - }, - sorry + let tmp : _ := _, + let ψ : adjoin_root (f * C (leading_coeff f)⁻¹) →ₗ[K] adjoin_root f := + { to_fun := ideal.quotient.lift _ mk tmp, + add := λ x y, is_ring_hom.map_add _, + smul := λ c x, is_ring_hom.map_mul _ }, + have trick := fg_of_monic _ (monic_mul_leading_coeff_inv hf), + { convert fg_map trick, swap, + { exact ψ }, + { refine (submodule.eq_top_iff'.mpr _).symm, + intros x, apply quotient.induction_on' x, clear x, + intro g, + rw mem_map, + use [mk g, mem_top, rfl] } }, + { intros g hg, erw quotient.mk_eq_zero, + rw ideal.mem_span_singleton' at hg ⊢, + rcases hg with ⟨g, rfl⟩, rw [mul_comm f, ← mul_assoc], + exact ⟨_, rfl⟩ }, +end +. + +lemma is_integral {f : polynomial K} (hf : f ≠ 0) (x : adjoin_root f) : is_integral K x := +begin + refine is_integral_of_mem_of_fg ⊤ _ x algebra.mem_top, + convert fg hf, + rw eq_top_iff', intro y, exact algebra.mem_top, end end adjoin_root @@ -547,6 +567,30 @@ local attribute [instance] aux_instance lemma adjoin_root.algebraic (x : adjoin_root f) : is_integral K x := begin + rcases adjoin_root.is_integral hif.ne_zero x with ⟨p, pmonic, hp⟩, + let S : set (adjoin_root f) := + (↑((finset.range (f.nat_degree + 1)).image (λ i, (f.coeff i : adjoin_root f))) : set (adjoin_root f)), + let B := algebra.adjoin K S, + have Bfg : submodule.fg (B : submodule K (adjoin_root f)), + { apply fg_adjoin_of_finite, + { apply finset.finite_to_set }, + { intros x hx, + rw [finset.mem_coe, finset.mem_image] at hx, + rcases hx with ⟨i, hi, rfl⟩, + rcases L.algebraic (f.coeff i) with ⟨q, qmonic, hq⟩, + use [q, qmonic], + replace hq := congr_arg (coe : L.carrier → adjoin_root f) hq, + convert hq using 1, + { symmetry, exact polynomial.hom_eval₂ _ _ _ _ }, + { symmetry, exact is_ring_hom.map_zero _ } } }, + + -- (1) we know that x is integral over L, by adjoin_root.is_integral + -- (2) we know that B is fg over K + -- (3) we want to say that adjoin B {x} is fg over B, because (1) + -- (4) hence adjoin B {x} is fg over K + -- (5) and then we are done by is_integral_of_mem_of_fg + + -- refine is_integral_of_mem_of_fg (algebra.adjoin K S) _ x (algebra.subset_adjoin $ mem_insert _ _), sorry end From f3620a68e9e1913ba42fa2951c44e70c63b37785 Mon Sep 17 00:00:00 2001 From: Johan Commelin Date: Tue, 27 Aug 2019 22:28:11 +0200 Subject: [PATCH 13/51] WIP --- src/field_theory/algebraic_closure.lean | 93 +++++++++++++++++++++---- 1 file changed, 80 insertions(+), 13 deletions(-) diff --git a/src/field_theory/algebraic_closure.lean b/src/field_theory/algebraic_closure.lean index 18142a403cb82..9368603e5c75c 100644 --- a/src/field_theory/algebraic_closure.lean +++ b/src/field_theory/algebraic_closure.lean @@ -43,8 +43,11 @@ begin end namespace polynomial -variables {α : Type*} {β : Type*} {γ : Type*} -variables [decidable_eq α] [comm_semiring α] [comm_semiring β] [comm_semiring γ] + +variables {α : Type*} {β : Type*} {γ : Type*} [comm_semiring α] [comm_semiring β] [comm_semiring γ] + +section +variables [decidable_eq α] variables (f : α → β) (g : β → γ) [is_semiring_hom f] [is_semiring_hom g] (p : polynomial α) (x : β) lemma hom_eval₂ : g (p.eval₂ f x) = p.eval₂ (g ∘ f) (g x) := @@ -57,7 +60,38 @@ begin eval₂_C, eval₂_mul, eval₂_X] at ih ⊢, rw ih } end +end +section injective +variables [decidable_eq α] [decidable_eq β] +variables {f : α → β} [is_semiring_hom f] + +lemma map_injective (hf : function.injective f) : + function.injective (map f : polynomial α → polynomial β) := +λ p q h, ext.mpr $ λ m, hf $ +begin + rw ext at h, + specialize h m, + rw [coeff_map f, coeff_map f] at h, + exact h +end + +lemma leading_coeff_of_injective (hf : injective f) (p : polynomial α) : + leading_coeff (p.map f) = f (leading_coeff p) := +begin + delta leading_coeff, + rw [coeff_map f, nat_degree_map' p hf] +end + +lemma monic_of_injective (hf : injective f) {p : polynomial α} (hp : (p.map f).monic) : p.monic := +begin + apply hf, + rw [← leading_coeff_of_injective hf, hp.leading_coeff, is_semiring_hom.map_one f] +end + +end injective +-- def mk' (n : ℕ) (coeffs : fin n → α) : polynomial α := +-- finset.univ.sum (λ i : fin n, C (coeffs i) * X^(i:ℕ)) end polynomial @@ -565,11 +599,13 @@ algebra.of_ring_hom (adjoin_root.of ∘ algebra_map _) (is_ring_hom.comp _ _) local attribute [instance] aux_instance +-- set_option class.instance_max_depth 80 + lemma adjoin_root.algebraic (x : adjoin_root f) : is_integral K x := begin rcases adjoin_root.is_integral hif.ne_zero x with ⟨p, pmonic, hp⟩, let S : set (adjoin_root f) := - (↑((finset.range (f.nat_degree + 1)).image (λ i, (f.coeff i : adjoin_root f))) : set (adjoin_root f)), + (↑((finset.range (p.nat_degree + 1)).image (λ i, (p.coeff i : adjoin_root f))) : set (adjoin_root f)), let B := algebra.adjoin K S, have Bfg : submodule.fg (B : submodule K (adjoin_root f)), { apply fg_adjoin_of_finite, @@ -577,21 +613,52 @@ begin { intros x hx, rw [finset.mem_coe, finset.mem_image] at hx, rcases hx with ⟨i, hi, rfl⟩, - rcases L.algebraic (f.coeff i) with ⟨q, qmonic, hq⟩, + rcases L.algebraic (p.coeff i) with ⟨q, qmonic, hq⟩, use [q, qmonic], replace hq := congr_arg (coe : L.carrier → adjoin_root f) hq, convert hq using 1, { symmetry, exact polynomial.hom_eval₂ _ _ _ _ }, { symmetry, exact is_ring_hom.map_zero _ } } }, - - -- (1) we know that x is integral over L, by adjoin_root.is_integral - -- (2) we know that B is fg over K - -- (3) we want to say that adjoin B {x} is fg over B, because (1) - -- (4) hence adjoin B {x} is fg over K - -- (5) and then we are done by is_integral_of_mem_of_fg - - -- refine is_integral_of_mem_of_fg (algebra.adjoin K S) _ x (algebra.subset_adjoin $ mem_insert _ _), - sorry + refine is_integral_of_mem_of_fg (algebra.adjoin K (S ∪ {x})) _ x (algebra.subset_adjoin _), + { apply algebra.fg_trans Bfg, + apply fg_adjoin_singleton_of_integral, + let pB : polynomial _ := _, + refine ⟨pB, _⟩, swap, + { refine (finset.range (p.nat_degree + 1)).sum + (λ i, C ⟨(p.coeff i : adjoin_root f), _⟩ * X^i), + by_cases hi : i ∈ finset.range (p.nat_degree + 1), + { apply algebra.subset_adjoin, + rw [finset.mem_coe, finset.mem_image], + use [i, hi, rfl] }, + { rw [finset.mem_range, not_lt] at hi, + rw polynomial.coeff_eq_zero_nat_degree_lt hi, + rw [is_ring_hom.map_zero (coe : L.carrier → adjoin_root f)], + have : (0 : adjoin_root f) ∈ (B : submodule K (adjoin_root f)) := + submodule.zero_mem _, + rwa ← submodule.mem_coe at this } }, + let tmp : _ := _, + have hpB : pB.map (algebra_map _ : _ → adjoin_root f) = p.map (coe : L.carrier → adjoin_root f) := _, + { split, + { have := polynomial.monic_map (coe : L.carrier → adjoin_root f) pmonic, + rw ← hpB at this, + refine @polynomial.monic_of_injective + _ _ _ _ _ _ _ tmp _ _ this, + { exact subtype.val_injective } }, + { rwa [aeval_def, ← @eval_map _ _ _ _ _ _ _ _ tmp _, hpB, eval_map], } }, + { rw [p.as_sum, ← finset.sum_hom (map _), ← finset.sum_hom (map _)], + all_goals { try { apply_instance } }, + { apply finset.sum_congr rfl, + intros i hi, + rw [@polynomial.map_mul _ _ _ _ _ _ _ _ _ tmp, + @polynomial.map_C _ _ _ _ _ _ _ _ tmp], + rw [polynomial.map_mul, polynomial.map_C], + refine congr_arg _ _, + rw [is_monoid_hom.map_pow (map _), is_monoid_hom.map_pow (map _)], + { rw [@polynomial.map_X _ _ _ _ _ _ _ tmp, polynomial.map_X] }, + all_goals {sorry} }, + all_goals {sorry} }, + { exact ⟨rfl, rfl, λ _ _, rfl, λ _ _, rfl⟩ } }, + { exact mem_union_right _ (mem_singleton _) } end private def adjoin_root_extension_map : adjoin_root f ↪ big_type K := From cac86328b3fd8ff53292aa4779b319f013fd331d Mon Sep 17 00:00:00 2001 From: Johan Commelin Date: Wed, 28 Aug 2019 14:03:28 +0200 Subject: [PATCH 14/51] Sorry-free definition of algebraic closure --- src/field_theory/algebraic.lean | 187 ++++++++++++++++++++++++ src/field_theory/algebraic_closure.lean | 118 +-------------- 2 files changed, 191 insertions(+), 114 deletions(-) create mode 100644 src/field_theory/algebraic.lean diff --git a/src/field_theory/algebraic.lean b/src/field_theory/algebraic.lean new file mode 100644 index 0000000000000..de12ca092d1ef --- /dev/null +++ b/src/field_theory/algebraic.lean @@ -0,0 +1,187 @@ +import ring_theory.adjoin_root +import ring_theory.integral_closure +import field_theory.minimal_polynomial + +namespace polynomial + +variables {α : Type*} {β : Type*} {γ : Type*} + +section +variables [decidable_eq α] [comm_semiring α] [comm_semiring β] [comm_semiring γ] +variables (f : α → β) (g : β → γ) [is_semiring_hom f] [is_semiring_hom g] (p : polynomial α) (x : β) + +lemma hom_eval₂ : g (p.eval₂ f x) = p.eval₂ (g ∘ f) (g x) := +begin + apply polynomial.induction_on p; clear p, + { intros a, rw [eval₂_C, eval₂_C] }, + { intros p q hp hq, simp only [hp, hq, eval₂_add, is_semiring_hom.map_add g] }, + { intros n a ih, + simp only [pow_succ', is_semiring_hom.map_mul g, (mul_assoc _ _ _).symm, + eval₂_C, eval₂_mul, eval₂_X] at ih ⊢, + rw ih } +end + +end + +section coeffs +variables [decidable_eq α] [comm_semiring α] (p : polynomial α) + +def nonzero_coeffs : finset α := p.support.image p.coeff + +lemma mem_nonzero_coeffs (i : ℕ) (hi : p.coeff i ≠ 0) : p.coeff i ∈ p.nonzero_coeffs := +finset.mem_image.mpr ⟨i, finsupp.mem_support_iff.mpr hi, rfl⟩ + +def coeffs : finset α := p.nonzero_coeffs ∪ {0} + +lemma mem_coeffs (i : ℕ) : p.coeff i ∈ p.coeffs := +if hi : p.coeff i = 0 +then finset.mem_union_right _ (by simp [hi]) +else finset.mem_union_left _ $ mem_nonzero_coeffs p i hi + +end coeffs + +section +variables [decidable_eq α] [comm_semiring α] + +lemma coeff_sum' (s : finset β) (f : β → polynomial α) (n : ℕ) : + coeff (s.sum f) n = s.sum (λ x, coeff (f x) n) := +(@finset.sum_hom _ _ _ s f _ _ (λ q, coeff q n) _).symm + +end + +section injective +open function +variables [decidable_eq α] [decidable_eq β] [comm_semiring α] [comm_semiring β] +variables {f : α → β} [is_semiring_hom f] + +lemma map_injective (hf : function.injective f) : + function.injective (map f : polynomial α → polynomial β) := +λ p q h, ext.mpr $ λ m, hf $ +begin + rw ext at h, + specialize h m, + rw [coeff_map f, coeff_map f] at h, + exact h +end + +lemma leading_coeff_of_injective (hf : injective f) (p : polynomial α) : + leading_coeff (p.map f) = f (leading_coeff p) := +begin + delta leading_coeff, + rw [coeff_map f, nat_degree_map' p hf] +end + +lemma monic_of_injective (hf : injective f) {p : polynomial α} (hp : (p.map f).monic) : p.monic := +begin + apply hf, + rw [← leading_coeff_of_injective hf, hp.leading_coeff, is_semiring_hom.map_one f] +end + +end injective +section lift +open function +variables [decidable_eq α] [comm_ring α] (s : set α) [is_subring s] +variables (p : polynomial α) (hp : ∀ i, p.coeff i ∈ s) + +def lift : polynomial s := +(finset.range (p.nat_degree + 1)).sum (λ i, C ⟨(p.coeff i), hp i⟩ * X^i) + +lemma map_coe_lift : map coe (p.lift s hp) = p := +begin + conv_rhs {rw p.as_sum}, + rw ext, intro n, + rw [coeff_map, lift, coeff_sum', coeff_sum', ← finset.sum_hom coe], + all_goals { try {apply_instance} }, + apply finset.sum_congr rfl, + intros i hi, + rw [coeff_C_mul, coeff_C_mul, is_ring_hom.map_mul coe, ← coeff_map coe, map_pow coe, map_X coe], + { refl }, + all_goals { apply_instance } +end + +end lift + +end polynomial + +namespace subalgebra +variables {R : Type*} {A : Type*} +variables [comm_ring R] [comm_ring A] [algebra R A] + +lemma zero_mem (S : subalgebra R A) : (0 : A) ∈ S := +submodule.zero_mem (S : submodule R A) + +variable [decidable_eq A] + +end subalgebra + +open function algebra polynomial +variables {R : Type*} {A : Type*} {B : Type*} +variables [decidable_eq R] [decidable_eq A] [decidable_eq B] +variables [comm_ring R] [comm_ring A] [comm_ring B] +variables [algebra R A] [algebra A B] +variables (A_alg : ∀ x : A, is_integral R x) + +include A_alg + +set_option class.instance_max_depth 50 + +lemma is_integral_trans_aux (x : B) {p : polynomial A} (pmonic : monic p) (hp : aeval A B x p = 0) + (S : set (comap R A B)) + (hS : S = (↑((finset.range (p.nat_degree + 1)).image + (λ i, to_comap R A B (p.coeff i))) : set (comap R A B))) : + is_integral (adjoin R S) (comap.to_comap R A B x) := +begin + have coeffs_mem : ∀ i, coeff (map (to_comap R A B) p) i ∈ adjoin R S, + { intro i, + by_cases hi : i ∈ finset.range (p.nat_degree + 1), + { apply algebra.subset_adjoin, subst S, + rw [finset.mem_coe, finset.mem_image, coeff_map], + exact ⟨i, hi, rfl⟩ }, + { rw [finset.mem_range, not_lt] at hi, + rw [coeff_map, coeff_eq_zero_nat_degree_lt hi, alg_hom.map_zero], + exact subalgebra.zero_mem _ } }, + let q : polynomial (adjoin R S) := polynomial.lift _ (p.map $ to_comap R A B) coeffs_mem, + have hq : (q.map (algebra_map (comap R A B))) = (p.map $ to_comap R A B) := + map_coe_lift _ (p.map $ to_comap R A B) coeffs_mem, + use q, + split, + { suffices h : (q.map (algebra_map (comap R A B))).monic, + { refine @monic_of_injective _ _ _ _ _ _ _ + (by exact is_ring_hom.is_semiring_hom _) _ _ h, + exact subtype.val_injective }, + { erw map_coe_lift, exact monic_map _ pmonic } }, + { convert hp using 1, + replace hq := congr_arg (eval (comap.to_comap R A B x)) hq, + convert hq using 1; symmetry, swap, + exact eval_map _ _, + refine @eval_map _ _ _ _ _ _ _ _ (by exact is_ring_hom.is_semiring_hom _) _ }, +end + +lemma is_integral_trans (x : B) (hx : is_integral A x) : + is_integral R (comap.to_comap R A B x) := +begin + rcases hx with ⟨p, pmonic, hp⟩, + let S : set (comap R A B) := + (↑((finset.range (p.nat_degree + 1)).image + (λ i, to_comap R A B (p.coeff i))) : set (comap R A B)), + refine is_integral_of_mem_of_fg (adjoin R (S ∪ {comap.to_comap R A B x})) _ _ _, + swap, { apply subset_adjoin, simp }, + apply fg_trans, + { apply fg_adjoin_of_finite, { apply finset.finite_to_set }, + intros x hx, + rw [finset.mem_coe, finset.mem_image] at hx, + rcases hx with ⟨i, hi, rfl⟩, + rcases A_alg (p.coeff i) with ⟨q, qmonic, hq⟩, + use [q, qmonic], + replace hq := congr_arg (to_comap R A B : A → (comap R A B)) hq, + rw alg_hom.map_zero at hq, + convert hq using 1, + symmetry, exact polynomial.hom_eval₂ _ _ _ _ }, + { apply fg_adjoin_singleton_of_integral, + exact is_integral_trans_aux A_alg _ pmonic hp _ rfl } +end +. + +lemma algebraic_trans (B_alg : ∀ x : B, is_integral A x) : + ∀ x : comap R A B, is_integral R x := +λ x, is_integral_trans A_alg x (B_alg x) diff --git a/src/field_theory/algebraic_closure.lean b/src/field_theory/algebraic_closure.lean index 9368603e5c75c..346de29ac1beb 100644 --- a/src/field_theory/algebraic_closure.lean +++ b/src/field_theory/algebraic_closure.lean @@ -6,7 +6,7 @@ Authors: Chris Hughes import algebra.direct_limit import set_theory.schroeder_bernstein -import field_theory.minimal_polynomial +import field_theory.algebraic universes u v w open polynomial zorn set function @@ -42,59 +42,6 @@ begin equiv.apply_symm_apply, is_ring_hom.map_zero e.to_equiv, hx], end -namespace polynomial - -variables {α : Type*} {β : Type*} {γ : Type*} [comm_semiring α] [comm_semiring β] [comm_semiring γ] - -section -variables [decidable_eq α] -variables (f : α → β) (g : β → γ) [is_semiring_hom f] [is_semiring_hom g] (p : polynomial α) (x : β) - -lemma hom_eval₂ : g (p.eval₂ f x) = p.eval₂ (g ∘ f) (g x) := -begin - apply polynomial.induction_on p; clear p, - { intros a, rw [eval₂_C, eval₂_C] }, - { intros p q hp hq, simp only [hp, hq, eval₂_add, is_semiring_hom.map_add g] }, - { intros n a ih, - simp only [pow_succ', is_semiring_hom.map_mul g, (mul_assoc _ _ _).symm, - eval₂_C, eval₂_mul, eval₂_X] at ih ⊢, - rw ih } -end -end - -section injective -variables [decidable_eq α] [decidable_eq β] -variables {f : α → β} [is_semiring_hom f] - -lemma map_injective (hf : function.injective f) : - function.injective (map f : polynomial α → polynomial β) := -λ p q h, ext.mpr $ λ m, hf $ -begin - rw ext at h, - specialize h m, - rw [coeff_map f, coeff_map f] at h, - exact h -end - -lemma leading_coeff_of_injective (hf : injective f) (p : polynomial α) : - leading_coeff (p.map f) = f (leading_coeff p) := -begin - delta leading_coeff, - rw [coeff_map f, nat_degree_map' p hf] -end - -lemma monic_of_injective (hf : injective f) {p : polynomial α} (hp : (p.map f).monic) : p.monic := -begin - apply hf, - rw [← leading_coeff_of_injective hf, hp.leading_coeff, is_semiring_hom.map_one f] -end - -end injective --- def mk' (n : ℕ) (coeffs : fin n → α) : polynomial α := --- finset.univ.sum (λ i : fin n, C (coeffs i) * X^(i:ℕ)) - -end polynomial - namespace alg_hom variables {α : Type u} {β : Type v} {γ : Type w} [comm_ring α] [ring β] [ring γ] [algebra α β] [algebra α γ] (f : β →ₐ[α] γ) @@ -583,6 +530,8 @@ section adjoin_root variables {L : extension K} (f : polynomial L.carrier) [hif : irreducible f] include hif +open algebra + instance adjoin_root_algebraic_closure.field : discrete_field (adjoin_root f) := adjoin_root.field @@ -599,67 +548,8 @@ algebra.of_ring_hom (adjoin_root.of ∘ algebra_map _) (is_ring_hom.comp _ _) local attribute [instance] aux_instance --- set_option class.instance_max_depth 80 - lemma adjoin_root.algebraic (x : adjoin_root f) : is_integral K x := -begin - rcases adjoin_root.is_integral hif.ne_zero x with ⟨p, pmonic, hp⟩, - let S : set (adjoin_root f) := - (↑((finset.range (p.nat_degree + 1)).image (λ i, (p.coeff i : adjoin_root f))) : set (adjoin_root f)), - let B := algebra.adjoin K S, - have Bfg : submodule.fg (B : submodule K (adjoin_root f)), - { apply fg_adjoin_of_finite, - { apply finset.finite_to_set }, - { intros x hx, - rw [finset.mem_coe, finset.mem_image] at hx, - rcases hx with ⟨i, hi, rfl⟩, - rcases L.algebraic (p.coeff i) with ⟨q, qmonic, hq⟩, - use [q, qmonic], - replace hq := congr_arg (coe : L.carrier → adjoin_root f) hq, - convert hq using 1, - { symmetry, exact polynomial.hom_eval₂ _ _ _ _ }, - { symmetry, exact is_ring_hom.map_zero _ } } }, - refine is_integral_of_mem_of_fg (algebra.adjoin K (S ∪ {x})) _ x (algebra.subset_adjoin _), - { apply algebra.fg_trans Bfg, - apply fg_adjoin_singleton_of_integral, - let pB : polynomial _ := _, - refine ⟨pB, _⟩, swap, - { refine (finset.range (p.nat_degree + 1)).sum - (λ i, C ⟨(p.coeff i : adjoin_root f), _⟩ * X^i), - by_cases hi : i ∈ finset.range (p.nat_degree + 1), - { apply algebra.subset_adjoin, - rw [finset.mem_coe, finset.mem_image], - use [i, hi, rfl] }, - { rw [finset.mem_range, not_lt] at hi, - rw polynomial.coeff_eq_zero_nat_degree_lt hi, - rw [is_ring_hom.map_zero (coe : L.carrier → adjoin_root f)], - have : (0 : adjoin_root f) ∈ (B : submodule K (adjoin_root f)) := - submodule.zero_mem _, - rwa ← submodule.mem_coe at this } }, - let tmp : _ := _, - have hpB : pB.map (algebra_map _ : _ → adjoin_root f) = p.map (coe : L.carrier → adjoin_root f) := _, - { split, - { have := polynomial.monic_map (coe : L.carrier → adjoin_root f) pmonic, - rw ← hpB at this, - refine @polynomial.monic_of_injective - _ _ _ _ _ _ _ tmp _ _ this, - { exact subtype.val_injective } }, - { rwa [aeval_def, ← @eval_map _ _ _ _ _ _ _ _ tmp _, hpB, eval_map], } }, - { rw [p.as_sum, ← finset.sum_hom (map _), ← finset.sum_hom (map _)], - all_goals { try { apply_instance } }, - { apply finset.sum_congr rfl, - intros i hi, - rw [@polynomial.map_mul _ _ _ _ _ _ _ _ _ tmp, - @polynomial.map_C _ _ _ _ _ _ _ _ tmp], - rw [polynomial.map_mul, polynomial.map_C], - refine congr_arg _ _, - rw [is_monoid_hom.map_pow (map _), is_monoid_hom.map_pow (map _)], - { rw [@polynomial.map_X _ _ _ _ _ _ _ tmp, polynomial.map_X] }, - all_goals {sorry} }, - all_goals {sorry} }, - { exact ⟨rfl, rfl, λ _ _, rfl, λ _ _, rfl⟩ } }, - { exact mem_union_right _ (mem_singleton _) } -end +is_integral_trans L.algebraic x $ adjoin_root.is_integral hif.ne_zero x private def adjoin_root_extension_map : adjoin_root f ↪ big_type K := thing (adjoin_root.of_embedding f) ⟨subtype.val, subtype.val_injective⟩ From 6915af4eae0fe128549a9a3ee7faff06890b0fe7 Mon Sep 17 00:00:00 2001 From: Johan Commelin Date: Wed, 28 Aug 2019 23:21:44 +0200 Subject: [PATCH 15/51] More or less sorries --- src/field_theory/algebraic.lean | 43 +++- src/field_theory/algebraic_closure.lean | 239 +++++++++++++++++------ src/field_theory/minimal_polynomial.lean | 25 ++- 3 files changed, 239 insertions(+), 68 deletions(-) diff --git a/src/field_theory/algebraic.lean b/src/field_theory/algebraic.lean index de12ca092d1ef..f1335c8253bac 100644 --- a/src/field_theory/algebraic.lean +++ b/src/field_theory/algebraic.lean @@ -104,16 +104,57 @@ end lift end polynomial namespace subalgebra +open polynomial variables {R : Type*} {A : Type*} variables [comm_ring R] [comm_ring A] [algebra R A] lemma zero_mem (S : subalgebra R A) : (0 : A) ∈ S := submodule.zero_mem (S : submodule R A) -variable [decidable_eq A] +variables [decidable_eq R] [decidable_eq A] + +lemma is_integral (S : subalgebra R A) (x : A) (hx : is_integral R x) : + is_integral S x := +begin + rcases hx with ⟨p, pmonic, hp⟩, + use p.map (algebra_map S), + split, + { exact monic_map _ pmonic }, + { rwa [aeval_def, eval₂_map] } +end end subalgebra +namespace algebra +open set polynomial +variables {R : Type*} {A : Type*} {B : Type*} +variables [decidable_eq R] [decidable_eq A] [decidable_eq B] +variables [comm_ring R] [comm_ring A] [algebra R A] [comm_ring B] + +def adjoin_singleton_desc (x : A) (hx : is_integral R x) + (f : R → B) [is_ring_hom f] (y : B) (hy : is_root ((minimal_polynomial hx).map f) y) : +(adjoin R ({x} : set A) : Type _) → B := +_ + +instance adjoin_singleton_desc.is_ring_hom (x : A) (hx : is_integral R x) + (f : R → B) [is_ring_hom f] (y : B) (hy : is_root ((minimal_polynomial hx).map f) y) : + is_ring_hom (adjoin_singleton_desc x hx f y hy) := +_ + +end algebra + +-- namespace subalgebra +-- open set lattice +-- variables {R : Type*} {A : Type*} {B : Type*} +-- variables [comm_ring R] [comm_ring A] [algebra R A] [comm_ring B] [algebra R B] + +-- def Sup.desc (Ss : set (subalgebra R A)) (f : Π S : Ss, (S : subalgebra R A) →ₐ[R] B) +-- (hf : ∀ S₁ S₂ : Ss, ∃ h : (S₁ : subalgebra R A) ≤ S₂, (f S₂) ∘ inclusion h = f S₁) : +-- (Sup Ss : subalgebra R A) →ₐ[R] B := +-- sorry + +-- end subalgebra + open function algebra polynomial variables {R : Type*} {A : Type*} {B : Type*} variables [decidable_eq R] [decidable_eq A] [decidable_eq B] diff --git a/src/field_theory/algebraic_closure.lean b/src/field_theory/algebraic_closure.lean index 346de29ac1beb..87698aa57b3ab 100644 --- a/src/field_theory/algebraic_closure.lean +++ b/src/field_theory/algebraic_closure.lean @@ -1,7 +1,7 @@ /- Copyright (c) 2019 Chris Hughes. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Chris Hughes +Authors: Chris Hughes, Johan Commelin -/ import algebra.direct_limit @@ -225,6 +225,12 @@ lemma is_algebraically_closed_of_irreducible_has_root let ⟨i, hi⟩ := hg.2 in ⟨x, by rw [hi, is_root.def, eval_mul, show _ = _, from hx, zero_mul]⟩⟩ +variables (L : Type*) [nonzero_comm_ring L] [decidable_eq L] [is_algebraically_closed L] + +lemma exists_root (f : polynomial L) (hf : 0 < f.degree) : + ∃ x, is_root f x := +is_algebraically_closed.exists_root f hf + -- /- An algebraic extension of -/ -- lemma equiv_of_algebraic @@ -673,91 +679,196 @@ protected def is_integral : ∀ x : algebraic_closure K, is_integral K x := attribute [irreducible] algebraic_closure closed_extension algebraic_closure.algebra -section lift +namespace lift /- In this section, the homomorphism from any algebraic extension into an algebraically closed extension is proven to exist. The assumption that M is algebraically closed could probably easily be switched to an assumption that M contains all the roots of polynomials in K -/ variables {L : Type v} {M : Type w} [discrete_field L] [algebra K L] [discrete_field M] [algebra K M] [is_algebraically_closed M] (hL : ∀ x : L, is_integral K x) -/-- This structure is used to prove the existence of a homomorphism from any algebraic extension - into an algebraic closure -/ variables (K L M) include hL -private structure subfield_and_hom extends extension K := -( to_algebraically_closed : carrier →ₐ[K] M ) -( to_field : carrier →ₐ[K] L ) +/-- This structure is used to prove the existence of a homomorphism from any algebraic extension +into an algebraic closure -/ +private structure subfield_with_hom := +(carrier : subalgebra K L) +(emb : (@alg_hom K _ M _ _ _ (subalgebra.algebra carrier) _)) -variables {K L M} +variables {K L M hL} -instance subfield_and_hom.preorder : preorder (subfield_and_hom K L M hL) := -preorder.lift subfield_and_hom.to_extension (by apply_instance) +namespace subfield_with_hom +variables {E₁ E₂ E₃ : subfield_with_hom K L M hL} -private def maximal_subfield_and_hom_chain (c : set (subfield_and_hom K L M hL)) (hc : chain (≤) c) : - ∃ ub : subfield_and_hom K L M hL, ∀ N, N ∈ c → N ≤ ub := -let ub := (maximal_extension_chain (subfield_and_hom.to_extension '' c) (chain.image (≤) _ _ (λ _ _, id) hc)) in -⟨{ to_algebraically_closed := sorry, --field in question is direct limit of a bunch of fields with - --algebra homs into M - to_field := sorry, -- direct limit of a bunch of subfields is also a subfield - ..ub.1 }, - λ n hN, ub.2 _ (mem_image_of_mem _ hN)⟩ +instance : has_le (subfield_with_hom K L M hL) := +{ le := λ E₁ E₂, ∃ h : E₁.carrier ≤ E₂.carrier, E₂.emb ∘ inclusion h = E₁.emb } -private lemma exists_maximal_subfield_and_hom : ∃ N : subfield_and_hom K L M hL, - ∀ O, N ≤ O → O ≤ N := -zorn (maximal_subfield_and_hom_chain _) (λ _ _ _, le_trans) +lemma le_def : E₁ ≤ E₂ ↔ ∃ h : E₁.carrier ≤ E₂.carrier, E₂.emb ∘ inclusion h = E₁.emb := iff.rfl -variable (M) +lemma subalgebra_le_of_le (h : E₁ ≤ E₂) : E₁.carrier ≤ E₂.carrier := +by { rw le_def at h, cases h, assumption } -private def maximal_subfield_and_hom : subfield_and_hom K L M hL := -classical.some (exists_maximal_subfield_and_hom hL) +lemma compat (h : E₁ ≤ E₂) : E₂.emb ∘ inclusion (subalgebra_le_of_le h) = E₁.emb := +by { rw le_def at h, cases h, assumption } -instance akgh : algebra (maximal_subfield_and_hom M hL).carrier L := -algebra.of_ring_hom (maximal_subfield_and_hom M hL).to_field (by apply_instance) +instance : preorder (subfield_with_hom K L M hL) := +{ le := λ E₁ E₂, ∃ h : E₁.carrier ≤ E₂.carrier, E₂.emb ∘ inclusion h = E₁.emb, + le_refl := λ E, ⟨le_refl _, by rw [inclusion_refl, comp.right_id]⟩, + le_trans := λ E₁ E₂ E₃ h₁₂ h₂₃, ⟨le_trans (subalgebra_le_of_le h₁₂) (subalgebra_le_of_le h₂₃), + begin + erw inclusion_trans (subalgebra_le_of_le h₁₂) (subalgebra_le_of_le h₂₃), + rw [← comp.assoc, compat, compat] + end⟩ } --- Given K:L:M, if M is algebraic over K it is algebraic over L (names are different) -private lemma is_integral_over_maximal (x : L) : is_integral (maximal_subfield_and_hom M hL).carrier x := sorry +end subfield_with_hom -variables (f : polynomial (maximal_subfield_and_hom M hL).carrier) [hif : irreducible f] - {x : L} (hxf : f.eval₂ (maximal_subfield_and_hom M hL).to_field x = 0) +open lattice -include hif hxf +def maximal_subfield_with_hom_chain (c : set (subfield_with_hom K L M hL)) (hc : chain (≤) c) : + ∃ ub : subfield_with_hom K L M hL, ∀ N, N ∈ c → N ≤ ub := +let ub : subfield_with_hom K L M hL := +{ carrier := Sup (subfield_with_hom.carrier '' c), + emb := sorry } in +⟨ub, λ N hN, +begin + refine ⟨lattice.le_Sup ⟨N, hN, rfl⟩, _⟩, + sorry +end⟩ -private def adjoin_root_subfield_and_hom : subfield_and_hom K L M hL := -{ to_algebraically_closed := sorry, -- should be adjoin_root.lift composed with an isomorphism - to_field := sorry, -- - ..adjoin_root_extension K f } +variables (hL M) -private lemma le_adjoin_root_subfield_and_hom : maximal_subfield_and_hom M hL ≤ - adjoin_root_subfield_and_hom M hL f hxf := -le_adjoin_root_extension _ _ +lemma exists_maximal_subfield_with_hom : ∃ E : subfield_with_hom K L M hL, + ∀ N, E ≤ N → N ≤ E := +zorn (maximal_subfield_with_hom_chain) (λ _ _ _, le_trans) -private def maximal_subfield_and_hom_equiv_adjoin_root := -equiv_adjoin_root_of_le K f $ - classical.some_spec (exists_maximal_subfield_and_hom hL) _ - (le_adjoin_root_subfield_and_hom M hL f hxf) - -omit hif hxf - -private lemma surjective_maximal_subfield_and_hom_to_field : - function.surjective (maximal_subfield_and_hom M hL).to_field := -λ x, let hx := is_integral_over_maximal M hL x in -by letI := minimal_polynomial.irreducible hx; exact -⟨_, minimal_polynomial.root hx - (exists_root_of_equiv (maximal_subfield_and_hom_equiv_adjoin_root M hL _ - (minimal_polynomial.aeval hx)) (adjoin_root.eval₂_root _))⟩ - -private def equiv_maximal_subfield_and_hom : - (maximal_subfield_and_hom M hL).carrier ≃ₐ[K] L := -{ ..(maximal_subfield_and_hom M hL).to_field, - ..equiv.of_bijective - ⟨is_field_hom.injective _, surjective_maximal_subfield_and_hom_to_field _ _⟩ } - -/-- The hom from an algebraic extension of K into an algebraic closure -/ -def lift : L →ₐ[K] M := -(maximal_subfield_and_hom M hL).to_algebraically_closed.comp -(equiv_maximal_subfield_and_hom M hL).symm.to_alg_hom +def maximal_subfield_with_hom : subfield_with_hom K L M hL := +classical.some (exists_maximal_subfield_with_hom M hL) + +lemma maximal_subfield_with_hom_is_maximal : + ∀ (N : subfield_with_hom K L M hL), (maximal_subfield_with_hom M hL) ≤ N → N ≤ (maximal_subfield_with_hom M hL) := +classical.some_spec (exists_maximal_subfield_with_hom M hL) + +lemma maximal_subfield_with_hom_eq_top : + ((maximal_subfield_with_hom M hL).carrier : subalgebra K L) = (⊤ : subalgebra K L) := +begin + rw eq_top_iff, + rintros x hx, + replace hx := (maximal_subfield_with_hom M hL).carrier.is_integral x (hL x), + let p := minimal_polynomial hx, + have H := exists_root M (p.map (maximal_subfield_with_hom M hL).emb) _, + swap, { sorry }, + let y := classical.some H, + let f := algebra.adjoin_singleton_desc x hx (maximal_subfield_with_hom M hL).emb y (classical.some_spec H), + let tmpa : subalgebra _ L := algebra.adjoin _ ({x} : set L), + let tmpb : _ := _, + let N : subfield_with_hom K L M hL := + { carrier := subalgebra.under (maximal_subfield_with_hom M hL).carrier tmpa, + emb := + { to_fun := f, + hom := algebra.adjoin_singleton_desc.is_ring_hom _ _ _ _ _, + commutes' := tmpb } }, + have hN : x ∈ N.carrier := algebra.subset_adjoin (mem_singleton x), + refine subfield_with_hom.subalgebra_le_of_le (maximal_subfield_with_hom_is_maximal M hL N _) hN, + { refine ⟨λ l hl, ring.subset_closure (mem_union_left _ _), _⟩, + { rw mem_range, refine ⟨⟨l, hl⟩, rfl⟩ }, + { sorry } }, + { sorry } +end + +def lift_fun : L → M := +λ l, (maximal_subfield_with_hom M hL).emb.to_fun ⟨l, +by { rw maximal_subfield_with_hom_eq_top, show l ∈ (⊤ : subalgebra K L), sorry }⟩ end lift -end algebraic_closure +variables {L : Type v} {M : Type w} [discrete_field L] [algebra K L] + [discrete_field M] [algebra K M] [is_algebraically_closed M] (hL : ∀ x : L, is_integral K x) + +variables (K L M) +include hL + +/-- A (random) hom from an algebraic extension of K into an algebraic closure -/ +def lift : L →ₐ[K] M := +{ to_fun := lift.lift_fun M hL, + hom := sorry, + commutes' := sorry } + +-- #exit + +-- private structure subfield_and_hom extends extension K := +-- (to_algebraically_closed : carrier →ₐ[K] M) +-- (to_field : carrier →ₐ[K] L) + + +-- variables {K L M} + +-- instance subfield_and_hom.preorder : preorder (subfield_and_hom K L M hL) := +-- preorder.lift subfield_and_hom.to_extension (by apply_instance) + +-- private def maximal_subfield_and_hom_chain (c : set (subfield_and_hom K L M hL)) (hc : chain (≤) c) : +-- ∃ ub : subfield_and_hom K L M hL, ∀ N, N ∈ c → N ≤ ub := +-- let ub := (maximal_extension_chain (subfield_and_hom.to_extension '' c) (chain.image (≤) _ _ (λ _ _, id) hc)) in +-- ⟨{ to_algebraically_closed := sorry, --field in question is direct limit of a bunch of fields with +-- --algebra homs into M +-- to_field := sorry, -- direct limit of a bunch of subfields is also a subfield +-- ..ub.1 }, +-- λ n hN, ub.2 _ (mem_image_of_mem _ hN)⟩ + +-- private lemma exists_maximal_subfield_and_hom : ∃ N : subfield_and_hom K L M hL, +-- ∀ O, N ≤ O → O ≤ N := +-- zorn (maximal_subfield_and_hom_chain _) (λ _ _ _, le_trans) + +-- variable (M) + +-- private def maximal_subfield_and_hom : subfield_and_hom K L M hL := +-- classical.some (exists_maximal_subfield_and_hom hL) + +-- instance akgh : algebra (maximal_subfield_and_hom M hL).carrier L := +-- algebra.of_ring_hom (maximal_subfield_and_hom M hL).to_field (by apply_instance) + +-- -- Given K:L:M, if M is algebraic over K it is algebraic over L (names are different) +-- private lemma is_integral_over_maximal (x : L) : is_integral (maximal_subfield_and_hom M hL).carrier x := sorry + +-- variables (f : polynomial (maximal_subfield_and_hom M hL).carrier) [hif : irreducible f] +-- {x : L} (hxf : f.eval₂ (maximal_subfield_and_hom M hL).to_field x = 0) + +-- include hif hxf + +-- private def adjoin_root_subfield_and_hom : subfield_and_hom K L M hL := +-- { to_algebraically_closed := sorry, -- should be adjoin_root.lift composed with an isomorphism +-- to_field := sorry, -- +-- ..adjoin_root_extension K f } + +-- private lemma le_adjoin_root_subfield_and_hom : maximal_subfield_and_hom M hL ≤ +-- adjoin_root_subfield_and_hom M hL f hxf := +-- le_adjoin_root_extension _ _ + +-- private def maximal_subfield_and_hom_equiv_adjoin_root := +-- equiv_adjoin_root_of_le K f $ +-- classical.some_spec (exists_maximal_subfield_and_hom hL) _ +-- (le_adjoin_root_subfield_and_hom M hL f hxf) + +-- omit hif hxf + +-- private lemma surjective_maximal_subfield_and_hom_to_field : +-- function.surjective (maximal_subfield_and_hom M hL).to_field := +-- λ x, let hx := is_integral_over_maximal M hL x in +-- by letI := minimal_polynomial.irreducible hx; exact +-- ⟨_, minimal_polynomial.root hx +-- (exists_root_of_equiv (maximal_subfield_and_hom_equiv_adjoin_root M hL _ +-- (minimal_polynomial.aeval hx)) (adjoin_root.eval₂_root _))⟩ + +-- private def equiv_maximal_subfield_and_hom : +-- (maximal_subfield_and_hom M hL).carrier ≃ₐ[K] L := +-- { ..(maximal_subfield_and_hom M hL).to_field, +-- ..equiv.of_bijective +-- ⟨is_field_hom.injective _, surjective_maximal_subfield_and_hom_to_field _ _⟩ } + +-- /-- The hom from an algebraic extension of K into an algebraic closure -/ +-- def lift : L →ₐ[K] M := +-- (maximal_subfield_and_hom M hL).to_algebraically_closed.comp +-- (equiv_maximal_subfield_and_hom M hL).symm.to_alg_hom + +-- end lift + +-- end algebraic_closure diff --git a/src/field_theory/minimal_polynomial.lean b/src/field_theory/minimal_polynomial.lean index 68a929d0b7b0e..f4643d009e248 100644 --- a/src/field_theory/minimal_polynomial.lean +++ b/src/field_theory/minimal_polynomial.lean @@ -86,12 +86,20 @@ noncomputable theory to avoid diamonds -/ local attribute [instance, priority 0] subtype.decidable_eq -variables {α : Type u} {β : Type v} [discrete_field α] [discrete_field β] [algebra α β] +variables {α : Type u} {β : Type v} + +section min_poly_def +variables [decidable_eq α] [decidable_eq β] [comm_ring α] [comm_ring β] [algebra α β] def minimal_polynomial {x : β} (hx : is_integral α x) : polynomial α := well_founded.min polynomial.degree_lt_wf _ (ne_empty_iff_exists_mem.mpr hx) +end min_poly_def + namespace minimal_polynomial + +section ring +variables [decidable_eq α] [decidable_eq β] [comm_ring α] [comm_ring β] [algebra α β] variables {x : β} (hx : is_integral α x) lemma monic : monic (minimal_polynomial hx) := @@ -100,12 +108,18 @@ lemma monic : monic (minimal_polynomial hx) := @[simp] lemma aeval : aeval α β x (minimal_polynomial hx) = 0 := (well_founded.min_mem degree_lt_wf _ (ne_empty_iff_exists_mem.mpr hx)).2 -lemma ne_zero : (minimal_polynomial hx) ≠ 0 := ne_zero_of_monic (monic hx) - lemma min {p : polynomial α} (pmonic : p.monic) (hp : polynomial.aeval α β x p = 0) : degree (minimal_polynomial hx) ≤ degree p := le_of_not_lt $ well_founded.not_lt_min degree_lt_wf _ (ne_empty_iff_exists_mem.mpr hx) ⟨pmonic, hp⟩ +end ring + +section field +variables [discrete_field α] [discrete_field β] [algebra α β] +variables {x : β} (hx : is_integral α x) + +lemma ne_zero : (minimal_polynomial hx) ≠ 0 := ne_zero_of_monic (monic hx) + lemma degree_le_of_ne_zero {p : polynomial α} (pnz : p ≠ 0) (hp : polynomial.aeval α β x p = 0) : degree (minimal_polynomial hx) ≤ degree p := begin @@ -156,6 +170,9 @@ begin exact degree_eq_zero_of_is_unit H end +lemma degree_pos : 0 < degree (minimal_polynomial hx) := +degree_pos_of_ne_zero_of_nonunit (ne_zero hx) (not_is_unit hx) + lemma prime : prime (minimal_polynomial hx) := begin refine ⟨ne_zero hx, not_is_unit hx, _⟩, @@ -187,4 +204,6 @@ begin simpa [ndeg_one, finset.sum_range_succ, coeff_one, aeval_def] using H end +end field + end minimal_polynomial From a4fb3750824c9c7e1280f1d6aa2ade5795ce3f13 Mon Sep 17 00:00:00 2001 From: Johan Commelin Date: Thu, 29 Aug 2019 09:05:07 +0200 Subject: [PATCH 16/51] Removing some sorries --- src/field_theory/algebraic_closure.lean | 104 +++--------------------- 1 file changed, 12 insertions(+), 92 deletions(-) diff --git a/src/field_theory/algebraic_closure.lean b/src/field_theory/algebraic_closure.lean index 87698aa57b3ab..b3c6b4d3c5a8e 100644 --- a/src/field_theory/algebraic_closure.lean +++ b/src/field_theory/algebraic_closure.lean @@ -745,20 +745,21 @@ def maximal_subfield_with_hom : subfield_with_hom K L M hL := classical.some (exists_maximal_subfield_with_hom M hL) lemma maximal_subfield_with_hom_is_maximal : - ∀ (N : subfield_with_hom K L M hL), (maximal_subfield_with_hom M hL) ≤ N → N ≤ (maximal_subfield_with_hom M hL) := + ∀ (N : subfield_with_hom K L M hL), + (maximal_subfield_with_hom M hL) ≤ N → N ≤ (maximal_subfield_with_hom M hL) := classical.some_spec (exists_maximal_subfield_with_hom M hL) -lemma maximal_subfield_with_hom_eq_top : - ((maximal_subfield_with_hom M hL).carrier : subalgebra K L) = (⊤ : subalgebra K L) := +lemma maximal_subfield_with_hom_surj : + surjective (maximal_subfield_with_hom M hL).carrier.val := begin - rw eq_top_iff, - rintros x hx, + intros x, refine ⟨⟨x, _⟩, rfl⟩, replace hx := (maximal_subfield_with_hom M hL).carrier.is_integral x (hL x), let p := minimal_polynomial hx, have H := exists_root M (p.map (maximal_subfield_with_hom M hL).emb) _, swap, { sorry }, let y := classical.some H, - let f := algebra.adjoin_singleton_desc x hx (maximal_subfield_with_hom M hL).emb y (classical.some_spec H), + let f := algebra.adjoin_singleton_desc x hx + (maximal_subfield_with_hom M hL).emb y (classical.some_spec H), let tmpa : subalgebra _ L := algebra.adjoin _ ({x} : set L), let tmpb : _ := _, let N : subfield_with_hom K L M hL := @@ -775,10 +776,6 @@ begin { sorry } end -def lift_fun : L → M := -λ l, (maximal_subfield_with_hom M hL).emb.to_fun ⟨l, -by { rw maximal_subfield_with_hom_eq_top, show l ∈ (⊤ : subalgebra K L), sorry }⟩ - end lift variables {L : Type v} {M : Type w} [discrete_field L] [algebra K L] @@ -789,86 +786,9 @@ include hL /-- A (random) hom from an algebraic extension of K into an algebraic closure -/ def lift : L →ₐ[K] M := -{ to_fun := lift.lift_fun M hL, - hom := sorry, - commutes' := sorry } - --- #exit - --- private structure subfield_and_hom extends extension K := --- (to_algebraically_closed : carrier →ₐ[K] M) --- (to_field : carrier →ₐ[K] L) - - --- variables {K L M} - --- instance subfield_and_hom.preorder : preorder (subfield_and_hom K L M hL) := --- preorder.lift subfield_and_hom.to_extension (by apply_instance) - --- private def maximal_subfield_and_hom_chain (c : set (subfield_and_hom K L M hL)) (hc : chain (≤) c) : --- ∃ ub : subfield_and_hom K L M hL, ∀ N, N ∈ c → N ≤ ub := --- let ub := (maximal_extension_chain (subfield_and_hom.to_extension '' c) (chain.image (≤) _ _ (λ _ _, id) hc)) in --- ⟨{ to_algebraically_closed := sorry, --field in question is direct limit of a bunch of fields with --- --algebra homs into M --- to_field := sorry, -- direct limit of a bunch of subfields is also a subfield --- ..ub.1 }, --- λ n hN, ub.2 _ (mem_image_of_mem _ hN)⟩ - --- private lemma exists_maximal_subfield_and_hom : ∃ N : subfield_and_hom K L M hL, --- ∀ O, N ≤ O → O ≤ N := --- zorn (maximal_subfield_and_hom_chain _) (λ _ _ _, le_trans) - --- variable (M) - --- private def maximal_subfield_and_hom : subfield_and_hom K L M hL := --- classical.some (exists_maximal_subfield_and_hom hL) - --- instance akgh : algebra (maximal_subfield_and_hom M hL).carrier L := --- algebra.of_ring_hom (maximal_subfield_and_hom M hL).to_field (by apply_instance) - --- -- Given K:L:M, if M is algebraic over K it is algebraic over L (names are different) --- private lemma is_integral_over_maximal (x : L) : is_integral (maximal_subfield_and_hom M hL).carrier x := sorry - --- variables (f : polynomial (maximal_subfield_and_hom M hL).carrier) [hif : irreducible f] --- {x : L} (hxf : f.eval₂ (maximal_subfield_and_hom M hL).to_field x = 0) - --- include hif hxf - --- private def adjoin_root_subfield_and_hom : subfield_and_hom K L M hL := --- { to_algebraically_closed := sorry, -- should be adjoin_root.lift composed with an isomorphism --- to_field := sorry, -- --- ..adjoin_root_extension K f } - --- private lemma le_adjoin_root_subfield_and_hom : maximal_subfield_and_hom M hL ≤ --- adjoin_root_subfield_and_hom M hL f hxf := --- le_adjoin_root_extension _ _ - --- private def maximal_subfield_and_hom_equiv_adjoin_root := --- equiv_adjoin_root_of_le K f $ --- classical.some_spec (exists_maximal_subfield_and_hom hL) _ --- (le_adjoin_root_subfield_and_hom M hL f hxf) - --- omit hif hxf - --- private lemma surjective_maximal_subfield_and_hom_to_field : --- function.surjective (maximal_subfield_and_hom M hL).to_field := --- λ x, let hx := is_integral_over_maximal M hL x in --- by letI := minimal_polynomial.irreducible hx; exact --- ⟨_, minimal_polynomial.root hx --- (exists_root_of_equiv (maximal_subfield_and_hom_equiv_adjoin_root M hL _ --- (minimal_polynomial.aeval hx)) (adjoin_root.eval₂_root _))⟩ - --- private def equiv_maximal_subfield_and_hom : --- (maximal_subfield_and_hom M hL).carrier ≃ₐ[K] L := --- { ..(maximal_subfield_and_hom M hL).to_field, --- ..equiv.of_bijective --- ⟨is_field_hom.injective _, surjective_maximal_subfield_and_hom_to_field _ _⟩ } - --- /-- The hom from an algebraic extension of K into an algebraic closure -/ --- def lift : L →ₐ[K] M := --- (maximal_subfield_and_hom M hL).to_algebraically_closed.comp --- (equiv_maximal_subfield_and_hom M hL).symm.to_alg_hom - --- end lift +(lift.maximal_subfield_with_hom M hL).emb.comp $ alg_equiv.to_alg_hom $ alg_equiv.symm +{ ..(lift.maximal_subfield_with_hom M hL).carrier.val, + ..equiv.of_bijective + ⟨subtype.val_injective, lift.maximal_subfield_with_hom_surj _ _⟩ } --- end algebraic_closure +end algebraic_closure From 3079d38d838c0619b6169e3ea9f1d1b157ce6f7f Mon Sep 17 00:00:00 2001 From: Johan Commelin Date: Thu, 29 Aug 2019 14:25:10 +0200 Subject: [PATCH 17/51] WIP --- src/field_theory/algebraic.lean | 41 ++++++++++++++++++++++-- src/field_theory/algebraic_closure.lean | 22 ++++++++++++- src/field_theory/minimal_polynomial.lean | 33 ++++++++++++------- 3 files changed, 80 insertions(+), 16 deletions(-) diff --git a/src/field_theory/algebraic.lean b/src/field_theory/algebraic.lean index f1335c8253bac..9a7320482568e 100644 --- a/src/field_theory/algebraic.lean +++ b/src/field_theory/algebraic.lean @@ -2,6 +2,8 @@ import ring_theory.adjoin_root import ring_theory.integral_closure import field_theory.minimal_polynomial +universe variables u v w + namespace polynomial variables {α : Type*} {β : Type*} {γ : Type*} @@ -113,7 +115,7 @@ submodule.zero_mem (S : submodule R A) variables [decidable_eq R] [decidable_eq A] -lemma is_integral (S : subalgebra R A) (x : A) (hx : is_integral R x) : +protected lemma is_integral (S : subalgebra R A) (x : A) (hx : is_integral R x) : is_integral S x := begin rcases hx with ⟨p, pmonic, hp⟩, @@ -125,6 +127,39 @@ end end subalgebra +section subfield +variables {K : Type u} [discrete_field K] +variables (s : set K) [is_subring s] + +def subfield_of_inv_mem (h : ∀ x ∈ s, x⁻¹ ∈ s) : discrete_field s := +{ inv := λ x, ⟨x⁻¹, h x.val x.property⟩, + zero_ne_one := λ h, zero_ne_one $ congr_arg subtype.val h, + mul_inv_cancel := λ x hx, subtype.val_injective $ mul_inv_cancel + $ λ hz, hx $ subtype.val_injective hz, + inv_mul_cancel := λ x hx, subtype.val_injective $ inv_mul_cancel + $ λ hz, hx $ subtype.val_injective hz, + inv_zero := subtype.val_injective $ inv_zero, + has_decidable_eq := subtype.decidable_eq, + .. @subtype.comm_ring K _ s _ } + +end subfield + +namespace subalgebra +open polynomial +variables {K : Type u} {L : Type v} [discrete_field K] [discrete_field L] [algebra K L] + +noncomputable def inv_poly (x : L) (hx : is_integral K x) : polynomial K := +let p := minimal_polynomial hx in +((p - C (p.coeff 0)) /ₘ X) * C (p.coeff 0)⁻¹ + +variables (L_alg : ∀ x:L, is_integral K x) +include L_alg + +instance (E : subalgebra K L) : discrete_field (subtype E.carrier) := +subfield_of_inv_mem _ $ λ x (hx : x ∈ E), _ + +end subalgebra + namespace algebra open set polynomial variables {R : Type*} {A : Type*} {B : Type*} @@ -134,12 +169,12 @@ variables [comm_ring R] [comm_ring A] [algebra R A] [comm_ring B] def adjoin_singleton_desc (x : A) (hx : is_integral R x) (f : R → B) [is_ring_hom f] (y : B) (hy : is_root ((minimal_polynomial hx).map f) y) : (adjoin R ({x} : set A) : Type _) → B := -_ +sorry instance adjoin_singleton_desc.is_ring_hom (x : A) (hx : is_integral R x) (f : R → B) [is_ring_hom f] (y : B) (hy : is_root ((minimal_polynomial hx).map f) y) : is_ring_hom (adjoin_singleton_desc x hx f y hy) := -_ +sorry end algebra diff --git a/src/field_theory/algebraic_closure.lean b/src/field_theory/algebraic_closure.lean index b3c6b4d3c5a8e..925feb8b5fdb9 100644 --- a/src/field_theory/algebraic_closure.lean +++ b/src/field_theory/algebraic_closure.lean @@ -749,6 +749,16 @@ lemma maximal_subfield_with_hom_is_maximal : (maximal_subfield_with_hom M hL) ≤ N → N ≤ (maximal_subfield_with_hom M hL) := classical.some_spec (exists_maximal_subfield_with_hom M hL) +lemma emb_injective (E : subfield_with_hom K L M hL) : + injective E.emb := +begin + let tmpa : _ := _, let tmpb : _ := _, + refine @is_field_hom.injective _ M tmpa _ _ tmpb, + swap, + { }, + { exact { .. E.emb.hom } } +end + lemma maximal_subfield_with_hom_surj : surjective (maximal_subfield_with_hom M hL).carrier.val := begin @@ -756,7 +766,17 @@ begin replace hx := (maximal_subfield_with_hom M hL).carrier.is_integral x (hL x), let p := minimal_polynomial hx, have H := exists_root M (p.map (maximal_subfield_with_hom M hL).emb) _, - swap, { sorry }, + swap, + { calc 0 < degree p : + begin + sorry -- minimal_polynomial.degree_pos gives diamonds + end + ... = degree (p.map (maximal_subfield_with_hom M hL).emb) : + begin + symmetry, + refine @polynomial.degree_map' _ _ _ _ _ _ p _ (by exact is_ring_hom.is_semiring_hom _) _, + sorry, + end }, let y := classical.some H, let f := algebra.adjoin_singleton_desc x hx (maximal_subfield_with_hom M hL).emb y (classical.some_spec H), diff --git a/src/field_theory/minimal_polynomial.lean b/src/field_theory/minimal_polynomial.lean index f4643d009e248..960b5c0f24335 100644 --- a/src/field_theory/minimal_polynomial.lean +++ b/src/field_theory/minimal_polynomial.lean @@ -66,17 +66,18 @@ begin end end --- lemma explicit_form_of_degree_eq_one {α : Type*} [decidable_eq α] [comm_semiring α] --- {p : polynomial α} (hp : degree p = 1) : p = C p.leading_coeff * X + C (coeff p 0) := --- begin --- have ndeg : p.nat_degree = 1, --- { }, --- rw polynomial.ext, intro n, --- by_cases hn : n ≤ 1, --- { cases n, {simp}, --- obtain rfl : n = 0, { rwa [nat.succ_le_succ_iff, nat.le_zero_iff] at hn }, --- simp [hp, leading_coeff], } --- end +lemma zero_is_root_of_coeff_zero_eq_zero {p : polynomial α} (hp : p.coeff 0 = 0) : + is_root p 0 := +begin + rw p.as_sum, + calc _ = _ : (finset.sum_hom _).symm + ... = _ : finset.sum_eq_zero $ + begin + intros i hi, clear hi, + by_cases hi : i = 0, { rw [hi, hp, C_0, zero_mul, eval_zero] }, + rw [eval_mul, eval_pow, eval_X, zero_pow (nat.pos_of_ne_zero hi), mul_zero] + end +end end polynomial @@ -186,7 +187,7 @@ lemma irreducible : irreducible (minimal_polynomial hx) := irreducible_of_prime (prime hx) lemma root {x : β} (hx : is_integral α x) {y : α} - (h : (minimal_polynomial hx).eval y = 0) : algebra_map β y = x := + (h : is_root (minimal_polynomial hx) y) : algebra_map β y = x := begin have ndeg_one : nat_degree (minimal_polynomial hx) = 1, { rw ← polynomial.degree_eq_iff_nat_degree_eq_of_pos (nat.zero_lt_one), @@ -204,6 +205,14 @@ begin simpa [ndeg_one, finset.sum_range_succ, coeff_one, aeval_def] using H end +lemma coeff_zero_ne_zero (h : x ≠ 0) : coeff (minimal_polynomial hx) 0 ≠ 0 := +begin + contrapose! h, + have zero_root := polynomial.zero_is_root_of_coeff_zero_eq_zero h, + rw ← root hx zero_root, + exact is_ring_hom.map_zero _ +end + end field end minimal_polynomial From a81fb91c7553984ea08bf9562e1d3cdb42d2cbe4 Mon Sep 17 00:00:00 2001 From: Johan Commelin Date: Mon, 7 Oct 2019 15:02:35 +0200 Subject: [PATCH 18/51] Fix algebraic.lean --- src/field_theory/algebraic.lean | 33 +++++++++------------------------ 1 file changed, 9 insertions(+), 24 deletions(-) diff --git a/src/field_theory/algebraic.lean b/src/field_theory/algebraic.lean index 9a7320482568e..8f358e32512ad 100644 --- a/src/field_theory/algebraic.lean +++ b/src/field_theory/algebraic.lean @@ -5,28 +5,13 @@ import field_theory.minimal_polynomial universe variables u v w namespace polynomial - variables {α : Type*} {β : Type*} {γ : Type*} -section -variables [decidable_eq α] [comm_semiring α] [comm_semiring β] [comm_semiring γ] -variables (f : α → β) (g : β → γ) [is_semiring_hom f] [is_semiring_hom g] (p : polynomial α) (x : β) - -lemma hom_eval₂ : g (p.eval₂ f x) = p.eval₂ (g ∘ f) (g x) := -begin - apply polynomial.induction_on p; clear p, - { intros a, rw [eval₂_C, eval₂_C] }, - { intros p q hp hq, simp only [hp, hq, eval₂_add, is_semiring_hom.map_add g] }, - { intros n a ih, - simp only [pow_succ', is_semiring_hom.map_mul g, (mul_assoc _ _ _).symm, - eval₂_C, eval₂_mul, eval₂_X] at ih ⊢, - rw ih } -end - -end +noncomputable theory +open_locale classical section coeffs -variables [decidable_eq α] [comm_semiring α] (p : polynomial α) +variables [comm_semiring α] (p : polynomial α) def nonzero_coeffs : finset α := p.support.image p.coeff @@ -58,9 +43,9 @@ variables {f : α → β} [is_semiring_hom f] lemma map_injective (hf : function.injective f) : function.injective (map f : polynomial α → polynomial β) := -λ p q h, ext.mpr $ λ m, hf $ +λ p q h, ext $ λ m, hf $ begin - rw ext at h, + rw ext_iff at h, specialize h m, rw [coeff_map f, coeff_map f] at h, exact h @@ -91,7 +76,7 @@ def lift : polynomial s := lemma map_coe_lift : map coe (p.lift s hp) = p := begin conv_rhs {rw p.as_sum}, - rw ext, intro n, + rw ext_iff, intro n, rw [coeff_map, lift, coeff_sum', coeff_sum', ← finset.sum_hom coe], all_goals { try {apply_instance} }, apply finset.sum_congr rfl, @@ -156,7 +141,7 @@ variables (L_alg : ∀ x:L, is_integral K x) include L_alg instance (E : subalgebra K L) : discrete_field (subtype E.carrier) := -subfield_of_inv_mem _ $ λ x (hx : x ∈ E), _ +subfield_of_inv_mem _ $ λ x (hx : x ∈ E), sorry end subalgebra @@ -214,7 +199,7 @@ begin rw [finset.mem_coe, finset.mem_image, coeff_map], exact ⟨i, hi, rfl⟩ }, { rw [finset.mem_range, not_lt] at hi, - rw [coeff_map, coeff_eq_zero_nat_degree_lt hi, alg_hom.map_zero], + rw [coeff_map, coeff_eq_zero_of_nat_degree_lt hi, alg_hom.map_zero], exact subalgebra.zero_mem _ } }, let q : polynomial (adjoin R S) := polynomial.lift _ (p.map $ to_comap R A B) coeffs_mem, have hq : (q.map (algebra_map (comap R A B))) = (p.map $ to_comap R A B) := @@ -230,7 +215,7 @@ begin replace hq := congr_arg (eval (comap.to_comap R A B x)) hq, convert hq using 1; symmetry, swap, exact eval_map _ _, - refine @eval_map _ _ _ _ _ _ _ _ (by exact is_ring_hom.is_semiring_hom _) _ }, + refine @eval_map _ _ _ _ _ _ (by exact is_ring_hom.is_semiring_hom _) _ }, end lemma is_integral_trans (x : B) (hx : is_integral A x) : From 62e7e1d665e5e551272095a8431d42e529e1bbf0 Mon Sep 17 00:00:00 2001 From: Johan Commelin Date: Mon, 7 Oct 2019 20:58:09 +0200 Subject: [PATCH 19/51] Fix build, mostly --- src/field_theory/algebraic_closure.lean | 104 ++++++++++++++---------- 1 file changed, 62 insertions(+), 42 deletions(-) diff --git a/src/field_theory/algebraic_closure.lean b/src/field_theory/algebraic_closure.lean index 925feb8b5fdb9..19e063e106053 100644 --- a/src/field_theory/algebraic_closure.lean +++ b/src/field_theory/algebraic_closure.lean @@ -8,10 +8,12 @@ import algebra.direct_limit import set_theory.schroeder_bernstein import field_theory.algebraic +noncomputable theory +open_locale classical + universes u v w open polynomial zorn set function variables {K : Type u} [discrete_field K] -noncomputable theory /- Turn down the instance priority for subtype.decidable_eq and use classical.dec_eq everywhere, to avoid diamonds -/ @@ -28,18 +30,29 @@ instance equiv.is_ring_hom.symm {α β : Type*} [ring β] (e : α ≃ β) : @is_ring_hom α β (equiv.ring e) _ e := by split; simp [equiv.mul_def, equiv.add_def, equiv.one_def] -def equiv.ring_equiv {α β : Type*} [ring β] (e : α ≃ β) : - @ring_equiv α β (equiv.ring e) _ := -{ hom := by apply_instance, ..e } - -lemma exists_root_of_equiv {α β : Type*} [comm_ring α] [comm_ring β] [decidable_eq α] - [decidable_eq β] (e : α ≃r β) {f : polynomial α} {x : β} (hx : f.eval₂ e.to_equiv x = 0) : - f.eval (e.symm.to_equiv x) = 0 := +def equiv.add_equiv {α β : Type*} [has_add β] (e : α ≃ β) : + @add_equiv α β e.has_add _ := +{ map_add' := λ x y, e.apply_symm_apply _, + ..e } + +def equiv.mul_equiv {α β : Type*} [has_mul β] (e : α ≃ β) : + @mul_equiv α β e.has_mul _ := +{ map_mul' := λ x y, e.apply_symm_apply _, + ..e } + +def equiv.ring_equiv {α β : Type*} [has_mul β] [has_add β] (e : α ≃ β) : + @ring_equiv α β e.has_mul e.has_add _ _ := +{ map_add' := λ x y, e.apply_symm_apply _, + map_mul' := λ x y, e.apply_symm_apply _, + ..e } + +lemma exists_root_of_equiv {α β : Type*} [comm_ring α] [comm_ring β] (e : α ≃+* β) + {f : polynomial α} {x : β} (hx : f.eval₂ e x = 0) : + f.eval (e.symm x) = 0 := begin - letI : is_ring_hom e.to_equiv := e.hom, - rw [← e.to_equiv.injective.eq_iff, - ← eval₂_hom e.to_equiv, ring_equiv.to_equiv_symm, - equiv.apply_symm_apply, is_ring_hom.map_zero e.to_equiv, hx], + have e_inj : function.injective e := e.to_equiv.injective, + apply e_inj, + rw [← eval₂_hom e, e.apply_symm_apply, is_ring_hom.map_zero e, hx], end namespace alg_hom @@ -137,25 +150,28 @@ end open submodule +def adjoin_root_of_monic (f : polynomial K) : + adjoin_root (f * C (leading_coeff f)⁻¹) →ₐ[K] adjoin_root f := +{ to_fun := ideal.quotient.lift _ mk $ λ g hg, + begin + erw quotient.mk_eq_zero, + rw ideal.mem_span_singleton' at hg ⊢, + rcases hg with ⟨g, rfl⟩, rw [mul_comm f, ← mul_assoc], + exact ⟨_, rfl⟩, + end, + hom := ideal.quotient.is_ring_hom, + commutes' := λ g, rfl } + lemma fg {f : polynomial K} (hf : f ≠ 0) : submodule.fg (⊤ : submodule K (adjoin_root f)) := begin - let tmp : _ := _, - let ψ : adjoin_root (f * C (leading_coeff f)⁻¹) →ₗ[K] adjoin_root f := - { to_fun := ideal.quotient.lift _ mk tmp, - add := λ x y, is_ring_hom.map_add _, - smul := λ c x, is_ring_hom.map_mul _ }, + let φ := adjoin_root_of_monic f, have trick := fg_of_monic _ (monic_mul_leading_coeff_inv hf), - { convert fg_map trick, swap, - { exact ψ }, - { refine (submodule.eq_top_iff'.mpr _).symm, - intros x, apply quotient.induction_on' x, clear x, - intro g, - rw mem_map, - use [mk g, mem_top, rfl] } }, - { intros g hg, erw quotient.mk_eq_zero, - rw ideal.mem_span_singleton' at hg ⊢, - rcases hg with ⟨g, rfl⟩, rw [mul_comm f, ← mul_assoc], - exact ⟨_, rfl⟩ }, + convert fg_map trick, swap, exact φ.to_linear_map, + { refine (submodule.eq_top_iff'.mpr _).symm, + intros x, apply quotient.induction_on' x, clear x, + intro g, + rw mem_map, + use [mk g, mem_top, rfl] } end . @@ -521,7 +537,7 @@ if h : nonempty c rw aeval_def at hp ⊢, replace hp := congr_arg (inclusion (set.subset_Union (λ i : c, i.1.carrier) L')) hp, convert hp using 1; symmetry, - { rw polynomial.hom_eval₂ _ (inclusion _) p _, + { rw p.hom_eval₂ _ (inclusion _) _, congr' 1, { exact Union_compat c L L' }, { apply_instance, }, @@ -606,19 +622,22 @@ private lemma le_adjoin_root_extension : L ≤ adjoin_root_extension K f := end⟩ private def equiv_adjoin_root_of_le (h : adjoin_root_extension K f ≤ L) : - L.carrier ≃r adjoin_root f := + L.carrier ≃+* adjoin_root f := have left_inv : left_inverse (inclusion h.fst ∘ (equiv.set.range _ (adjoin_root_extension_map K f).2)) adjoin_root.of, - from λ _, by simp [adjoin_root_extension_map_apply, inclusion], +from λ _, by simp [adjoin_root_extension_map_apply, inclusion], +have hom : is_ring_hom (coe : L.carrier → adjoin_root f), by apply_instance, { to_fun := coe, inv_fun := inclusion h.fst ∘ (equiv.set.range _ (adjoin_root_extension_map K f).2), left_inv := left_inv, right_inv := right_inverse_of_injective_of_left_inverse (injective_comp (inclusion_injective _) (equiv.injective _)) left_inv, - hom := by apply_instance } + map_add' := hom.map_add, + map_mul' := hom.map_mul } -private def adjoin_root_equiv_adjoin_root_extension : adjoin_root f ≃r (adjoin_root_extension K f).carrier := +private def adjoin_root_equiv_adjoin_root_extension : + adjoin_root f ≃+* (adjoin_root_extension K f).carrier := (equiv.set.range _ (adjoin_root_extension_map K f).2).symm.ring_equiv.symm end adjoin_root @@ -645,12 +664,12 @@ include hif variable (K) -def algebraic_closure_equiv_adjoin_root : algebraic_closure K ≃r adjoin_root f := +def algebraic_closure_equiv_adjoin_root : algebraic_closure K ≃+* adjoin_root f := equiv_adjoin_root_of_le K f $ classical.some_spec (exists_algebraic_closure K) _ (le_adjoin_root_extension _ _) -instance ring_equiv.is_semiring_hom {α β : Type*} [ring α] [ring β] (e : α ≃r β) : - is_semiring_hom (e.to_equiv : α → β) := +instance ring_equiv.is_semiring_hom {α β : Type*} [ring α] [ring β] (e : α ≃+* β) : + is_semiring_hom (e : α → β) := is_ring_hom.is_semiring_hom _ omit hif @@ -752,11 +771,12 @@ classical.some_spec (exists_maximal_subfield_with_hom M hL) lemma emb_injective (E : subfield_with_hom K L M hL) : injective E.emb := begin - let tmpa : _ := _, let tmpb : _ := _, - refine @is_field_hom.injective _ M tmpa _ _ tmpb, - swap, - { }, - { exact { .. E.emb.hom } } + sorry + -- let tmpa : _ := _, let tmpb : _ := _, + -- refine @is_field_hom.injective _ M tmpa _ _ tmpb, + -- swap, + -- { sorry }, + -- { exact { .. E.emb.hom } } end lemma maximal_subfield_with_hom_surj : @@ -774,7 +794,7 @@ begin ... = degree (p.map (maximal_subfield_with_hom M hL).emb) : begin symmetry, - refine @polynomial.degree_map' _ _ _ _ _ _ p _ (by exact is_ring_hom.is_semiring_hom _) _, + -- refine @polynomial.degree_map' _ _ _ _ _ _ p _ (by exact is_ring_hom.is_semiring_hom _) _, sorry, end }, let y := classical.some H, From 57985f3d45a27a7247a25217d19a7c2c8817f2bd Mon Sep 17 00:00:00 2001 From: Chris Hughes Date: Sat, 4 Sep 2021 23:26:53 +0100 Subject: [PATCH 20/51] feat(algebra/algebra/subalgebra): inclusion --- src/algebra/algebra/subalgebra.lean | 15 +++++++++++++++ 1 file changed, 15 insertions(+) diff --git a/src/algebra/algebra/subalgebra.lean b/src/algebra/algebra/subalgebra.lean index 68ef6408a8d70..4f07b5c146321 100644 --- a/src/algebra/algebra/subalgebra.lean +++ b/src/algebra/algebra/subalgebra.lean @@ -648,6 +648,21 @@ instance : unique (subalgebra R R) := end .. algebra.subalgebra.inhabited } +def inclusion {S T : subalgebra R A} (h : S ≤ T) : S →ₐ[R] T := +{ to_fun := set.inclusion h, + map_one' := rfl, + map_add' := λ _ _, rfl, + map_mul' := λ _ _, rfl, + map_zero' := rfl, + commutes' := λ _, rfl } + +lemma inclusion_injective {S T : subalgebra R A} (h : S ≤ T) : + function.injective (inclusion h) := +λ _ _, subtype.ext ∘ subtype.mk.inj + +@[simp] lemma coe_inclusion {S T : subalgebra R A} (h : S ≤ T) (s : S) : + (inclusion h s : A) = s := rfl + /-- Two subalgebras that are equal are also equivalent as algebras. This is the `subalgebra` version of `linear_equiv.of_eq` and `equiv.set.of_eq`. -/ From 23dbec20400cfe113bbf313e2dc530a7d4b048c2 Mon Sep 17 00:00:00 2001 From: Chris Hughes <33847686+ChrisHughes24@users.noreply.github.com> Date: Sat, 4 Sep 2021 23:42:32 +0100 Subject: [PATCH 21/51] feat(algebra/algebra/subalgebra): inclusion map of subalgebras --- src/algebra/algebra/subalgebra.lean | 15 +++++++++++++++ 1 file changed, 15 insertions(+) diff --git a/src/algebra/algebra/subalgebra.lean b/src/algebra/algebra/subalgebra.lean index 68ef6408a8d70..4f07b5c146321 100644 --- a/src/algebra/algebra/subalgebra.lean +++ b/src/algebra/algebra/subalgebra.lean @@ -648,6 +648,21 @@ instance : unique (subalgebra R R) := end .. algebra.subalgebra.inhabited } +def inclusion {S T : subalgebra R A} (h : S ≤ T) : S →ₐ[R] T := +{ to_fun := set.inclusion h, + map_one' := rfl, + map_add' := λ _ _, rfl, + map_mul' := λ _ _, rfl, + map_zero' := rfl, + commutes' := λ _, rfl } + +lemma inclusion_injective {S T : subalgebra R A} (h : S ≤ T) : + function.injective (inclusion h) := +λ _ _, subtype.ext ∘ subtype.mk.inj + +@[simp] lemma coe_inclusion {S T : subalgebra R A} (h : S ≤ T) (s : S) : + (inclusion h s : A) = s := rfl + /-- Two subalgebras that are equal are also equivalent as algebras. This is the `subalgebra` version of `linear_equiv.of_eq` and `equiv.set.of_eq`. -/ From f4db504ab4bcf242b0e4830e33ca06bd12521e2c Mon Sep 17 00:00:00 2001 From: Chris Hughes <33847686+ChrisHughes24@users.noreply.github.com> Date: Sat, 4 Sep 2021 23:43:50 +0100 Subject: [PATCH 22/51] Update subalgebra.lean --- src/algebra/algebra/subalgebra.lean | 1 + 1 file changed, 1 insertion(+) diff --git a/src/algebra/algebra/subalgebra.lean b/src/algebra/algebra/subalgebra.lean index 4f07b5c146321..fea2b8f9bc1b1 100644 --- a/src/algebra/algebra/subalgebra.lean +++ b/src/algebra/algebra/subalgebra.lean @@ -648,6 +648,7 @@ instance : unique (subalgebra R R) := end .. algebra.subalgebra.inhabited } +/-- The map `S → T` when `S` is a subalgebra contained in the subalgebra `T`. -/ def inclusion {S T : subalgebra R A} (h : S ≤ T) : S →ₐ[R] T := { to_fun := set.inclusion h, map_one' := rfl, From 888fef8a7f4d194e87370cdee04334918e556c8b Mon Sep 17 00:00:00 2001 From: Chris Hughes Date: Sun, 5 Sep 2021 01:24:01 +0100 Subject: [PATCH 23/51] uncomment stuff --- src/field_theory/algebraic_closure.lean | 883 ++++++++++++++---------- 1 file changed, 509 insertions(+), 374 deletions(-) diff --git a/src/field_theory/algebraic_closure.lean b/src/field_theory/algebraic_closure.lean index c1a103e6a2422..d90eba46b75ca 100644 --- a/src/field_theory/algebraic_closure.lean +++ b/src/field_theory/algebraic_closure.lean @@ -1,3 +1,512 @@ +/- +Copyright (c) 2020 Kenny Lau. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Kenny Lau +-/ + +import algebra.direct_limit +import field_theory.splitting_field + +/-! +# Algebraic Closure + +In this file we define the typeclass for algebraically closed fields and algebraic closures. +We also construct an algebraic closure for any field. + +## Main Definitions + +- `is_alg_closed k` is the typeclass saying `k` is an algebraically closed field, i.e. every +polynomial in `k` splits. + +- `is_alg_closure k K` is the typeclass saying `K` is an algebraic closure of `k`. + +- `algebraic_closure k` is an algebraic closure of `k` (in the same universe). + It is constructed by taking the polynomial ring generated by indeterminates `x_f` + corresponding to monic irreducible polynomials `f` with coefficients in `k`, and quotienting + out by a maximal ideal containing every `f(x_f)`, and then repeating this step countably + many times. See Exercise 1.13 in Atiyah--Macdonald. + +## TODO + +Show that any algebraic extension embeds into any algebraically closed extension (via Zorn's lemma). + +## Tags + +algebraic closure, algebraically closed + +-/ + +universes u v w +noncomputable theory +open_locale classical big_operators +open polynomial + +variables (k : Type u) [field k] + +/-- Typeclass for algebraically closed fields. + +To show `polynomial.splits p f` for an arbitrary ring homomorphism `f`, +see `is_alg_closed.splits_codomain` and `is_alg_closed.splits_domain`. +-/ +class is_alg_closed : Prop := +(splits : ∀ p : polynomial k, p.splits $ ring_hom.id k) + +/-- Every polynomial splits in the field extension `f : K →+* k` if `k` is algebraically closed. + +See also `is_alg_closed.splits_domain` for the case where `K` is algebraically closed. +-/ +theorem is_alg_closed.splits_codomain {k K : Type*} [field k] [is_alg_closed k] [field K] + {f : K →+* k} (p : polynomial K) : p.splits f := +by { convert is_alg_closed.splits (p.map f), simp [splits_map_iff] } + +/-- Every polynomial splits in the field extension `f : K →+* k` if `K` is algebraically closed. + +See also `is_alg_closed.splits_codomain` for the case where `k` is algebraically closed. +-/ +theorem is_alg_closed.splits_domain {k K : Type*} [field k] [is_alg_closed k] [field K] + {f : k →+* K} (p : polynomial k) : p.splits f := +polynomial.splits_of_splits_id _ $ is_alg_closed.splits _ + +namespace is_alg_closed + +theorem of_exists_root (H : ∀ p : polynomial k, p.monic → irreducible p → ∃ x, p.eval x = 0) : + is_alg_closed k := +⟨λ p, or.inr $ λ q hq hqp, + have irreducible (q * C (leading_coeff q)⁻¹), + by { rw ← coe_norm_unit_of_ne_zero hq.ne_zero, + exact (associated_normalize _).irreducible hq }, + let ⟨x, hx⟩ := H (q * C (leading_coeff q)⁻¹) (monic_mul_leading_coeff_inv hq.ne_zero) this in + degree_mul_leading_coeff_inv q hq.ne_zero ▸ degree_eq_one_of_irreducible_of_root this hx⟩ + +lemma degree_eq_one_of_irreducible [is_alg_closed k] {p : polynomial k} (h_nz : p ≠ 0) + (hp : irreducible p) : + p.degree = 1 := +degree_eq_one_of_irreducible_of_splits h_nz hp (is_alg_closed.splits_codomain _) + +lemma algebra_map_surjective_of_is_integral {k K : Type*} [field k] [domain K] + [hk : is_alg_closed k] [algebra k K] (hf : algebra.is_integral k K) : + function.surjective (algebra_map k K) := +begin + refine λ x, ⟨-((minpoly k x).coeff 0), _⟩, + have hq : (minpoly k x).leading_coeff = 1 := minpoly.monic (hf x), + have h : (minpoly k x).degree = 1 := degree_eq_one_of_irreducible k + (minpoly.ne_zero (hf x)) (minpoly.irreducible (hf x)), + have : (aeval x (minpoly k x)) = 0 := minpoly.aeval k x, + rw [eq_X_add_C_of_degree_eq_one h, hq, C_1, one_mul, + aeval_add, aeval_X, aeval_C, add_eq_zero_iff_eq_neg] at this, + exact (ring_hom.map_neg (algebra_map k K) ((minpoly k x).coeff 0)).symm ▸ this.symm, +end + +lemma algebra_map_surjective_of_is_integral' {k K : Type*} [field k] [integral_domain K] + [hk : is_alg_closed k] (f : k →+* K) (hf : f.is_integral) : function.surjective f := +@algebra_map_surjective_of_is_integral k K _ _ _ f.to_algebra hf + +lemma algebra_map_surjective_of_is_algebraic {k K : Type*} [field k] [domain K] + [hk : is_alg_closed k] [algebra k K] (hf : algebra.is_algebraic k K) : + function.surjective (algebra_map k K) := +algebra_map_surjective_of_is_integral ((is_algebraic_iff_is_integral' k).mp hf) + +end is_alg_closed + +/-- Typeclass for an extension being an algebraic closure. -/ +class is_alg_closure (K : Type v) [field K] [algebra k K] : Prop := +(alg_closed : is_alg_closed K) +(algebraic : algebra.is_algebraic k K) + +theorem is_alg_closure_iff (K : Type v) [field K] [algebra k K] : + is_alg_closure k K ↔ is_alg_closed K ∧ algebra.is_algebraic k K := +⟨λ h, ⟨h.1, h.2⟩, λ h, ⟨h.1, h.2⟩⟩ + +namespace algebraic_closure + +open mv_polynomial + +/-- The subtype of monic irreducible polynomials -/ +@[reducible] def monic_irreducible : Type u := +{ f : polynomial k // monic f ∧ irreducible f } + +/-- Sends a monic irreducible polynomial `f` to `f(x_f)` where `x_f` is a formal indeterminate. -/ +def eval_X_self (f : monic_irreducible k) : mv_polynomial (monic_irreducible k) k := +polynomial.eval₂ mv_polynomial.C (X f) f + +/-- The span of `f(x_f)` across monic irreducible polynomials `f` where `x_f` is an +indeterminate. -/ +def span_eval : ideal (mv_polynomial (monic_irreducible k) k) := +ideal.span $ set.range $ eval_X_self k + +/-- Given a finset of monic irreducible polynomials, construct an algebra homomorphism to the +splitting field of the product of the polynomials sending each indeterminate `x_f` represented by +the polynomial `f` in the finset to a root of `f`. -/ +def to_splitting_field (s : finset (monic_irreducible k)) : + mv_polynomial (monic_irreducible k) k →ₐ[k] splitting_field (∏ x in s, x : polynomial k) := +mv_polynomial.aeval $ λ f, + if hf : f ∈ s + then root_of_splits _ + ((splits_prod_iff _ $ λ (j : monic_irreducible k) _, j.2.2.ne_zero).1 + (splitting_field.splits _) f hf) + (mt is_unit_iff_degree_eq_zero.2 f.2.2.not_unit) + else 37 + +theorem to_splitting_field_eval_X_self {s : finset (monic_irreducible k)} {f} (hf : f ∈ s) : + to_splitting_field k s (eval_X_self k f) = 0 := +by { rw [to_splitting_field, eval_X_self, ← alg_hom.coe_to_ring_hom, hom_eval₂, + alg_hom.coe_to_ring_hom, mv_polynomial.aeval_X, dif_pos hf, + ← algebra_map_eq, alg_hom.comp_algebra_map], + exact map_root_of_splits _ _ _ } + +theorem span_eval_ne_top : span_eval k ≠ ⊤ := +begin + rw [ideal.ne_top_iff_one, span_eval, ideal.span, ← set.image_univ, + finsupp.mem_span_image_iff_total], + rintros ⟨v, _, hv⟩, + replace hv := congr_arg (to_splitting_field k v.support) hv, + rw [alg_hom.map_one, finsupp.total_apply, finsupp.sum, alg_hom.map_sum, finset.sum_eq_zero] at hv, + { exact zero_ne_one hv }, + intros j hj, + rw [smul_eq_mul, alg_hom.map_mul, to_splitting_field_eval_X_self k hj, mul_zero] +end + +/-- A random maximal ideal that contains `span_eval k` -/ +def max_ideal : ideal (mv_polynomial (monic_irreducible k) k) := +classical.some $ ideal.exists_le_maximal _ $ span_eval_ne_top k + +instance max_ideal.is_maximal : (max_ideal k).is_maximal := +(classical.some_spec $ ideal.exists_le_maximal _ $ span_eval_ne_top k).1 + +theorem le_max_ideal : span_eval k ≤ max_ideal k := +(classical.some_spec $ ideal.exists_le_maximal _ $ span_eval_ne_top k).2 + +/-- The first step of constructing `algebraic_closure`: adjoin a root of all monic polynomials -/ +def adjoin_monic : Type u := +(max_ideal k).quotient + +instance adjoin_monic.field : field (adjoin_monic k) := +ideal.quotient.field _ + +instance adjoin_monic.inhabited : inhabited (adjoin_monic k) := ⟨37⟩ + +/-- The canonical ring homomorphism to `adjoin_monic k`. -/ +def to_adjoin_monic : k →+* adjoin_monic k := +(ideal.quotient.mk _).comp C + +instance adjoin_monic.algebra : algebra k (adjoin_monic k) := +(to_adjoin_monic k).to_algebra + +theorem adjoin_monic.algebra_map : algebra_map k (adjoin_monic k) = (ideal.quotient.mk _).comp C := +rfl + +theorem adjoin_monic.is_integral (z : adjoin_monic k) : is_integral k z := +let ⟨p, hp⟩ := ideal.quotient.mk_surjective z in hp ▸ +mv_polynomial.induction_on p (λ x, is_integral_algebra_map) (λ p q, is_integral_add) + (λ p f ih, @is_integral_mul _ _ _ _ _ _ (ideal.quotient.mk _ _) ih ⟨f, f.2.1, + by { erw [adjoin_monic.algebra_map, ← hom_eval₂, + ideal.quotient.eq_zero_iff_mem], + exact le_max_ideal k (ideal.subset_span ⟨f, rfl⟩) }⟩) + +theorem adjoin_monic.exists_root {f : polynomial k} (hfm : f.monic) (hfi : irreducible f) : + ∃ x : adjoin_monic k, f.eval₂ (to_adjoin_monic k) x = 0 := +⟨ideal.quotient.mk _ $ X (⟨f, hfm, hfi⟩ : monic_irreducible k), + by { rw [to_adjoin_monic, ← hom_eval₂, ideal.quotient.eq_zero_iff_mem], + exact le_max_ideal k (ideal.subset_span $ ⟨_, rfl⟩) }⟩ + +/-- The `n`th step of constructing `algebraic_closure`, together with its `field` instance. -/ +def step_aux (n : ℕ) : Σ α : Type u, field α := +nat.rec_on n ⟨k, infer_instance⟩ $ λ n ih, ⟨@adjoin_monic ih.1 ih.2, @adjoin_monic.field ih.1 ih.2⟩ + +/-- The `n`th step of constructing `algebraic_closure`. -/ +def step (n : ℕ) : Type u := +(step_aux k n).1 + +instance step.field (n : ℕ) : field (step k n) := +(step_aux k n).2 + +instance step.inhabited (n) : inhabited (step k n) := ⟨37⟩ + +/-- The canonical inclusion to the `0`th step. -/ +def to_step_zero : k →+* step k 0 := +ring_hom.id k + +/-- The canonical ring homomorphism to the next step. -/ +def to_step_succ (n : ℕ) : step k n →+* step k (n + 1) := +@to_adjoin_monic (step k n) (step.field k n) + +instance step.algebra_succ (n) : algebra (step k n) (step k (n + 1)) := +(to_step_succ k n).to_algebra + +theorem to_step_succ.exists_root {n} {f : polynomial (step k n)} + (hfm : f.monic) (hfi : irreducible f) : + ∃ x : step k (n + 1), f.eval₂ (to_step_succ k n) x = 0 := +@adjoin_monic.exists_root _ (step.field k n) _ hfm hfi + +/-- The canonical ring homomorphism to a step with a greater index. -/ +def to_step_of_le (m n : ℕ) (h : m ≤ n) : step k m →+* step k n := +{ to_fun := nat.le_rec_on h (λ n, to_step_succ k n), + map_one' := begin + induction h with n h ih, { exact nat.le_rec_on_self 1 }, + rw [nat.le_rec_on_succ h, ih, ring_hom.map_one] + end, + map_mul' := λ x y, begin + induction h with n h ih, { simp_rw nat.le_rec_on_self }, + simp_rw [nat.le_rec_on_succ h, ih, ring_hom.map_mul] + end, + map_zero' := begin + induction h with n h ih, { exact nat.le_rec_on_self 0 }, + rw [nat.le_rec_on_succ h, ih, ring_hom.map_zero] + end, + map_add' := λ x y, begin + induction h with n h ih, { simp_rw nat.le_rec_on_self }, + simp_rw [nat.le_rec_on_succ h, ih, ring_hom.map_add] + end } + +@[simp] lemma coe_to_step_of_le (m n : ℕ) (h : m ≤ n) : + (to_step_of_le k m n h : step k m → step k n) = nat.le_rec_on h (λ n, to_step_succ k n) := +rfl + +instance step.algebra (n) : algebra k (step k n) := +(to_step_of_le k 0 n n.zero_le).to_algebra + +instance step.scalar_tower (n) : is_scalar_tower k (step k n) (step k (n + 1)) := +is_scalar_tower.of_algebra_map_eq $ λ z, + @nat.le_rec_on_succ (step k) 0 n n.zero_le (n + 1).zero_le (λ n, to_step_succ k n) z + +theorem step.is_integral (n) : ∀ z : step k n, is_integral k z := +nat.rec_on n (λ z, is_integral_algebra_map) $ λ n ih z, + is_integral_trans ih _ (adjoin_monic.is_integral (step k n) z : _) + +instance to_step_of_le.directed_system : + directed_system (step k) (λ i j h, to_step_of_le k i j h) := +⟨λ i x h, nat.le_rec_on_self x, λ i₁ i₂ i₃ h₁₂ h₂₃ x, (nat.le_rec_on_trans h₁₂ h₂₃ x).symm⟩ + +end algebraic_closure + +/-- The canonical algebraic closure of a field, the direct limit of adding roots to the field for +each polynomial over the field. -/ +def algebraic_closure : Type u := +ring.direct_limit (algebraic_closure.step k) (λ i j h, algebraic_closure.to_step_of_le k i j h) + +namespace algebraic_closure + +instance : field (algebraic_closure k) := +field.direct_limit.field _ _ + +instance : inhabited (algebraic_closure k) := ⟨37⟩ + +/-- The canonical ring embedding from the `n`th step to the algebraic closure. -/ +def of_step (n : ℕ) : step k n →+* algebraic_closure k := +ring.direct_limit.of _ _ _ + +instance algebra_of_step (n) : algebra (step k n) (algebraic_closure k) := +(of_step k n).to_algebra + +theorem of_step_succ (n : ℕ) : (of_step k (n + 1)).comp (to_step_succ k n) = of_step k n := +ring_hom.ext $ λ x, show ring.direct_limit.of (step k) (λ i j h, to_step_of_le k i j h) _ _ = _, + by { convert ring.direct_limit.of_f n.le_succ x, ext x, exact (nat.le_rec_on_succ' x).symm } + +theorem exists_of_step (z : algebraic_closure k) : ∃ n x, of_step k n x = z := +ring.direct_limit.exists_of z + +-- slow +theorem exists_root {f : polynomial (algebraic_closure k)} + (hfm : f.monic) (hfi : irreducible f) : + ∃ x : algebraic_closure k, f.eval x = 0 := +begin + have : ∃ n p, polynomial.map (of_step k n) p = f, + { convert ring.direct_limit.polynomial.exists_of f }, + unfreezingI { obtain ⟨n, p, rfl⟩ := this }, + rw monic_map_iff at hfm, + have := hfm.irreducible_of_irreducible_map (of_step k n) p hfi, + obtain ⟨x, hx⟩ := to_step_succ.exists_root k hfm this, + refine ⟨of_step k (n + 1) x, _⟩, + rw [← of_step_succ k n, eval_map, ← hom_eval₂, hx, ring_hom.map_zero] +end + +instance : is_alg_closed (algebraic_closure k) := +is_alg_closed.of_exists_root _ $ λ f, exists_root k + +instance {R : Type*} [comm_semiring R] [alg : algebra R k] : + algebra R (algebraic_closure k) := +((of_step k 0).comp (@algebra_map _ _ _ _ alg)).to_algebra + +lemma algebra_map_def {R : Type*} [comm_semiring R] [alg : algebra R k] : + algebra_map R (algebraic_closure k) = ((of_step k 0 : k →+* _).comp (@algebra_map _ _ _ _ alg)) := +rfl + +instance {R S : Type*} [comm_semiring R] [comm_semiring S] + [algebra R S] [algebra S k] [algebra R k] [is_scalar_tower R S k] : + is_scalar_tower R S (algebraic_closure k) := +is_scalar_tower.of_algebra_map_eq (λ x, + ring_hom.congr_arg _ (is_scalar_tower.algebra_map_apply R S k x : _)) + +/-- Canonical algebra embedding from the `n`th step to the algebraic closure. -/ +def of_step_hom (n) : step k n →ₐ[k] algebraic_closure k := +{ commutes' := λ x, ring.direct_limit.of_f n.zero_le x, + .. of_step k n } + +theorem is_algebraic : algebra.is_algebraic k (algebraic_closure k) := +λ z, (is_algebraic_iff_is_integral _).2 $ let ⟨n, x, hx⟩ := exists_of_step k z in +hx ▸ is_integral_alg_hom (of_step_hom k n) (step.is_integral k n x) + +instance : is_alg_closure k (algebraic_closure k) := +⟨algebraic_closure.is_alg_closed k, is_algebraic k⟩ + +end algebraic_closure + +/-- +Every element `f` in a nontrivial finite-dimensional algebra `A` +over an algebraically closed field `K` +has non-empty spectrum: +that is, there is some `c : K` so `f - c • 1` is not invertible. +-/ +-- We will use this both to show eigenvalues exist, and to prove Schur's lemma. +lemma exists_spectrum_of_is_alg_closed_of_finite_dimensional (𝕜 : Type*) [field 𝕜] [is_alg_closed 𝕜] + {A : Type*} [nontrivial A] [ring A] [algebra 𝕜 A] [I : finite_dimensional 𝕜 A] (f : A) : + ∃ c : 𝕜, ¬ is_unit (f - algebra_map 𝕜 A c) := +begin + obtain ⟨p, ⟨h_mon, h_eval_p⟩⟩ := is_integral_of_noetherian I f, + have nu : ¬ is_unit (aeval f p), { rw [←aeval_def] at h_eval_p, rw h_eval_p, simp, }, + rw [eq_prod_roots_of_monic_of_splits_id h_mon (is_alg_closed.splits p), + ←multiset.prod_to_list, alg_hom.map_list_prod] at nu, + replace nu := mt list.prod_is_unit nu, + simp only [not_forall, exists_prop, aeval_C, multiset.mem_to_list, + list.mem_map, aeval_X, exists_exists_and_eq_and, multiset.mem_map, alg_hom.map_sub] at nu, + exact ⟨nu.some, nu.some_spec.2⟩, +end + +namespace lift +/- In this section, the homomorphism from any algebraic extension into an algebraically + closed extension is proven to exist. The assumption that M is algebraically closed could probably + easily be switched to an assumption that M contains all the roots of polynomials in K -/ +variables {K : Type u} {L : Type v} {M : Type w} [field K] [field L] [algebra K L] + [field M] [algebra K M] [is_alg_closed M] (hL : ∀ x : L, is_integral K x) + +variables (K L M) +include hL +open zorn subalgebra alg_hom function + +/-- This structure is used to prove the existence of a homomorphism from any algebraic extension +into an algebraic closure -/ +private structure subfield_with_hom := +(carrier : subalgebra K L) +(emb : (@alg_hom K _ M _ _ _ (subalgebra.algebra carrier) _)) + +variables {K L M hL} + +namespace subfield_with_hom +variables {E₁ E₂ E₃ : subfield_with_hom K L M hL} + +instance : has_le (subfield_with_hom K L M hL) := +{ le := λ E₁ E₂, ∃ h : E₁.carrier ≤ E₂.carrier, E₂.emb ∘ inclusion h = E₁.emb } + +lemma le_def : E₁ ≤ E₂ ↔ ∃ h : E₁.carrier ≤ E₂.carrier, E₂.emb ∘ inclusion h = E₁.emb := iff.rfl + +lemma subalgebra_le_of_le (h : E₁ ≤ E₂) : E₁.carrier ≤ E₂.carrier := +by { rw le_def at h, cases h, assumption } + +lemma compat (h : E₁ ≤ E₂) : E₂.emb ∘ inclusion (subalgebra_le_of_le h) = E₁.emb := +by { rw le_def at h, cases h, assumption } + +instance : preorder (subfield_with_hom K L M hL) := +{ le := λ E₁ E₂, ∃ h : E₁.carrier ≤ E₂.carrier, E₂.emb ∘ inclusion h = E₁.emb, + le_refl := λ E, ⟨le_refl _, by rw [inclusion_refl, comp.right_id]⟩, + le_trans := λ E₁ E₂ E₃ h₁₂ h₂₃, ⟨le_trans (subalgebra_le_of_le h₁₂) (subalgebra_le_of_le h₂₃), + begin + erw inclusion_trans (subalgebra_le_of_le h₁₂) (subalgebra_le_of_le h₂₃), + rw [← comp.assoc, compat, compat] + end⟩ } + +end subfield_with_hom + +open lattice + +def maximal_subfield_with_hom_chain (c : set (subfield_with_hom K L M hL)) (hc : chain (≤) c) : + ∃ ub : subfield_with_hom K L M hL, ∀ N, N ∈ c → N ≤ ub := +let ub : subfield_with_hom K L M hL := +{ carrier := Sup (subfield_with_hom.carrier '' c), + emb := } in +⟨ub, λ N hN, +begin + refine ⟨le_Sup ⟨N, hN, rfl⟩, _⟩, + sorry +end⟩ + +variables (hL M) + +lemma exists_maximal_subfield_with_hom : ∃ E : subfield_with_hom K L M hL, + ∀ N, E ≤ N → N ≤ E := +zorn (maximal_subfield_with_hom_chain) (λ _ _ _, le_trans) + +def maximal_subfield_with_hom : subfield_with_hom K L M hL := +classical.some (exists_maximal_subfield_with_hom M hL) + +lemma maximal_subfield_with_hom_is_maximal : + ∀ (N : subfield_with_hom K L M hL), + (maximal_subfield_with_hom M hL) ≤ N → N ≤ (maximal_subfield_with_hom M hL) := +classical.some_spec (exists_maximal_subfield_with_hom M hL) + +lemma emb_injective (E : subfield_with_hom K L M hL) : + injective E.emb := +begin + sorry + -- let tmpa : _ := _, let tmpb : _ := _, + -- refine @is_field_hom.injective _ M tmpa _ _ tmpb, + -- swap, + -- { sorry }, + -- { exact { .. E.emb.hom } } +end + +lemma maximal_subfield_with_hom_surj : + surjective (maximal_subfield_with_hom M hL).carrier.val := +begin + intros x, refine ⟨⟨x, _⟩, rfl⟩, + replace hx := (maximal_subfield_with_hom M hL).carrier.is_integral x (hL x), + let p := minimal_polynomial hx, + have H := exists_root M (p.map (maximal_subfield_with_hom M hL).emb) _, + swap, + { calc 0 < degree p : + begin + sorry -- minimal_polynomial.degree_pos gives diamonds + end + ... = degree (p.map (maximal_subfield_with_hom M hL).emb) : + begin + symmetry, + -- refine @polynomial.degree_map' _ _ _ _ _ _ p _ (by exact is_ring_hom.is_semiring_hom _) _, + sorry, + end }, + let y := classical.some H, + let f := algebra.adjoin_singleton_desc x hx + (maximal_subfield_with_hom M hL).emb y (classical.some_spec H), + let tmpa : subalgebra _ L := algebra.adjoin _ ({x} : set L), + let tmpb : _ := _, + let N : subfield_with_hom K L M hL := + { carrier := subalgebra.under (maximal_subfield_with_hom M hL).carrier tmpa, + emb := + { to_fun := f, + hom := algebra.adjoin_singleton_desc.is_ring_hom _ _ _ _ _, + commutes' := tmpb } }, + have hN : x ∈ N.carrier := algebra.subset_adjoin (mem_singleton x), + refine subfield_with_hom.subalgebra_le_of_le (maximal_subfield_with_hom_is_maximal M hL N _) hN, + { refine ⟨λ l hl, ring.subset_closure (mem_union_left _ _), _⟩, + { rw mem_range, refine ⟨⟨l, hl⟩, rfl⟩ }, + { sorry } }, + { sorry } +end + +end lift + +variables {L : Type v} {M : Type w} [discrete_field L] [algebra K L] + [discrete_field M] [algebra K M] [is_algebraically_closed M] (hL : ∀ x : L, is_integral K x) + +variables (K L M) +include hL + +/-- A (random) hom from an algebraic extension of K into an algebraic closure -/ +def lift : L →ₐ[K] M := +(lift.maximal_subfield_with_hom M hL).emb.comp $ alg_equiv.to_alg_hom $ alg_equiv.symm +{ ..(lift.maximal_subfield_with_hom M hL).carrier.val, + ..equiv.of_bijective + ⟨subtype.val_injective, lift.maximal_subfield_with_hom_surj _ _⟩ } + + -- /- -- Copyright (c) 2019 Chris Hughes. All rights reserved. -- Released under Apache 2.0 license as described in the file LICENSE. @@ -830,377 +1339,3 @@ -- { ..(lift.maximal_subfield_with_hom M hL).carrier.val, -- ..equiv.of_bijective -- ⟨subtype.val_injective, lift.maximal_subfield_with_hom_surj _ _⟩ } - --- end algebraic_closure --- Copyright (c) 2020 Kenny Lau. All rights reserved. --- Released under Apache 2.0 license as described in the file LICENSE. --- Authors: Kenny Lau --- -/ - --- import algebra.direct_limit --- import field_theory.splitting_field - --- /-! --- # Algebraic Closure - --- In this file we define the typeclass for algebraically closed fields and algebraic closures. --- We also construct an algebraic closure for any field. - --- ## Main Definitions - --- - `is_alg_closed k` is the typeclass saying `k` is an algebraically closed field, i.e. every --- polynomial in `k` splits. - --- - `is_alg_closure k K` is the typeclass saying `K` is an algebraic closure of `k`. - --- - `algebraic_closure k` is an algebraic closure of `k` (in the same universe). --- It is constructed by taking the polynomial ring generated by indeterminates `x_f` --- corresponding to monic irreducible polynomials `f` with coefficients in `k`, and quotienting --- out by a maximal ideal containing every `f(x_f)`, and then repeating this step countably --- many times. See Exercise 1.13 in Atiyah--Macdonald. - --- ## TODO - --- Show that any algebraic extension embeds into any algebraically closed extension (via Zorn's lemma). - --- ## Tags - --- algebraic closure, algebraically closed - --- -/ - --- universes u v w --- noncomputable theory --- open_locale classical big_operators --- open polynomial - --- variables (k : Type u) [field k] - --- /-- Typeclass for algebraically closed fields. - --- To show `polynomial.splits p f` for an arbitrary ring homomorphism `f`, --- see `is_alg_closed.splits_codomain` and `is_alg_closed.splits_domain`. --- -/ --- class is_alg_closed : Prop := --- (splits : ∀ p : polynomial k, p.splits $ ring_hom.id k) - --- /-- Every polynomial splits in the field extension `f : K →+* k` if `k` is algebraically closed. - --- See also `is_alg_closed.splits_domain` for the case where `K` is algebraically closed. --- -/ --- theorem is_alg_closed.splits_codomain {k K : Type*} [field k] [is_alg_closed k] [field K] --- {f : K →+* k} (p : polynomial K) : p.splits f := --- by { convert is_alg_closed.splits (p.map f), simp [splits_map_iff] } - --- /-- Every polynomial splits in the field extension `f : K →+* k` if `K` is algebraically closed. - --- See also `is_alg_closed.splits_codomain` for the case where `k` is algebraically closed. --- -/ --- theorem is_alg_closed.splits_domain {k K : Type*} [field k] [is_alg_closed k] [field K] --- {f : k →+* K} (p : polynomial k) : p.splits f := --- polynomial.splits_of_splits_id _ $ is_alg_closed.splits _ - --- namespace is_alg_closed - --- theorem of_exists_root (H : ∀ p : polynomial k, p.monic → irreducible p → ∃ x, p.eval x = 0) : --- is_alg_closed k := --- ⟨λ p, or.inr $ λ q hq hqp, --- have irreducible (q * C (leading_coeff q)⁻¹), --- by { rw ← coe_norm_unit_of_ne_zero hq.ne_zero, --- exact (associated_normalize _).irreducible hq }, --- let ⟨x, hx⟩ := H (q * C (leading_coeff q)⁻¹) (monic_mul_leading_coeff_inv hq.ne_zero) this in --- degree_mul_leading_coeff_inv q hq.ne_zero ▸ degree_eq_one_of_irreducible_of_root this hx⟩ - --- lemma degree_eq_one_of_irreducible [is_alg_closed k] {p : polynomial k} (h_nz : p ≠ 0) --- (hp : irreducible p) : --- p.degree = 1 := --- degree_eq_one_of_irreducible_of_splits h_nz hp (is_alg_closed.splits_codomain _) - --- lemma algebra_map_surjective_of_is_integral {k K : Type*} [field k] [domain K] --- [hk : is_alg_closed k] [algebra k K] (hf : algebra.is_integral k K) : --- function.surjective (algebra_map k K) := --- begin --- refine λ x, ⟨-((minpoly k x).coeff 0), _⟩, --- have hq : (minpoly k x).leading_coeff = 1 := minpoly.monic (hf x), --- have h : (minpoly k x).degree = 1 := degree_eq_one_of_irreducible k --- (minpoly.ne_zero (hf x)) (minpoly.irreducible (hf x)), --- have : (aeval x (minpoly k x)) = 0 := minpoly.aeval k x, --- rw [eq_X_add_C_of_degree_eq_one h, hq, C_1, one_mul, --- aeval_add, aeval_X, aeval_C, add_eq_zero_iff_eq_neg] at this, --- exact (ring_hom.map_neg (algebra_map k K) ((minpoly k x).coeff 0)).symm ▸ this.symm, --- end - --- lemma algebra_map_surjective_of_is_integral' {k K : Type*} [field k] [integral_domain K] --- [hk : is_alg_closed k] (f : k →+* K) (hf : f.is_integral) : function.surjective f := --- @algebra_map_surjective_of_is_integral k K _ _ _ f.to_algebra hf - --- lemma algebra_map_surjective_of_is_algebraic {k K : Type*} [field k] [domain K] --- [hk : is_alg_closed k] [algebra k K] (hf : algebra.is_algebraic k K) : --- function.surjective (algebra_map k K) := --- algebra_map_surjective_of_is_integral ((is_algebraic_iff_is_integral' k).mp hf) - --- end is_alg_closed - --- /-- Typeclass for an extension being an algebraic closure. -/ --- class is_alg_closure (K : Type v) [field K] [algebra k K] : Prop := --- (alg_closed : is_alg_closed K) --- (algebraic : algebra.is_algebraic k K) - --- theorem is_alg_closure_iff (K : Type v) [field K] [algebra k K] : --- is_alg_closure k K ↔ is_alg_closed K ∧ algebra.is_algebraic k K := --- ⟨λ h, ⟨h.1, h.2⟩, λ h, ⟨h.1, h.2⟩⟩ - --- namespace algebraic_closure - --- open mv_polynomial - --- /-- The subtype of monic irreducible polynomials -/ --- @[reducible] def monic_irreducible : Type u := --- { f : polynomial k // monic f ∧ irreducible f } - --- /-- Sends a monic irreducible polynomial `f` to `f(x_f)` where `x_f` is a formal indeterminate. -/ --- def eval_X_self (f : monic_irreducible k) : mv_polynomial (monic_irreducible k) k := --- polynomial.eval₂ mv_polynomial.C (X f) f - --- /-- The span of `f(x_f)` across monic irreducible polynomials `f` where `x_f` is an --- indeterminate. -/ --- def span_eval : ideal (mv_polynomial (monic_irreducible k) k) := --- ideal.span $ set.range $ eval_X_self k - --- /-- Given a finset of monic irreducible polynomials, construct an algebra homomorphism to the --- splitting field of the product of the polynomials sending each indeterminate `x_f` represented by --- the polynomial `f` in the finset to a root of `f`. -/ --- def to_splitting_field (s : finset (monic_irreducible k)) : --- mv_polynomial (monic_irreducible k) k →ₐ[k] splitting_field (∏ x in s, x : polynomial k) := --- mv_polynomial.aeval $ λ f, --- if hf : f ∈ s --- then root_of_splits _ --- ((splits_prod_iff _ $ λ (j : monic_irreducible k) _, j.2.2.ne_zero).1 --- (splitting_field.splits _) f hf) --- (mt is_unit_iff_degree_eq_zero.2 f.2.2.not_unit) --- else 37 - --- theorem to_splitting_field_eval_X_self {s : finset (monic_irreducible k)} {f} (hf : f ∈ s) : --- to_splitting_field k s (eval_X_self k f) = 0 := --- by { rw [to_splitting_field, eval_X_self, ← alg_hom.coe_to_ring_hom, hom_eval₂, --- alg_hom.coe_to_ring_hom, mv_polynomial.aeval_X, dif_pos hf, --- ← algebra_map_eq, alg_hom.comp_algebra_map], --- exact map_root_of_splits _ _ _ } - --- theorem span_eval_ne_top : span_eval k ≠ ⊤ := --- begin --- rw [ideal.ne_top_iff_one, span_eval, ideal.span, ← set.image_univ, --- finsupp.mem_span_image_iff_total], --- rintros ⟨v, _, hv⟩, --- replace hv := congr_arg (to_splitting_field k v.support) hv, --- rw [alg_hom.map_one, finsupp.total_apply, finsupp.sum, alg_hom.map_sum, finset.sum_eq_zero] at hv, --- { exact zero_ne_one hv }, --- intros j hj, --- rw [smul_eq_mul, alg_hom.map_mul, to_splitting_field_eval_X_self k hj, mul_zero] --- end - --- /-- A random maximal ideal that contains `span_eval k` -/ --- def max_ideal : ideal (mv_polynomial (monic_irreducible k) k) := --- classical.some $ ideal.exists_le_maximal _ $ span_eval_ne_top k - --- instance max_ideal.is_maximal : (max_ideal k).is_maximal := --- (classical.some_spec $ ideal.exists_le_maximal _ $ span_eval_ne_top k).1 - --- theorem le_max_ideal : span_eval k ≤ max_ideal k := --- (classical.some_spec $ ideal.exists_le_maximal _ $ span_eval_ne_top k).2 - --- /-- The first step of constructing `algebraic_closure`: adjoin a root of all monic polynomials -/ --- def adjoin_monic : Type u := --- (max_ideal k).quotient - --- instance adjoin_monic.field : field (adjoin_monic k) := --- ideal.quotient.field _ - --- instance adjoin_monic.inhabited : inhabited (adjoin_monic k) := ⟨37⟩ - --- /-- The canonical ring homomorphism to `adjoin_monic k`. -/ --- def to_adjoin_monic : k →+* adjoin_monic k := --- (ideal.quotient.mk _).comp C - --- instance adjoin_monic.algebra : algebra k (adjoin_monic k) := --- (to_adjoin_monic k).to_algebra - --- theorem adjoin_monic.algebra_map : algebra_map k (adjoin_monic k) = (ideal.quotient.mk _).comp C := --- rfl - --- theorem adjoin_monic.is_integral (z : adjoin_monic k) : is_integral k z := --- let ⟨p, hp⟩ := ideal.quotient.mk_surjective z in hp ▸ --- mv_polynomial.induction_on p (λ x, is_integral_algebra_map) (λ p q, is_integral_add) --- (λ p f ih, @is_integral_mul _ _ _ _ _ _ (ideal.quotient.mk _ _) ih ⟨f, f.2.1, --- by { erw [adjoin_monic.algebra_map, ← hom_eval₂, --- ideal.quotient.eq_zero_iff_mem], --- exact le_max_ideal k (ideal.subset_span ⟨f, rfl⟩) }⟩) - --- theorem adjoin_monic.exists_root {f : polynomial k} (hfm : f.monic) (hfi : irreducible f) : --- ∃ x : adjoin_monic k, f.eval₂ (to_adjoin_monic k) x = 0 := --- ⟨ideal.quotient.mk _ $ X (⟨f, hfm, hfi⟩ : monic_irreducible k), --- by { rw [to_adjoin_monic, ← hom_eval₂, ideal.quotient.eq_zero_iff_mem], --- exact le_max_ideal k (ideal.subset_span $ ⟨_, rfl⟩) }⟩ - --- /-- The `n`th step of constructing `algebraic_closure`, together with its `field` instance. -/ --- def step_aux (n : ℕ) : Σ α : Type u, field α := --- nat.rec_on n ⟨k, infer_instance⟩ $ λ n ih, ⟨@adjoin_monic ih.1 ih.2, @adjoin_monic.field ih.1 ih.2⟩ - --- /-- The `n`th step of constructing `algebraic_closure`. -/ --- def step (n : ℕ) : Type u := --- (step_aux k n).1 - --- instance step.field (n : ℕ) : field (step k n) := --- (step_aux k n).2 - --- instance step.inhabited (n) : inhabited (step k n) := ⟨37⟩ - --- /-- The canonical inclusion to the `0`th step. -/ --- def to_step_zero : k →+* step k 0 := --- ring_hom.id k - --- /-- The canonical ring homomorphism to the next step. -/ --- def to_step_succ (n : ℕ) : step k n →+* step k (n + 1) := --- @to_adjoin_monic (step k n) (step.field k n) - --- instance step.algebra_succ (n) : algebra (step k n) (step k (n + 1)) := --- (to_step_succ k n).to_algebra - --- theorem to_step_succ.exists_root {n} {f : polynomial (step k n)} --- (hfm : f.monic) (hfi : irreducible f) : --- ∃ x : step k (n + 1), f.eval₂ (to_step_succ k n) x = 0 := --- @adjoin_monic.exists_root _ (step.field k n) _ hfm hfi - --- /-- The canonical ring homomorphism to a step with a greater index. -/ --- def to_step_of_le (m n : ℕ) (h : m ≤ n) : step k m →+* step k n := --- { to_fun := nat.le_rec_on h (λ n, to_step_succ k n), --- map_one' := begin --- induction h with n h ih, { exact nat.le_rec_on_self 1 }, --- rw [nat.le_rec_on_succ h, ih, ring_hom.map_one] --- end, --- map_mul' := λ x y, begin --- induction h with n h ih, { simp_rw nat.le_rec_on_self }, --- simp_rw [nat.le_rec_on_succ h, ih, ring_hom.map_mul] --- end, --- map_zero' := begin --- induction h with n h ih, { exact nat.le_rec_on_self 0 }, --- rw [nat.le_rec_on_succ h, ih, ring_hom.map_zero] --- end, --- map_add' := λ x y, begin --- induction h with n h ih, { simp_rw nat.le_rec_on_self }, --- simp_rw [nat.le_rec_on_succ h, ih, ring_hom.map_add] --- end } - --- @[simp] lemma coe_to_step_of_le (m n : ℕ) (h : m ≤ n) : --- (to_step_of_le k m n h : step k m → step k n) = nat.le_rec_on h (λ n, to_step_succ k n) := --- rfl - --- instance step.algebra (n) : algebra k (step k n) := --- (to_step_of_le k 0 n n.zero_le).to_algebra - --- instance step.scalar_tower (n) : is_scalar_tower k (step k n) (step k (n + 1)) := --- is_scalar_tower.of_algebra_map_eq $ λ z, --- @nat.le_rec_on_succ (step k) 0 n n.zero_le (n + 1).zero_le (λ n, to_step_succ k n) z - --- theorem step.is_integral (n) : ∀ z : step k n, is_integral k z := --- nat.rec_on n (λ z, is_integral_algebra_map) $ λ n ih z, --- is_integral_trans ih _ (adjoin_monic.is_integral (step k n) z : _) - --- instance to_step_of_le.directed_system : --- directed_system (step k) (λ i j h, to_step_of_le k i j h) := --- ⟨λ i x h, nat.le_rec_on_self x, λ i₁ i₂ i₃ h₁₂ h₂₃ x, (nat.le_rec_on_trans h₁₂ h₂₃ x).symm⟩ - --- end algebraic_closure - --- /-- The canonical algebraic closure of a field, the direct limit of adding roots to the field for --- each polynomial over the field. -/ --- def algebraic_closure : Type u := --- ring.direct_limit (algebraic_closure.step k) (λ i j h, algebraic_closure.to_step_of_le k i j h) - --- namespace algebraic_closure - --- instance : field (algebraic_closure k) := --- field.direct_limit.field _ _ - --- instance : inhabited (algebraic_closure k) := ⟨37⟩ - --- /-- The canonical ring embedding from the `n`th step to the algebraic closure. -/ --- def of_step (n : ℕ) : step k n →+* algebraic_closure k := --- ring.direct_limit.of _ _ _ - --- instance algebra_of_step (n) : algebra (step k n) (algebraic_closure k) := --- (of_step k n).to_algebra - --- theorem of_step_succ (n : ℕ) : (of_step k (n + 1)).comp (to_step_succ k n) = of_step k n := --- ring_hom.ext $ λ x, show ring.direct_limit.of (step k) (λ i j h, to_step_of_le k i j h) _ _ = _, --- by { convert ring.direct_limit.of_f n.le_succ x, ext x, exact (nat.le_rec_on_succ' x).symm } - --- theorem exists_of_step (z : algebraic_closure k) : ∃ n x, of_step k n x = z := --- ring.direct_limit.exists_of z - --- -- slow --- theorem exists_root {f : polynomial (algebraic_closure k)} --- (hfm : f.monic) (hfi : irreducible f) : --- ∃ x : algebraic_closure k, f.eval x = 0 := --- begin --- have : ∃ n p, polynomial.map (of_step k n) p = f, --- { convert ring.direct_limit.polynomial.exists_of f }, --- unfreezingI { obtain ⟨n, p, rfl⟩ := this }, --- rw monic_map_iff at hfm, --- have := hfm.irreducible_of_irreducible_map (of_step k n) p hfi, --- obtain ⟨x, hx⟩ := to_step_succ.exists_root k hfm this, --- refine ⟨of_step k (n + 1) x, _⟩, --- rw [← of_step_succ k n, eval_map, ← hom_eval₂, hx, ring_hom.map_zero] --- end - --- instance : is_alg_closed (algebraic_closure k) := --- is_alg_closed.of_exists_root _ $ λ f, exists_root k - --- instance {R : Type*} [comm_semiring R] [alg : algebra R k] : --- algebra R (algebraic_closure k) := --- ((of_step k 0).comp (@algebra_map _ _ _ _ alg)).to_algebra - --- lemma algebra_map_def {R : Type*} [comm_semiring R] [alg : algebra R k] : --- algebra_map R (algebraic_closure k) = ((of_step k 0 : k →+* _).comp (@algebra_map _ _ _ _ alg)) := --- rfl - --- instance {R S : Type*} [comm_semiring R] [comm_semiring S] --- [algebra R S] [algebra S k] [algebra R k] [is_scalar_tower R S k] : --- is_scalar_tower R S (algebraic_closure k) := --- is_scalar_tower.of_algebra_map_eq (λ x, --- ring_hom.congr_arg _ (is_scalar_tower.algebra_map_apply R S k x : _)) - --- /-- Canonical algebra embedding from the `n`th step to the algebraic closure. -/ --- def of_step_hom (n) : step k n →ₐ[k] algebraic_closure k := --- { commutes' := λ x, ring.direct_limit.of_f n.zero_le x, --- .. of_step k n } - --- theorem is_algebraic : algebra.is_algebraic k (algebraic_closure k) := --- λ z, (is_algebraic_iff_is_integral _).2 $ let ⟨n, x, hx⟩ := exists_of_step k z in --- hx ▸ is_integral_alg_hom (of_step_hom k n) (step.is_integral k n x) - --- instance : is_alg_closure k (algebraic_closure k) := --- ⟨algebraic_closure.is_alg_closed k, is_algebraic k⟩ - --- end algebraic_closure - --- /-- --- Every element `f` in a nontrivial finite-dimensional algebra `A` --- over an algebraically closed field `K` --- has non-empty spectrum: --- that is, there is some `c : K` so `f - c • 1` is not invertible. --- -/ --- -- We will use this both to show eigenvalues exist, and to prove Schur's lemma. --- lemma exists_spectrum_of_is_alg_closed_of_finite_dimensional (𝕜 : Type*) [field 𝕜] [is_alg_closed 𝕜] --- {A : Type*} [nontrivial A] [ring A] [algebra 𝕜 A] [I : finite_dimensional 𝕜 A] (f : A) : --- ∃ c : 𝕜, ¬ is_unit (f - algebra_map 𝕜 A c) := --- begin --- obtain ⟨p, ⟨h_mon, h_eval_p⟩⟩ := is_integral_of_noetherian I f, --- have nu : ¬ is_unit (aeval f p), { rw [←aeval_def] at h_eval_p, rw h_eval_p, simp, }, --- rw [eq_prod_roots_of_monic_of_splits_id h_mon (is_alg_closed.splits p), --- ←multiset.prod_to_list, alg_hom.map_list_prod] at nu, --- replace nu := mt list.prod_is_unit nu, --- simp only [not_forall, exists_prop, aeval_C, multiset.mem_to_list, --- list.mem_map, aeval_X, exists_exists_and_eq_and, multiset.mem_map, alg_hom.map_sub] at nu, --- exact ⟨nu.some, nu.some_spec.2⟩, --- end From af35a255e508da77f725c26461caa5ee180ed98f Mon Sep 17 00:00:00 2001 From: Chris Hughes Date: Sun, 5 Sep 2021 22:11:47 +0100 Subject: [PATCH 24/51] feat(data/set/Union_lift): lift functions to Unions of sets --- src/algebra/algebra/subalgebra.lean | 67 ++++++++ src/data/set/Union_lift.lean | 235 ++++++++++++++++++++++++++++ 2 files changed, 302 insertions(+) create mode 100644 src/data/set/Union_lift.lean diff --git a/src/algebra/algebra/subalgebra.lean b/src/algebra/algebra/subalgebra.lean index 68ef6408a8d70..afd01b7db36e9 100644 --- a/src/algebra/algebra/subalgebra.lean +++ b/src/algebra/algebra/subalgebra.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Kenny Lau, Yury Kudryashov -/ import algebra.algebra.operations +import data.set.Union_lift /-! # Subalgebras over Commutative Semiring @@ -648,6 +649,21 @@ instance : unique (subalgebra R R) := end .. algebra.subalgebra.inhabited } +def inclusion {S T : subalgebra R A} (h : S ≤ T) : S →ₐ[R] T := +{ to_fun := set.inclusion h, + map_one' := rfl, + map_add' := λ _ _, rfl, + map_mul' := λ _ _, rfl, + map_zero' := rfl, + commutes' := λ _, rfl } + +lemma inclusion_injective {S T : subalgebra R A} (h : S ≤ T) : + function.injective (inclusion h) := +λ _ _, subtype.ext ∘ subtype.mk.inj + +@[simp] lemma coe_inclusion {S T : subalgebra R A} (h : S ≤ T) (s : S) : + (inclusion h s : A) = s := rfl + /-- Two subalgebras that are equal are also equivalent as algebras. This is the `subalgebra` version of `linear_equiv.of_eq` and `equiv.set.of_eq`. -/ @@ -702,6 +718,57 @@ set_like.coe_injective set.prod_inter_prod end prod +section supr_lift +variables {ι : Type*} + +lemma coe_supr_of_directed [nonempty ι] {S : ι → subalgebra R A} + (dir : directed (≤) S) : + ((supr S : subalgebra R A) : set A) = ⋃ i, (S i : set A) := +let K : subalgebra R A := + { carrier := ⋃ i, (S i), + mul_mem' := λ x y hx hy, + let ⟨i, hi⟩ := set.mem_Union.1 hx in + let ⟨j, hj⟩ := set.mem_Union.1 hy in + let ⟨k, hik, hjk⟩ := dir i j in + set.mem_Union.2 ⟨k, subalgebra.mul_mem (S k) (hik hi) (hjk hj)⟩ , + add_mem' := λ x y hx hy, + let ⟨i, hi⟩ := set.mem_Union.1 hx in + let ⟨j, hj⟩ := set.mem_Union.1 hy in + let ⟨k, hik, hjk⟩ := dir i j in + set.mem_Union.2 ⟨k, subalgebra.add_mem (S k) (hik hi) (hjk hj)⟩, + algebra_map_mem' := λ r, let i := @nonempty.some ι infer_instance in + set.mem_Union.2 ⟨i, subalgebra.algebra_map_mem _ _⟩ } in +have supr S = K, + from le_antisymm (supr_le (λ i, set.subset_Union (λ i, ↑(S i)) i)) + (set_like.coe_subset_coe.1 + (set.Union_subset (λ i, set_like.coe_subset_coe.2 (le_supr _ _)))), +this.symm ▸ rfl + +/-- Define an algebra homomorphism on a directed supremum of subalgebras by defining +it on each subalgebra, and proving that it agrees on the intersection of subalgebras. -/ +noncomputable def supr_lift [nonempty ι] + (S : ι → subalgebra R A) + (dir : directed (≤) S) + (f : Π i, S i →ₐ[R] B) + (hf : ∀ (i j : ι) (h : S i ≤ S j), f i = (f j).comp (inclusion h)) : + (supr S : subalgebra R A) →ₐ[R] B := +{ to_fun := set.lift_of_eq_Union (λ i, ↑(S i)) (λ i x, f i x) + (λ i j x hxi hxj, + let ⟨k, hik, hjk⟩ := dir i j in + begin + rw [hf i k hik, hf j k hjk], + refl + end) _ + (by exactI coe_supr_of_directed dir), + map_one' := set.lift_of_eq_Union_const _ (λ _, 1) (λ _, rfl) _ (by simp), + map_zero' := set.lift_of_eq_Union_const _ (λ _, 0) (λ _, rfl) _ (by simp), + map_mul' := set.lift_of_eq_Union_binary dir _ (λ _, (*)) (λ _ _ _, rfl) _ (by simp), + map_add' := set.lift_of_eq_Union_binary dir _ (λ _, (+)) (λ _ _ _, rfl) _ (by simp), + commutes' := λ r, set.lift_of_eq_Union_const _ (λ _, algebra_map _ _ r) + (λ _, rfl) _ (λ i, by erw [alg_hom.commutes (f i)]) } + +end supr_lift + end subalgebra section nat diff --git a/src/data/set/Union_lift.lean b/src/data/set/Union_lift.lean new file mode 100644 index 0000000000000..d8b5ce30a4f98 --- /dev/null +++ b/src/data/set/Union_lift.lean @@ -0,0 +1,235 @@ +/- +Copyright (c) 2021 Chris Hughes. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Chris Hughes +-/ +import data.set.lattice +import order.directed +/-! +# Union lift +This file defines `Union_lift` to lift function defined on each of a collection of sets +to the Union of those sets. + +## Main definitions + +* `Union_lift` - Given a Union of sets `Union S`, define a function on the Union by defining +it on each component, and proving that it agrees on the intersections. +* `lift_of_eq_Union` Version of `Union_lift` for a set that is propositionally equal to `Union S`. + This is mainly motivated by the need to define functions on directed supremums of e.g. subgroups, + where the supremum is propositionally, but not definitionally equal to the Union. + +## Main statements + +There are proofs of the obvious properties of `Union_lift`, i.e. what it does to elements of +each of the sets in the `Union`, stated in different ways. + +There are also three lemmas about `Union_lift` and three lemmas about `lift_of_eq_Union` intended +to aid with proving that `Union_lift` of `lift_of_eq_Union` is a homomorphism when defined +on a Union of substructures. There is one lemma each to show that constants, unary functions, +or binary functions are preserved. These lemmas are: + +*`Union_lift_const` +*`Union_lift_unary` +*`Union_lift_binary` +*`lift_of_eq_Union_const` +*`lift_of_eq_Union_unary` +*`lift_of_eq_Union_binary` + +## Tags + +directed union, directed supremum +-/ + +variables {α ι : Type*} {β : Type*} + +namespace set + +/-- Given a Union of sets `Union S`, define a function on the Union by defining +it on each component, and proving that it agrees on the intersections. -/ +noncomputable def Union_lift (S : ι → set α) + (f : Π i (x : S i), β) + (hf : ∀ i j (x : α) (hxi : x ∈ S i) (hxj : x ∈ S j), f i ⟨x, hxi⟩ = f j ⟨x, hxj⟩) + (x : Union S) : β := +let i : ι := classical.some (mem_Union.1 x.2) in +have hi : (x : α) ∈ S i := classical.some_spec (mem_Union.1 x.2), +f i ⟨x, hi⟩ + +@[simp] lemma Union_lift_inclusion {S : ι → set α} + {f : Π i (x : S i), β} + {hf : ∀ i j (x : α) (hxi : x ∈ S i) (hxj : x ∈ S j), f i ⟨x, hxi⟩ = f j ⟨x, hxj⟩ } + {i : ι} (x : S i) (h : S i ⊆ Union S := set.subset_Union S i) : + Union_lift S f hf (set.inclusion h x) = f i x := +let j : ι := classical.some (mem_Union.1 (h x.2)) in +have hj : (x : α) ∈ S j := classical.some_spec (mem_Union.1 (h x.2)), +by cases x with x hx; exact hf j i x hj hx + +@[simp] lemma Union_lift_mk + {S : ι → set α} + {f : Π i (x : S i), β} + {hf : ∀ i j (x : α) (hxi : x ∈ S i) (hxj : x ∈ S j), f i ⟨x, hxi⟩ = f j ⟨x, hxj⟩ } + {i : ι} (x : S i) (hx : (x : α) ∈ Union S := set.subset_Union S i x.2) : + Union_lift S f hf ⟨x, hx⟩ = f i x := +Union_lift_inclusion x + +lemma Union_lift_of_mem + {S : ι → set α} + {f : Π i (x : S i), β} + {hf : ∀ i j (x : α) (hxi : x ∈ S i) (hxj : x ∈ S j), f i ⟨x, hxi⟩ = f j ⟨x, hxj⟩ } + (x : Union S) {i : ι} (hx : (x : α) ∈ S i) : + Union_lift S f hf x = f i ⟨x, hx⟩ := +by cases x with x hx; exact hf _ _ _ _ _ + +/-- `Union_lift_const` is useful for proving that `Union_lift` is a homomorphism + of algebraic structures when defined on the Union of algebraic subobjects. + For example, it could be used to prove that the lift of a collection + of group homomorphisms on a union of subgroups preserves `1`. See also + `lift_of_eq_Union_const` -/ +lemma Union_lift_const + {S : ι → set α} + {f : Π i (x : S i), β} + {hf : ∀ i j (x : α) (hxi : x ∈ S i) (hxj : x ∈ S j), f i ⟨x, hxi⟩ = f j ⟨x, hxj⟩} + (c : Union S) + (ci : Π i, S i) + (hci : ∀ i, (ci i : α) = c) + (cβ : β) + (h : ∀ i, f i (ci i) = cβ) : + Union_lift S f hf c = cβ := +let ⟨i, hi⟩ := set.mem_Union.1 c.prop in +have (ci i) = ⟨c, hi⟩, from subtype.ext (hci i), +by rw [Union_lift_of_mem _ hi, ← this, h] + +/-- `Union_lift_unary` is useful for proving that `Union_lift` is a homomorphism + of algebraic structures when defined on the Union of algebraic subobjects. + For example, it could be used to prove that the lift of a collection + of linear_maps on a union of submodules preserves scalar multiplication. See also + `lift_of_eq_Union_unary` -/ +lemma Union_lift_unary + {S : ι → set α} + {f : Π i (x : S i), β} + {hf : ∀ i j (x : α) (hxi : x ∈ S i) (hxj : x ∈ S j), f i ⟨x, hxi⟩ = f j ⟨x, hxj⟩} + (u : (Union S) → (Union S)) + (ui : Π i, S i → S i) + (hui : ∀ i x, set.inclusion (set.subset_Union S i) (ui i x) = + u (set.inclusion (set.subset_Union S i) x)) + (uβ : β → β) + (h : ∀ i (x : S i), (f i (ui i x)) = uβ (f i x)) + (x : Union S) : + Union_lift S f hf (u x) = uβ (Union_lift S f hf x) := +begin + cases set.mem_Union.1 x.prop with i hi, + rw [Union_lift_of_mem x hi, ← h i], + have : x = (set.inclusion (set.subset_Union S i) ⟨x, hi⟩), { cases x, refl }, + have hx' : (set.inclusion (set.subset_Union S i) (ui i ⟨x, hi⟩) : α) ∈ S i, + from (ui i ⟨x, hi⟩).prop, + conv_lhs { rw [this, ← hui, Union_lift_of_mem _ hx'] }, + simp only [coe_inclusion, subtype.coe_eta] +end + +/-- `Union_lift_binary` is useful for proving that `Union_lift` is a homomorphism + of algebraic structures when defined on the Union of algebraic subobjects. + For example, it could be used to prove that the lift of a collection + of group homomorphisms on a union of subgroups preserves `*`. See also + `lift_of_eq_Union_binary` -/ +lemma Union_lift_binary + {S : ι → set α} + {f : Π i (x : S i), β} + {hf : ∀ i j (x : α) (hxi : x ∈ S i) (hxj : x ∈ S j), f i ⟨x, hxi⟩ = f j ⟨x, hxj⟩} + (dir: directed (≤) S) + (op : (Union S) → (Union S) → (Union S)) + (opi : Π i, S i → S i → S i) + (hopi : ∀ i x y, set.inclusion (set.subset_Union S i) (opi i x y) = + op (set.inclusion (set.subset_Union S i) x) + (set.inclusion (set.subset_Union S i) y)) + (opβ : β → β → β) + (h : ∀ i (x y : S i), (f i (opi i x y)) = opβ (f i x) (f i y)) + (x y : Union S) : + Union_lift S f hf (op x y) = opβ (Union_lift S f hf x) (Union_lift S f hf y) := +begin + cases set.mem_Union.1 x.prop with i hi, + cases set.mem_Union.1 y.prop with j hj, + rcases dir i j with ⟨k, hik, hjk⟩, + rw [Union_lift_of_mem x (hik hi), Union_lift_of_mem y (hjk hj), ← h k], + have hx : x = (set.inclusion (set.subset_Union S k) ⟨x, hik hi⟩), { cases x, refl }, + have hy : y = (set.inclusion (set.subset_Union S k) ⟨y, hjk hj⟩), { cases y, refl }, + have hxy : (set.inclusion (set.subset_Union S k) (opi k ⟨x, hik hi⟩ ⟨y, hjk hj⟩) : α) ∈ S k, + from (opi k ⟨x, hik hi⟩ ⟨y, hjk hj⟩).prop, + conv_lhs { rw [hx, hy, ← hopi, Union_lift_of_mem _ hxy] }, + simp only [coe_inclusion, subtype.coe_eta] +end + +/-- Version of `Union_lift` for sets that are propositionally equal to the Union. + Define a map on a set equal to the `Union` of sets by defining it on each of the sets + in the `Union`. -/ +noncomputable def lift_of_eq_Union + (S : ι → set α) + (f : Π i (x : S i), β) + (hf : ∀ i j (x : α) (hxi : x ∈ S i) (hxj : x ∈ S j), f i ⟨x, hxi⟩ = f j ⟨x, hxj⟩) + (T : set α) (hT : T = ⋃ i, S i) + (x : T) : β := +by subst hT; exact Union_lift S f hf x + +/-- Version of `Union_lift_const` for `lift_of_eq_Union`. -/ +lemma lift_of_eq_Union_const + {S : ι → set α} + {f : Π i (x : S i), β} + {hf : ∀ i j (x : α) (hxi : x ∈ S i) (hxj : x ∈ S j), f i ⟨x, hxi⟩ = f j ⟨x, hxj⟩} + {T : set α} {hT : T = Union S} + (c : T) + (ci : Π i, S i) + (hci : ∀ i, (ci i : α) = c) + (cβ : β) + (h : ∀ i, f i (ci i) = cβ) : + lift_of_eq_Union S f hf T hT c = cβ := +begin + subst T, + delta lift_of_eq_Union, + exact Union_lift_const c ci hci cβ h +end + +/-- Version of `Union_lift_unary` for `lift_of_eq_Union`. -/ +lemma lift_of_eq_Union_unary + {S : ι → set α} + {f : Π i (x : S i), β} + {hf : ∀ i j (x : α) (hxi : x ∈ S i) (hxj : x ∈ S j), f i ⟨x, hxi⟩ = f j ⟨x, hxj⟩} + {T : set α} {hT : T = Union S} + (u : T → T) + (ui : Π i, S i → S i) + (hui : ∀ i x, set.inclusion (show S i ⊆ T, from hT.symm ▸ set.subset_Union S i) (ui i x) = + u (set.inclusion (show S i ⊆ T, from hT.symm ▸ set.subset_Union S i) x)) + (uβ : β → β) + (h : ∀ i (x : S i), (f i (ui i x)) = uβ (f i x)) + (x : T) : + lift_of_eq_Union S f hf T hT (u x) = uβ (lift_of_eq_Union S f hf T hT x) := +begin + subst T, + delta lift_of_eq_Union, + exact Union_lift_unary u ui hui uβ h x +end + +/-- Version of `Union_lift_binary` for `lift_of_eq_Union`. -/ +lemma lift_of_eq_Union_binary + {S : ι → set α} + {f : Π i (x : S i), β} + {hf : ∀ i j (x : α) (hxi : x ∈ S i) (hxj : x ∈ S j), f i ⟨x, hxi⟩ = f j ⟨x, hxj⟩} + {T : set α} {hT : T = ⋃ i, S i} + (dir: directed (≤) S) + (op : T → T → T) + (opi : Π i, S i → S i → S i) + (hopi : ∀ i x y, set.inclusion (show S i ⊆ T, from hT.symm ▸ set.subset_Union S i) (opi i x y) = + op (set.inclusion (show S i ⊆ T, from hT.symm ▸ set.subset_Union S i) x) + (set.inclusion (show S i ⊆ T, from hT.symm ▸ set.subset_Union S i) y)) + (opβ : β → β → β) + (h : ∀ i (x y : S i), (f i (opi i x y)) = opβ (f i x) (f i y)) + (x y : T) : + lift_of_eq_Union S f hf T hT (op x y) = + opβ (lift_of_eq_Union S f hf T hT x) + (lift_of_eq_Union S f hf T hT y) := +begin + subst T, + delta lift_of_eq_Union, + exact Union_lift_binary dir op opi hopi opβ h x y +end + +attribute [irreducible] lift_of_eq_Union Union_lift + +end set From 10697b57a8a5bb74d7ed23536bd93628281c4a02 Mon Sep 17 00:00:00 2001 From: Chris Hughes Date: Sun, 5 Sep 2021 22:33:13 +0100 Subject: [PATCH 25/51] add docstrings and more lemmas --- src/algebra/algebra/subalgebra.lean | 14 +++++++++++++- 1 file changed, 13 insertions(+), 1 deletion(-) diff --git a/src/algebra/algebra/subalgebra.lean b/src/algebra/algebra/subalgebra.lean index fea2b8f9bc1b1..bbd060f2809a0 100644 --- a/src/algebra/algebra/subalgebra.lean +++ b/src/algebra/algebra/subalgebra.lean @@ -648,7 +648,9 @@ instance : unique (subalgebra R R) := end .. algebra.subalgebra.inhabited } -/-- The map `S → T` when `S` is a subalgebra contained in the subalgebra `T`. -/ +/-- The map `S → T` when `S` is a subalgebra contained in the subalgebra `T`. + +This is the subalgebra version of `submodule.of_le`, or `subring.inclusion` -/ def inclusion {S T : subalgebra R A} (h : S ≤ T) : S →ₐ[R] T := { to_fun := set.inclusion h, map_one' := rfl, @@ -661,6 +663,16 @@ lemma inclusion_injective {S T : subalgebra R A} (h : S ≤ T) : function.injective (inclusion h) := λ _ _, subtype.ext ∘ subtype.mk.inj +@[simp] lemma inclusion_self {S : subalgebra R A} (x : S) : + inclusion (le_refl S) x = x := subtype.ext rfl + +@[simp] lemma inclusion_right {S T : subalgebra R A} (h : S ≤ T) (x : T) + (m : (x : A) ∈ S) : inclusion h ⟨x, m⟩ = x := subtype.ext rfl + +@[simp] lemma inclusion_inclusion {S T U : subalgebra R A} (hst : S ≤ T) (htu : T ≤ U) + (x : S) : inclusion htu (inclusion hst x) = inclusion (set.subset.trans hst htu) x := +subtype.ext rfl + @[simp] lemma coe_inclusion {S T : subalgebra R A} (h : S ≤ T) (s : S) : (inclusion h s : A) = s := rfl From 31023b83f69b97b0eb28550436a64cb2b9480001 Mon Sep 17 00:00:00 2001 From: Chris Hughes Date: Mon, 6 Sep 2021 00:38:54 +0100 Subject: [PATCH 26/51] slightly change inclusion_self statement --- src/algebra/algebra/subalgebra.lean | 5 +++-- 1 file changed, 3 insertions(+), 2 deletions(-) diff --git a/src/algebra/algebra/subalgebra.lean b/src/algebra/algebra/subalgebra.lean index bbd060f2809a0..81e61859ad66a 100644 --- a/src/algebra/algebra/subalgebra.lean +++ b/src/algebra/algebra/subalgebra.lean @@ -663,8 +663,9 @@ lemma inclusion_injective {S T : subalgebra R A} (h : S ≤ T) : function.injective (inclusion h) := λ _ _, subtype.ext ∘ subtype.mk.inj -@[simp] lemma inclusion_self {S : subalgebra R A} (x : S) : - inclusion (le_refl S) x = x := subtype.ext rfl +@[simp] lemma inclusion_self {S : subalgebra R A}: + inclusion (le_refl S) = alg_hom.id R S := +alg_hom.ext $ λ x, subtype.ext rfl @[simp] lemma inclusion_right {S T : subalgebra R A} (h : S ≤ T) (x : T) (m : (x : A) ∈ S) : inclusion h ⟨x, m⟩ = x := subtype.ext rfl From c0a2875977fd660bb5da0cda25a86977566ea97a Mon Sep 17 00:00:00 2001 From: Chris Hughes <33847686+ChrisHughes24@users.noreply.github.com> Date: Mon, 6 Sep 2021 00:43:49 +0100 Subject: [PATCH 27/51] change statement of inclusion_self --- src/algebra/algebra/subalgebra.lean | 5 +++-- 1 file changed, 3 insertions(+), 2 deletions(-) diff --git a/src/algebra/algebra/subalgebra.lean b/src/algebra/algebra/subalgebra.lean index bbd060f2809a0..81e61859ad66a 100644 --- a/src/algebra/algebra/subalgebra.lean +++ b/src/algebra/algebra/subalgebra.lean @@ -663,8 +663,9 @@ lemma inclusion_injective {S T : subalgebra R A} (h : S ≤ T) : function.injective (inclusion h) := λ _ _, subtype.ext ∘ subtype.mk.inj -@[simp] lemma inclusion_self {S : subalgebra R A} (x : S) : - inclusion (le_refl S) x = x := subtype.ext rfl +@[simp] lemma inclusion_self {S : subalgebra R A}: + inclusion (le_refl S) = alg_hom.id R S := +alg_hom.ext $ λ x, subtype.ext rfl @[simp] lemma inclusion_right {S T : subalgebra R A} (h : S ≤ T) (x : T) (m : (x : A) ∈ S) : inclusion h ⟨x, m⟩ = x := subtype.ext rfl From 912a1c74c2f99c3a91b129d92f7fc308618a4bc3 Mon Sep 17 00:00:00 2001 From: Chris Hughes Date: Mon, 6 Sep 2021 11:56:51 +0100 Subject: [PATCH 28/51] work on algebraic closure --- src/field_theory/algebraic_closure.lean | 67 +++++++++++++++++-------- 1 file changed, 46 insertions(+), 21 deletions(-) diff --git a/src/field_theory/algebraic_closure.lean b/src/field_theory/algebraic_closure.lean index d90eba46b75ca..9d9f4e83e1de6 100644 --- a/src/field_theory/algebraic_closure.lean +++ b/src/field_theory/algebraic_closure.lean @@ -395,39 +395,64 @@ namespace subfield_with_hom variables {E₁ E₂ E₃ : subfield_with_hom K L M hL} instance : has_le (subfield_with_hom K L M hL) := -{ le := λ E₁ E₂, ∃ h : E₁.carrier ≤ E₂.carrier, E₂.emb ∘ inclusion h = E₁.emb } +{ le := λ E₁ E₂, ∃ h : E₁.carrier ≤ E₂.carrier, ∀ x, E₂.emb (inclusion h x) = E₁.emb x } -lemma le_def : E₁ ≤ E₂ ↔ ∃ h : E₁.carrier ≤ E₂.carrier, E₂.emb ∘ inclusion h = E₁.emb := iff.rfl +lemma le_def : E₁ ≤ E₂ ↔ ∃ h : E₁.carrier ≤ E₂.carrier, ∀ x, E₂.emb (inclusion h x) = E₁.emb x := +iff.rfl -lemma subalgebra_le_of_le (h : E₁ ≤ E₂) : E₁.carrier ≤ E₂.carrier := -by { rw le_def at h, cases h, assumption } - -lemma compat (h : E₁ ≤ E₂) : E₂.emb ∘ inclusion (subalgebra_le_of_le h) = E₁.emb := +lemma compat (h : E₁ ≤ E₂) : ∀ x, E₂.emb (inclusion h.fst x) = E₁.emb x := by { rw le_def at h, cases h, assumption } instance : preorder (subfield_with_hom K L M hL) := -{ le := λ E₁ E₂, ∃ h : E₁.carrier ≤ E₂.carrier, E₂.emb ∘ inclusion h = E₁.emb, - le_refl := λ E, ⟨le_refl _, by rw [inclusion_refl, comp.right_id]⟩, - le_trans := λ E₁ E₂ E₃ h₁₂ h₂₃, ⟨le_trans (subalgebra_le_of_le h₁₂) (subalgebra_le_of_le h₂₃), - begin - erw inclusion_trans (subalgebra_le_of_le h₁₂) (subalgebra_le_of_le h₂₃), - rw [← comp.assoc, compat, compat] - end⟩ } +{ le := (≤), + le_refl := λ E, ⟨le_refl _, by simp⟩, + le_trans := λ E₁ E₂ E₃ h₁₂ h₂₃, + ⟨le_trans h₁₂.fst h₂₃.fst, + λ _, by erw [← inclusion_inclusion h₁₂.fst h₂₃.fst, compat, compat]⟩, + -- le_antisymm := λ a b hab hba, begin + -- cases a with a aemb, + -- cases b with b bemb, + -- have h : a = b, from le_antisymm hab.fst hba.fst, + -- subst b, + -- simp only [true_and, eq_self_iff_true, heq_iff_eq], + -- ext x, + -- refine eq.trans (hab.snd x).symm _, + -- simp + -- end + } end subfield_with_hom open lattice -def maximal_subfield_with_hom_chain (c : set (subfield_with_hom K L M hL)) (hc : chain (≤) c) : +def maximal_subfield_with_hom_chain (c : set (subfield_with_hom K L M hL)) + [nonempty c] + (hc : chain (≤) c) : ∃ ub : subfield_with_hom K L M hL, ∀ N, N ∈ c → N ≤ ub := let ub : subfield_with_hom K L M hL := -{ carrier := Sup (subfield_with_hom.carrier '' c), - emb := } in -⟨ub, λ N hN, -begin - refine ⟨le_Sup ⟨N, hN, rfl⟩, _⟩, - sorry -end⟩ +{ carrier := ⨆ i : c, (i : subfield_with_hom K L M hL).carrier, + emb := subalgebra.supr_lift + (λ i : c, (i : subfield_with_hom K L M hL).carrier) + (λ i j, let ⟨k, hik, hjk⟩ := directed_on_iff_directed.1 hc.directed_on i j in + ⟨k, hik.fst, hjk.fst⟩) + (λ i, (i : subfield_with_hom K L M hL).emb) + begin + assume i j h, + ext x, + cases hc.total i.prop j.prop with hij hji, + { simp [← hij.snd x] }, + { erw [alg_hom.comp_apply, ← hji.snd (inclusion h x), + inclusion_inclusion, inclusion_self, alg_hom.id_apply x] } + end } in +⟨ub, λ N hN, ⟨(le_supr (λ i : c, (i : subfield_with_hom K L M hL).carrier) ⟨N, hN⟩ : _), + begin + intro x, + simp [ub], + + end⟩⟩ + -- ⟨le_supr (λ i : c, (i : subfield_with_hom K L M hL).carrier) ⟨N, hN⟩, begin + + --end⟩⟩ variables (hL M) From ef381806323a3ddfbc0b4cec3b1b4877969d724f Mon Sep 17 00:00:00 2001 From: Chris Hughes Date: Mon, 6 Sep 2021 12:31:15 +0100 Subject: [PATCH 29/51] Add lemmas and change docstrings --- src/algebra/algebra/subalgebra.lean | 35 ++++++++++++--- src/data/set/Union_lift.lean | 67 ++++++++++++++++++++++------- 2 files changed, 81 insertions(+), 21 deletions(-) diff --git a/src/algebra/algebra/subalgebra.lean b/src/algebra/algebra/subalgebra.lean index 298fac72ceacf..dc363371c0483 100644 --- a/src/algebra/algebra/subalgebra.lean +++ b/src/algebra/algebra/subalgebra.lean @@ -760,12 +760,12 @@ this.symm ▸ rfl /-- Define an algebra homomorphism on a directed supremum of subalgebras by defining it on each subalgebra, and proving that it agrees on the intersection of subalgebras. -/ noncomputable def supr_lift [nonempty ι] - (S : ι → subalgebra R A) - (dir : directed (≤) S) - (f : Π i, S i →ₐ[R] B) - (hf : ∀ (i j : ι) (h : S i ≤ S j), f i = (f j).comp (inclusion h)) : - (supr S : subalgebra R A) →ₐ[R] B := -{ to_fun := set.lift_of_eq_Union (λ i, ↑(S i)) (λ i x, f i x) + (K : ι → subalgebra R A) + (dir : directed (≤) K) + (f : Π i, K i →ₐ[R] B) + (hf : ∀ (i j : ι) (h : K i ≤ K j), f i = (f j).comp (inclusion h)) : + (supr K : subalgebra R A) →ₐ[R] B := +{ to_fun := set.lift_of_eq_Union (λ i, ↑(K i)) (λ i x, f i x) (λ i j x hxi hxj, let ⟨k, hik, hjk⟩ := dir i j in begin @@ -780,6 +780,29 @@ noncomputable def supr_lift [nonempty ι] commutes' := λ r, set.lift_of_eq_Union_const _ (λ _, algebra_map _ _ r) (λ _, rfl) _ (λ i, by erw [alg_hom.commutes (f i)]) } +variables [nonempty ι] {K : ι → subalgebra R A} {dir : directed (≤) K} + {f : Π i, K i →ₐ[R] B} + {hf : ∀ (i j : ι) (h : K i ≤ K j), f i = (f j).comp (inclusion h)} + +@[simp] lemma supr_lift_inclusion {i : ι} (x : K i) + (h : K i ≤ supr K := le_supr K i) : + supr_lift K dir f hf (inclusion h x) = f i x := +set.lift_of_eq_Union_inclusion _ + +@[simp] lemma supr_lift_comp_inclusion {i : ι} + (h : K i ≤ supr K := le_supr K i) : + (supr_lift K dir f hf).comp (inclusion h) = f i := +by ext; simp + +@[simp] lemma supr_lift_mk {i : ι} (x : K i) + (hx : (x : A) ∈ supr K := set_like.le_def.2 (le_supr K i) x.prop) : + supr_lift K dir f hf ⟨x, hx⟩ = f i x := +set.lift_of_eq_Union_mk x hx + +lemma supr_lift_of_mem {i : ι} (x : supr K) (hx : (x : A) ∈ K i) : + supr_lift K dir f hf x = f i ⟨x, hx⟩ := +set.lift_of_eq_Union_of_mem x hx + end supr_lift end subalgebra diff --git a/src/data/set/Union_lift.lean b/src/data/set/Union_lift.lean index d8b5ce30a4f98..90ac5f5cdd284 100644 --- a/src/data/set/Union_lift.lean +++ b/src/data/set/Union_lift.lean @@ -7,14 +7,14 @@ import data.set.lattice import order.directed /-! # Union lift -This file defines `Union_lift` to lift function defined on each of a collection of sets +This file defines `set.Union_lift` to lift function defined on each of a collection of sets to the Union of those sets. ## Main definitions -* `Union_lift` - Given a Union of sets `Union S`, define a function on the Union by defining +* `set.Union_lift` - Given a Union of sets `Union S`, define a function on the Union by defining it on each component, and proving that it agrees on the intersections. -* `lift_of_eq_Union` Version of `Union_lift` for a set that is propositionally equal to `Union S`. +* `set.lift_of_eq_Union` Version of `Union_lift` for a set that is propositionally equal to `Union S`. This is mainly motivated by the need to define functions on directed supremums of e.g. subgroups, where the supremum is propositionally, but not definitionally equal to the Union. @@ -28,12 +28,12 @@ to aid with proving that `Union_lift` of `lift_of_eq_Union` is a homomorphism wh on a Union of substructures. There is one lemma each to show that constants, unary functions, or binary functions are preserved. These lemmas are: -*`Union_lift_const` -*`Union_lift_unary` -*`Union_lift_binary` -*`lift_of_eq_Union_const` -*`lift_of_eq_Union_unary` -*`lift_of_eq_Union_binary` +*`set.Union_lift_const` +*`set.Union_lift_unary` +*`set.Union_lift_binary` +*`set.lift_of_eq_Union_const` +*`set.lift_of_eq_Union_unary` +*`set.lift_of_eq_Union_binary` ## Tags @@ -50,18 +50,16 @@ noncomputable def Union_lift (S : ι → set α) (f : Π i (x : S i), β) (hf : ∀ i j (x : α) (hxi : x ∈ S i) (hxj : x ∈ S j), f i ⟨x, hxi⟩ = f j ⟨x, hxj⟩) (x : Union S) : β := -let i : ι := classical.some (mem_Union.1 x.2) in -have hi : (x : α) ∈ S i := classical.some_spec (mem_Union.1 x.2), -f i ⟨x, hi⟩ +let i := classical.indefinite_description _ (mem_Union.1 x.2) in +f i ⟨x, i.prop⟩ @[simp] lemma Union_lift_inclusion {S : ι → set α} {f : Π i (x : S i), β} {hf : ∀ i j (x : α) (hxi : x ∈ S i) (hxj : x ∈ S j), f i ⟨x, hxi⟩ = f j ⟨x, hxj⟩ } {i : ι} (x : S i) (h : S i ⊆ Union S := set.subset_Union S i) : Union_lift S f hf (set.inclusion h x) = f i x := -let j : ι := classical.some (mem_Union.1 (h x.2)) in -have hj : (x : α) ∈ S j := classical.some_spec (mem_Union.1 (h x.2)), -by cases x with x hx; exact hf j i x hj hx +let j := classical.indefinite_description _ (mem_Union.1 (h x.2)) in +by cases x with x hx; exact hf j i x j.2 hx @[simp] lemma Union_lift_mk {S : ι → set α} @@ -168,6 +166,45 @@ noncomputable def lift_of_eq_Union (x : T) : β := by subst hT; exact Union_lift S f hf x +@[simp] lemma lift_of_eq_Union_inclusion {S : ι → set α} + {f : Π i (x : S i), β} + {hf : ∀ i j (x : α) (hxi : x ∈ S i) (hxj : x ∈ S j), f i ⟨x, hxi⟩ = f j ⟨x, hxj⟩ } + {T : set α} {hT : T = Union S} + {i : ι} (x : S i) (h : S i ⊆ T := hT.symm ▸ set.subset_Union S i) : + lift_of_eq_Union S f hf T hT (set.inclusion h x) = f i x := +begin + subst hT, + delta lift_of_eq_Union, + exact Union_lift_inclusion x h +end + +@[simp] lemma lift_of_eq_Union_mk + {S : ι → set α} + {f : Π i (x : S i), β} + {hf : ∀ i j (x : α) (hxi : x ∈ S i) (hxj : x ∈ S j), f i ⟨x, hxi⟩ = f j ⟨x, hxj⟩ } + {T : set α} {hT : T = Union S} + {i : ι} (x : S i) + (hx : (x : α) ∈ T := hT.symm ▸ set.subset_Union S i x.prop) : + lift_of_eq_Union S f hf T hT ⟨x, hx⟩ = f i x := +begin + subst hT, + delta lift_of_eq_Union, + exact Union_lift_mk x hx +end + +lemma lift_of_eq_Union_of_mem + {S : ι → set α} + {f : Π i (x : S i), β} + {hf : ∀ i j (x : α) (hxi : x ∈ S i) (hxj : x ∈ S j), f i ⟨x, hxi⟩ = f j ⟨x, hxj⟩ } + {T : set α} {hT : T = Union S} + (x : T) {i : ι} (hx : (x : α) ∈ S i) : + lift_of_eq_Union S f hf T hT x = f i ⟨x, hx⟩ := +begin + subst hT, + delta lift_of_eq_Union, + exact Union_lift_of_mem x hx +end + /-- Version of `Union_lift_const` for `lift_of_eq_Union`. -/ lemma lift_of_eq_Union_const {S : ι → set α} From 6c5bf79d11c00723d58653d68ef514c9c50d5914 Mon Sep 17 00:00:00 2001 From: Chris Hughes Date: Mon, 6 Sep 2021 12:57:08 +0100 Subject: [PATCH 30/51] nolint and explanation --- src/data/set/Union_lift.lean | 4 ++++ 1 file changed, 4 insertions(+) diff --git a/src/data/set/Union_lift.lean b/src/data/set/Union_lift.lean index 90ac5f5cdd284..1b072f970a996 100644 --- a/src/data/set/Union_lift.lean +++ b/src/data/set/Union_lift.lean @@ -44,8 +44,12 @@ variables {α ι : Type*} {β : Type*} namespace set +/- The unused argument `hf` is left in the definition so that the `simp` lemmas +`Union_lift_inclusion` will work without the user having to provide `hf` explicitly to +simplify terms involving `Union_lift`. -/ /-- Given a Union of sets `Union S`, define a function on the Union by defining it on each component, and proving that it agrees on the intersections. -/ +@[nolint unused_arguments] noncomputable def Union_lift (S : ι → set α) (f : Π i (x : S i), β) (hf : ∀ i j (x : α) (hxi : x ∈ S i) (hxj : x ∈ S j), f i ⟨x, hxi⟩ = f j ⟨x, hxj⟩) From 1176213da99cfbd3de3b60f3a4983606feff4449 Mon Sep 17 00:00:00 2001 From: Chris Hughes Date: Mon, 6 Sep 2021 12:58:02 +0100 Subject: [PATCH 31/51] nolint and explanation --- src/data/set/Union_lift.lean | 6 +++++- 1 file changed, 5 insertions(+), 1 deletion(-) diff --git a/src/data/set/Union_lift.lean b/src/data/set/Union_lift.lean index 90ac5f5cdd284..053e918ccd57d 100644 --- a/src/data/set/Union_lift.lean +++ b/src/data/set/Union_lift.lean @@ -40,12 +40,16 @@ or binary functions are preserved. These lemmas are: directed union, directed supremum -/ -variables {α ι : Type*} {β : Type*} +variables {α ι β : Type*} namespace set +/- The unused argument `hf` is left in the definition so that the `simp` lemmas +`Union_lift_inclusion` will work without the user having to provide `hf` explicitly to +simplify terms involving `Union_lift`. -/ /-- Given a Union of sets `Union S`, define a function on the Union by defining it on each component, and proving that it agrees on the intersections. -/ +@[nolint unused_arguments] noncomputable def Union_lift (S : ι → set α) (f : Π i (x : S i), β) (hf : ∀ i j (x : α) (hxi : x ∈ S i) (hxj : x ∈ S j), f i ⟨x, hxi⟩ = f j ⟨x, hxj⟩) From 5a4a4b8e0d91e4ee72e9cbafa2e05491d034a0f6 Mon Sep 17 00:00:00 2001 From: Chris Hughes Date: Tue, 7 Sep 2021 11:56:55 +0100 Subject: [PATCH 32/51] more work on lift --- src/field_theory/algebraic_closure.lean | 102 ++++++++++++++---------- 1 file changed, 59 insertions(+), 43 deletions(-) diff --git a/src/field_theory/algebraic_closure.lean b/src/field_theory/algebraic_closure.lean index 9d9f4e83e1de6..ed8a60d087d67 100644 --- a/src/field_theory/algebraic_closure.lean +++ b/src/field_theory/algebraic_closure.lean @@ -40,6 +40,21 @@ universes u v w noncomputable theory open_locale classical big_operators open polynomial +-- #print ideal.quotient.lift_mk +-- open ideal + +-- @[simp] lemma ideal.quotient.lift_comp_mk {R : Type*} [comm_ring R] +-- {S : Type v} [comm_ring S] {I : ideal R} (f : R →+* S) (H : I ≤ f.ker) : +-- (ideal.quotient.lift I f H).comp I^.quotient.mk = f := +-- by ext; simp + +-- lemma ker_quotient_lift {R : Type} [comm_ring R] +-- {S : Type v} [comm_ring S] {I : ideal R} (f : R →+* S) (H : I ≤ f.ker) : +-- (ideal.quotient.lift I f H).ker = (f.ker).map I^.quotient.mk := +-- by rwa [← (comap_injective_of_surjective _ quotient.mk_surjective).eq_iff, +-- ring_hom.ker_eq_comap_bot, comap_comap, ← ring_hom.ker_eq_comap_bot, +-- comap_map_of_surjective _ quotient.mk_surjective, +-- ideal.quotient.lift_comp_mk, ← ring_hom.ker_eq_comap_bot, mk_ker, left_eq_sup] variables (k : Type u) [field k] @@ -69,6 +84,9 @@ polynomial.splits_of_splits_id _ $ is_alg_closed.splits _ namespace is_alg_closed +theorem exists_root [is_alg_closed k] (p : polynomial k) (hp : p.degree ≠ 0) : ∃ x, is_root p x := +exists_root_of_splits _ (is_alg_closed.splits p) hp + theorem of_exists_root (H : ∀ p : polynomial k, p.monic → irreducible p → ∃ x, p.eval x = 0) : is_alg_closed k := ⟨λ p, or.inr $ λ q hq hqp, @@ -397,6 +415,10 @@ variables {E₁ E₂ E₃ : subfield_with_hom K L M hL} instance : has_le (subfield_with_hom K L M hL) := { le := λ E₁ E₂, ∃ h : E₁.carrier ≤ E₂.carrier, ∀ x, E₂.emb (inclusion h x) = E₁.emb x } +instance : inhabited (subfield_with_hom K L M hL) := +⟨{ carrier := ⊥, + emb := (algebra.of_id K M).comp (algebra.bot_equiv K L).to_alg_hom }⟩ + lemma le_def : E₁ ≤ E₂ ↔ ∃ h : E₁.carrier ≤ E₂.carrier, ∀ x, E₂.emb (inclusion h x) = E₁.emb x := iff.rfl @@ -425,11 +447,11 @@ end subfield_with_hom open lattice -def maximal_subfield_with_hom_chain (c : set (subfield_with_hom K L M hL)) - [nonempty c] - (hc : chain (≤) c) : +lemma maximal_subfield_with_hom_chain_bounded (c : set (subfield_with_hom K L M hL)) + (hc : chain (≤) c) (hcn : c.nonempty) : ∃ ub : subfield_with_hom K L M hL, ∀ N, N ∈ c → N ≤ ub := let ub : subfield_with_hom K L M hL := +by haveI : nonempty c := set.nonempty.to_subtype hcn; exact { carrier := ⨆ i : c, (i : subfield_with_hom K L M hL).carrier, emb := subalgebra.supr_lift (λ i : c, (i : subfield_with_hom K L M hL).carrier) @@ -448,17 +470,15 @@ let ub : subfield_with_hom K L M hL := begin intro x, simp [ub], - + refl end⟩⟩ - -- ⟨le_supr (λ i : c, (i : subfield_with_hom K L M hL).carrier) ⟨N, hN⟩, begin - - --end⟩⟩ variables (hL M) lemma exists_maximal_subfield_with_hom : ∃ E : subfield_with_hom K L M hL, ∀ N, E ≤ N → N ≤ E := -zorn (maximal_subfield_with_hom_chain) (λ _ _ _, le_trans) +zorn.exists_maximal_of_nonempty_chains_bounded + maximal_subfield_with_hom_chain_bounded (λ _ _ _, le_trans) def maximal_subfield_with_hom : subfield_with_hom K L M hL := classical.some (exists_maximal_subfield_with_hom M hL) @@ -478,42 +498,38 @@ begin -- { sorry }, -- { exact { .. E.emb.hom } } end - -lemma maximal_subfield_with_hom_surj : - surjective (maximal_subfield_with_hom M hL).carrier.val := +-- N : subalgebra K L +-- #print power_basis +lemma maximal_subfield_with_hom_surj (x : L) : + x ∈ (maximal_subfield_with_hom M hL).carrier := begin - intros x, refine ⟨⟨x, _⟩, rfl⟩, - replace hx := (maximal_subfield_with_hom M hL).carrier.is_integral x (hL x), - let p := minimal_polynomial hx, - have H := exists_root M (p.map (maximal_subfield_with_hom M hL).emb) _, - swap, - { calc 0 < degree p : - begin - sorry -- minimal_polynomial.degree_pos gives diamonds - end - ... = degree (p.map (maximal_subfield_with_hom M hL).emb) : - begin - symmetry, - -- refine @polynomial.degree_map' _ _ _ _ _ _ p _ (by exact is_ring_hom.is_semiring_hom _) _, - sorry, - end }, - let y := classical.some H, - let f := algebra.adjoin_singleton_desc x hx - (maximal_subfield_with_hom M hL).emb y (classical.some_spec H), - let tmpa : subalgebra _ L := algebra.adjoin _ ({x} : set L), - let tmpb : _ := _, - let N : subfield_with_hom K L M hL := - { carrier := subalgebra.under (maximal_subfield_with_hom M hL).carrier tmpa, - emb := - { to_fun := f, - hom := algebra.adjoin_singleton_desc.is_ring_hom _ _ _ _ _, - commutes' := tmpb } }, - have hN : x ∈ N.carrier := algebra.subset_adjoin (mem_singleton x), - refine subfield_with_hom.subalgebra_le_of_le (maximal_subfield_with_hom_is_maximal M hL N _) hN, - { refine ⟨λ l hl, ring.subset_closure (mem_union_left _ _), _⟩, - { rw mem_range, refine ⟨⟨l, hl⟩, rfl⟩ }, - { sorry } }, - { sorry } + let p := minpoly K x, + let N : subalgebra K L := (maximal_subfield_with_hom M hL).carrier, + letI : field N := is_field.to_field _ (subalgebra.is_field_of_algebraic N + (λ y, (is_algebraic_iff_is_integral K).2 (hL y))), + letI : algebra N M := (maximal_subfield_with_hom M hL).emb.to_ring_hom.to_algebra, + cases is_alg_closed.exists_root M ((minpoly N x).map (algebra_map _ _)) sorry with y hy, + let larger_carrier : subalgebra K L := N.under (algebra.adjoin N ({x} : set L)), + letI : is_scalar_tower K N L := begin apply_instance end, + letI : algebra K N := infer_instance, + let O : subalgebra N L := algebra.adjoin N {(x : L)}, + letI : is_scalar_tower N O L := by apply_instance, + -- letI : module K O.to_submodule := begin dsimp [O], apply_instance end, + -- have larger_emb := ((adjoin_root.lift_hom (minpoly N x) y sorry).comp + -- (alg_equiv.adjoin_singleton_equiv_adjoin_root_minpoly N x).to_alg_hom), + -- let f : tmpa →ₐ[↥(maximal_subfield_with_hom M hL).carrier] M := _, + -- let N : subfield_with_hom K L M hL := + -- { carrier := subalgebra.under (maximal_subfield_with_hom M hL).carrier tmpa, + -- emb := + -- { to_fun := f, + -- hom := algebra.adjoin_singleton_desc.is_ring_hom _ _ _ _ _, + -- commutes' := tmpb } }, + -- have hN : x ∈ N.carrier := algebra.subset_adjoin (mem_singleton x), + -- refine subfield_with_hom.subalgebra_le_of_le (maximal_subfield_with_hom_is_maximal M hL N _) hN, + -- { refine ⟨λ l hl, ring.subset_closure (mem_union_left _ _), _⟩, + -- { rw mem_range, refine ⟨⟨l, hl⟩, rfl⟩ }, + -- { sorry } }, + -- { sorry } end end lift From ab2f3d57c4664f31b34116473f45081b4712b6e9 Mon Sep 17 00:00:00 2001 From: Chris Hughes <33847686+ChrisHughes24@users.noreply.github.com> Date: Thu, 9 Sep 2021 08:09:30 +0100 Subject: [PATCH 33/51] feat(algebra/algebra/subalgebra): mem_under --- src/algebra/algebra/subalgebra.lean | 4 ++++ 1 file changed, 4 insertions(+) diff --git a/src/algebra/algebra/subalgebra.lean b/src/algebra/algebra/subalgebra.lean index e1d34399da93c..f1cd58b90719e 100644 --- a/src/algebra/algebra/subalgebra.lean +++ b/src/algebra/algebra/subalgebra.lean @@ -798,6 +798,10 @@ def under {R : Type u} {A : Type v} [comm_semiring R] [comm_semiring A] { algebra_map_mem' := λ r, T.algebra_map_mem ⟨algebra_map R A r, S.algebra_map_mem r⟩, .. T } +@[simp] lemma mem_under {R : Type u} {A : Type v} [comm_semiring R] [comm_semiring A] + {i : algebra R A} {S : subalgebra R A} {T : subalgebra S A} {x : A} : + x ∈ S.under T ↔ x ∈ T := iff.rfl + end actions end subalgebra From f9a9ea6a9424334ddbfb2ef85c87a3ceb7d80303 Mon Sep 17 00:00:00 2001 From: Chris Hughes Date: Thu, 9 Sep 2021 09:46:30 +0100 Subject: [PATCH 34/51] feat(field_theory/algebraic_closure): map from algebraic extensions into the algebraic closure --- src/field_theory/algebraic_closure.lean | 947 ++---------------------- src/ring_theory/algebraic.lean | 5 + 2 files changed, 57 insertions(+), 895 deletions(-) diff --git a/src/field_theory/algebraic_closure.lean b/src/field_theory/algebraic_closure.lean index ed8a60d087d67..a7e015e8bc9bb 100644 --- a/src/field_theory/algebraic_closure.lean +++ b/src/field_theory/algebraic_closure.lean @@ -84,9 +84,22 @@ polynomial.splits_of_splits_id _ $ is_alg_closed.splits _ namespace is_alg_closed +variables {k} + theorem exists_root [is_alg_closed k] (p : polynomial k) (hp : p.degree ≠ 0) : ∃ x, is_root p x := exists_root_of_splits _ (is_alg_closed.splits p) hp +theorem exists_eval₂_eq_zero {R : Type*} [field R] [is_alg_closed k] (f : R →+* k) + (p : polynomial R) (hp : p.degree ≠ 0) : ∃ x, p.eval₂ f x = 0 := +let ⟨x, hx⟩ := exists_root (p.map f) (by rwa [degree_map]) in +⟨x, by rwa [eval₂_eq_eval_map, ← is_root]⟩ + +variables (k) + +theorem exists_aeval_eq_zero {R : Type*} [field R] [is_alg_closed k] [algebra R k] + (p : polynomial R) (hp : p.degree ≠ 0) : ∃ x : k, aeval x p = 0 := +exists_eval₂_eq_zero (algebra_map R k) p hp + theorem of_exists_root (H : ∀ p : polynomial k, p.monic → irreducible p → ∃ x, p.eval x = 0) : is_alg_closed k := ⟨λ p, or.inr $ λ q hq hqp, @@ -395,7 +408,7 @@ namespace lift closed extension is proven to exist. The assumption that M is algebraically closed could probably easily be switched to an assumption that M contains all the roots of polynomials in K -/ variables {K : Type u} {L : Type v} {M : Type w} [field K] [field L] [algebra K L] - [field M] [algebra K M] [is_alg_closed M] (hL : ∀ x : L, is_integral K x) + [field M] [algebra K M] [is_alg_closed M] (hL : algebra.is_algebraic K L) variables (K L M) include hL @@ -403,9 +416,9 @@ open zorn subalgebra alg_hom function /-- This structure is used to prove the existence of a homomorphism from any algebraic extension into an algebraic closure -/ -private structure subfield_with_hom := +structure subfield_with_hom := (carrier : subalgebra K L) -(emb : (@alg_hom K _ M _ _ _ (subalgebra.algebra carrier) _)) +(emb : carrier →ₐ[K] M) variables {K L M hL} @@ -430,20 +443,7 @@ instance : preorder (subfield_with_hom K L M hL) := le_refl := λ E, ⟨le_refl _, by simp⟩, le_trans := λ E₁ E₂ E₃ h₁₂ h₂₃, ⟨le_trans h₁₂.fst h₂₃.fst, - λ _, by erw [← inclusion_inclusion h₁₂.fst h₂₃.fst, compat, compat]⟩, - -- le_antisymm := λ a b hab hba, begin - -- cases a with a aemb, - -- cases b with b bemb, - -- have h : a = b, from le_antisymm hab.fst hba.fst, - -- subst b, - -- simp only [true_and, eq_self_iff_true, heq_iff_eq], - -- ext x, - -- refine eq.trans (hab.snd x).symm _, - -- simp - -- end - } - -end subfield_with_hom + λ _, by erw [← inclusion_inclusion h₁₂.fst h₂₃.fst, compat, compat]⟩ } open lattice @@ -480,6 +480,7 @@ lemma exists_maximal_subfield_with_hom : ∃ E : subfield_with_hom K L M hL, zorn.exists_maximal_of_nonempty_chains_bounded maximal_subfield_with_hom_chain_bounded (λ _ _ _, le_trans) +/-- The maximal `subfield_with_hom`. We later prove that this is equal to `⊤`. -/ def maximal_subfield_with_hom : subfield_with_hom K L M hL := classical.some (exists_maximal_subfield_with_hom M hL) @@ -488,895 +489,51 @@ lemma maximal_subfield_with_hom_is_maximal : (maximal_subfield_with_hom M hL) ≤ N → N ≤ (maximal_subfield_with_hom M hL) := classical.some_spec (exists_maximal_subfield_with_hom M hL) -lemma emb_injective (E : subfield_with_hom K L M hL) : - injective E.emb := -begin - sorry - -- let tmpa : _ := _, let tmpb : _ := _, - -- refine @is_field_hom.injective _ M tmpa _ _ tmpb, - -- swap, - -- { sorry }, - -- { exact { .. E.emb.hom } } -end --- N : subalgebra K L --- #print power_basis -lemma maximal_subfield_with_hom_surj (x : L) : - x ∈ (maximal_subfield_with_hom M hL).carrier := +lemma maximal_subfield_with_hom_eq_top : + (maximal_subfield_with_hom M hL).carrier = ⊤ := begin + rw [eq_top_iff], + intros x _, let p := minpoly K x, let N : subalgebra K L := (maximal_subfield_with_hom M hL).carrier, - letI : field N := is_field.to_field _ (subalgebra.is_field_of_algebraic N - (λ y, (is_algebraic_iff_is_integral K).2 (hL y))), + letI : field N := is_field.to_field _ (subalgebra.is_field_of_algebraic N hL), letI : algebra N M := (maximal_subfield_with_hom M hL).emb.to_ring_hom.to_algebra, - cases is_alg_closed.exists_root M ((minpoly N x).map (algebra_map _ _)) sorry with y hy, - let larger_carrier : subalgebra K L := N.under (algebra.adjoin N ({x} : set L)), - letI : is_scalar_tower K N L := begin apply_instance end, - letI : algebra K N := infer_instance, + cases is_alg_closed.exists_aeval_eq_zero M (minpoly N x) + (ne_of_gt (minpoly.degree_pos + ((is_algebraic_iff_is_integral _).1 (algebra.is_algebraic_of_larger_field hL x)))) with y hy, let O : subalgebra N L := algebra.adjoin N {(x : L)}, - letI : is_scalar_tower N O L := by apply_instance, - -- letI : module K O.to_submodule := begin dsimp [O], apply_instance end, - -- have larger_emb := ((adjoin_root.lift_hom (minpoly N x) y sorry).comp - -- (alg_equiv.adjoin_singleton_equiv_adjoin_root_minpoly N x).to_alg_hom), - -- let f : tmpa →ₐ[↥(maximal_subfield_with_hom M hL).carrier] M := _, - -- let N : subfield_with_hom K L M hL := - -- { carrier := subalgebra.under (maximal_subfield_with_hom M hL).carrier tmpa, - -- emb := - -- { to_fun := f, - -- hom := algebra.adjoin_singleton_desc.is_ring_hom _ _ _ _ _, - -- commutes' := tmpb } }, - -- have hN : x ∈ N.carrier := algebra.subset_adjoin (mem_singleton x), - -- refine subfield_with_hom.subalgebra_le_of_le (maximal_subfield_with_hom_is_maximal M hL N _) hN, - -- { refine ⟨λ l hl, ring.subset_closure (mem_union_left _ _), _⟩, - -- { rw mem_range, refine ⟨⟨l, hl⟩, rfl⟩ }, - -- { sorry } }, - -- { sorry } + let larger_emb := ((adjoin_root.lift_hom (minpoly N x) y hy).comp + (alg_equiv.adjoin_singleton_equiv_adjoin_root_minpoly N x).to_alg_hom), + have hNO : N ≤ N.under O, + { intros z hz, + show algebra_map N L ⟨z, hz⟩ ∈ O, + exact O.algebra_map_mem _ }, + let O' : subfield_with_hom K L M hL := + { carrier := N.under O, + emb := larger_emb.restrict_scalars K }, + have hO' : maximal_subfield_with_hom M hL ≤ O', + { refine ⟨hNO, _⟩, + intros z, + show O'.emb (algebra_map N O z) = algebra_map N M z, + simp only [O', restrict_scalars_apply, alg_hom.commutes] }, + refine (maximal_subfield_with_hom_is_maximal M hL O' hO').fst _, + exact algebra.subset_adjoin (set.mem_singleton x), end +end subfield_with_hom end lift -variables {L : Type v} {M : Type w} [discrete_field L] [algebra K L] - [discrete_field M] [algebra K M] [is_algebraically_closed M] (hL : ∀ x : L, is_integral K x) +namespace is_alg_closed + +variables {K : Type u} [field K] {L : Type v} {M : Type w} [field L] [algebra K L] + [field M] [algebra K M] [is_alg_closed M] (hL : algebra.is_algebraic K L) variables (K L M) include hL -/-- A (random) hom from an algebraic extension of K into an algebraic closure -/ -def lift : L →ₐ[K] M := -(lift.maximal_subfield_with_hom M hL).emb.comp $ alg_equiv.to_alg_hom $ alg_equiv.symm -{ ..(lift.maximal_subfield_with_hom M hL).carrier.val, - ..equiv.of_bijective - ⟨subtype.val_injective, lift.maximal_subfield_with_hom_surj _ _⟩ } - - --- /- --- Copyright (c) 2019 Chris Hughes. All rights reserved. --- Released under Apache 2.0 license as described in the file LICENSE. --- Authors: Chris Hughes, Johan Commelin --- -/ - --- import algebra.direct_limit --- import set_theory.schroeder_bernstein --- import field_theory.algebraic +/-- A (random) hom from an algebraic extension of K into an algebraically closed extension of K -/ +@[irreducible] def lift : L →ₐ[K] M := +(lift.subfield_with_hom.maximal_subfield_with_hom M hL).emb.comp $ + eq.rec_on (lift.subfield_with_hom.maximal_subfield_with_hom_eq_top M hL).symm algebra.to_top --- noncomputable theory --- open_locale classical - --- universes u v w --- open polynomial zorn set function --- variables {K : Type u} [discrete_field K] - --- /- Turn down the instance priority for subtype.decidable_eq and use classical.dec_eq everywhere, --- to avoid diamonds -/ --- local attribute [instance, priority 0] subtype.decidable_eq - --- lemma injective_eq {α : Sort*} : injective (eq : α → α → Prop) := --- λ _ _ h, h.symm ▸ rfl - --- @[instance] lemma equiv.is_ring_hom {α β : Type*} [ring β] (e : α ≃ β) : --- @is_ring_hom β α _ (equiv.ring e) e.symm := --- by split; simp [equiv.mul_def, equiv.add_def, equiv.one_def] - --- instance equiv.is_ring_hom.symm {α β : Type*} [ring β] (e : α ≃ β) : --- @is_ring_hom α β (equiv.ring e) _ e := --- by split; simp [equiv.mul_def, equiv.add_def, equiv.one_def] - --- def equiv.add_equiv {α β : Type*} [has_add β] (e : α ≃ β) : --- @add_equiv α β e.has_add _ := --- { map_add' := λ x y, e.apply_symm_apply _, --- ..e } - --- def equiv.mul_equiv {α β : Type*} [has_mul β] (e : α ≃ β) : --- @mul_equiv α β e.has_mul _ := --- { map_mul' := λ x y, e.apply_symm_apply _, --- ..e } - --- def equiv.ring_equiv {α β : Type*} [has_mul β] [has_add β] (e : α ≃ β) : --- @ring_equiv α β e.has_mul e.has_add _ _ := --- { map_add' := λ x y, e.apply_symm_apply _, --- map_mul' := λ x y, e.apply_symm_apply _, --- ..e } - --- lemma exists_root_of_equiv {α β : Type*} [comm_ring α] [comm_ring β] (e : α ≃+* β) --- {f : polynomial α} {x : β} (hx : f.eval₂ e x = 0) : --- f.eval (e.symm x) = 0 := --- begin --- have e_inj : function.injective e := e.to_equiv.injective, --- apply e_inj, --- rw [← eval₂_hom e, e.apply_symm_apply, is_ring_hom.map_zero e, hx], --- end - --- namespace alg_hom --- variables {α : Type u} {β : Type v} {γ : Type w} [comm_ring α] [ring β] [ring γ] --- [algebra α β] [algebra α γ] (f : β →ₐ[α] γ) - --- def inverse (g : γ → β) (h₁ : left_inverse g f) (h₂ : right_inverse g f) : γ →ₐ[α] β := --- by dsimp [left_inverse, function.right_inverse] at h₁ h₂; exact --- { to_fun := g, --- hom := show is_ring_hom g, from --- { map_add := λ x y, by rw [← h₁ (g x + g y), f.map_add, h₂, h₂], --- map_mul := λ x y, by rw [← h₁ (g x * g y), f.map_mul, h₂, h₂], --- map_one := by rw [← h₁ 1, f.map_one] }, --- commutes' := λ a, by rw [← h₁ (algebra_map β a), f.commutes] } - --- instance quotient.algebra (I : ideal α) : algebra α I.quotient := --- algebra.of_ring_hom (ideal.quotient.mk I) (ideal.quotient.is_ring_hom_mk I) - --- def induced_quotient_map (I J : ideal α) (h : I ≤ J) : --- I.quotient →ₐ[α] J.quotient := --- { to_fun := ideal.quotient.lift I (ideal.quotient.mk J) $ --- by { intros i hi, erw submodule.quotient.mk_eq_zero, exact h hi }, --- commutes' := λ a, by { erw ideal.quotient.lift_mk, refl } } - --- end alg_hom - --- set_option old_structure_cmd true - --- structure alg_equiv (α β γ : Type*) [comm_ring α] [ring β] [ring γ] --- [algebra α β] [algebra α γ] extends alg_hom α β γ, equiv β γ - --- set_option old_structure_cmd false - --- infix ` ≃ₐ `:25 := alg_equiv _ --- notation A ` ≃ₐ[`:25 R `] ` B := alg_equiv R A B - --- namespace alg_equiv --- variables {α : Type u} {β : Type v} {γ : Type w} [comm_ring α] [ring β] [ring γ] --- [algebra α β] [algebra α γ] - --- protected def refl : β ≃ₐ[α] β := --- { hom := is_ring_hom.id, --- commutes' := λ b, rfl, --- .. equiv.refl β } - --- protected def symm (e : β ≃ₐ[α] γ) : γ ≃ₐ[α] β := --- { .. e.to_alg_hom.inverse e.inv_fun e.left_inv e.right_inv, --- .. e.to_equiv.symm } - --- -- TODO: trans - --- end alg_equiv - --- namespace adjoin_root --- variables {α : Type*} [comm_ring α] [decidable_eq α] (f : polynomial α) - --- instance : algebra α (adjoin_root f) := --- algebra.of_ring_hom coe $ by apply_instance - --- lemma fg_of_monic (hf : f.monic) : submodule.fg (⊤ : submodule α (adjoin_root f)) := --- begin --- refine ⟨(finset.range (f.nat_degree + 1)).image (λ i, adjoin_root.mk (X^i)), _⟩, --- rw submodule.eq_top_iff', --- intro a, --- apply quotient.induction_on' a, --- intro p, show mk p ∈ _, --- rw [← mod_by_monic_add_div p hf, is_ring_hom.map_add mk, is_ring_hom.map_mul mk], --- { refine submodule.add_mem _ _ _, --- { apply_instance }, --- { rw [(p %ₘ f).as_sum, ← finset.sum_hom mk], --- { refine submodule.sum_mem _ _, --- intros i hi, --- rw [is_ring_hom.map_mul mk], --- { show algebra_map _ (coeff (p %ₘ f) i) * _ ∈ _, --- rw ← algebra.smul_def, --- refine submodule.smul_mem _ _ (submodule.subset_span _), --- rw [finset.mem_coe, finset.mem_image], --- refine ⟨i, _, rfl⟩, --- rw finset.mem_range at hi ⊢, --- refine lt_of_lt_of_le hi (nat.succ_le_succ _), --- by_cases (p %ₘ f) = 0, { simp [h] }, --- rw ← with_bot.coe_le_coe, --- rw ← degree_eq_nat_degree h, --- apply le_trans (degree_mod_by_monic_le p hf), --- exact degree_le_nat_degree }, --- { apply_instance } }, --- { apply_instance } }, --- { convert submodule.zero_mem _, --- convert zero_mul _ using 2, --- erw [submodule.quotient.mk_eq_zero _], --- apply submodule.subset_span, --- exact mem_singleton _ } }, --- all_goals { apply_instance } --- end - --- open submodule - --- def adjoin_root_of_monic (f : polynomial K) : --- adjoin_root (f * C (leading_coeff f)⁻¹) →ₐ[K] adjoin_root f := --- { to_fun := ideal.quotient.lift _ mk $ λ g hg, --- begin --- erw quotient.mk_eq_zero, --- rw ideal.mem_span_singleton' at hg ⊢, --- rcases hg with ⟨g, rfl⟩, rw [mul_comm f, ← mul_assoc], --- exact ⟨_, rfl⟩, --- end, --- hom := ideal.quotient.is_ring_hom, --- commutes' := λ g, rfl } - --- lemma fg {f : polynomial K} (hf : f ≠ 0) : submodule.fg (⊤ : submodule K (adjoin_root f)) := --- begin --- let φ := adjoin_root_of_monic f, --- have trick := fg_of_monic _ (monic_mul_leading_coeff_inv hf), --- convert fg_map trick, swap, exact φ.to_linear_map, --- { refine (submodule.eq_top_iff'.mpr _).symm, --- intros x, apply quotient.induction_on' x, clear x, --- intro g, --- rw mem_map, --- use [mk g, mem_top, rfl] } --- end --- . - --- lemma is_integral {f : polynomial K} (hf : f ≠ 0) (x : adjoin_root f) : is_integral K x := --- begin --- refine is_integral_of_mem_of_fg ⊤ _ x algebra.mem_top, --- convert fg hf, --- rw eq_top_iff', intro y, exact algebra.mem_top, --- end - --- end adjoin_root - --- section thing - --- local attribute [instance] classical.dec - --- private lemma thing_aux {X : Type u} {Y : Type v} {Z : Type w} (fxy : X ↪ Y) (fxz : X ↪ Z) --- (hYZ : (Z ↪ Y) → false) : ↥-range fxy.1 ↪ ↥-range fxz.1 := --- classical.choice $ or.resolve_left embedding.total $ --- λ ⟨f⟩, hYZ $ --- calc Z ↪ range fxz ⊕ ↥-range fxz : --- (equiv.set.sum_compl _).symm.to_embedding --- ... ↪ range fxy ⊕ ↥-range fxy : --- embedding.sum_congr --- (((equiv.set.range _ fxz.2).symm.to_embedding).trans --- (equiv.set.range _ fxy.2).to_embedding) --- f --- ... ↪ Y : (equiv.set.sum_compl _).to_embedding - --- private def thing {X : Type u} {Y : Type v} {Z : Type w} (fxy : X ↪ Y) (fxz : X ↪ Z) --- (hYZ : (Z ↪ Y) → false) : Y ↪ Z := --- calc Y ↪ range fxy ⊕ ↥-range fxy : (equiv.set.sum_compl _).symm.to_embedding --- ... ↪ range fxz ⊕ ↥-range fxz : embedding.sum_congr --- ((equiv.set.range _ fxy.2).symm.to_embedding.trans --- (equiv.set.range _ fxz.2).to_embedding) --- (thing_aux fxy fxz hYZ) --- ... ↪ Z : (equiv.set.sum_compl _).to_embedding - --- private lemma thing_commutes {X : Type u} {Y : Type v} {Z : Type w} (fxy : X ↪ Y) (fxz : X ↪ Z) --- (hYZ : (Z ↪ Y) → false) (x : X) : thing fxy fxz hYZ (fxy x) = fxz x := --- have (⟨fxy x, mem_range_self _⟩ : range fxy) = equiv.set.range _ fxy.2 x, from rfl, --- begin --- dsimp only [thing, embedding.trans_apply, equiv.trans_apply, function.comp, --- equiv.to_embedding_coe_fn], --- simp only [equiv.set.sum_compl_symm_apply_of_mem (mem_range_self _), --- embedding.sum_congr_apply_inl, equiv.set.sum_compl_apply_inl, --- embedding.trans_apply, equiv.to_embedding_coe_fn, this, equiv.symm_apply_apply], --- refl --- end - --- end thing - --- class is_algebraically_closed (K : Type u) [nonzero_comm_ring K] [decidable_eq K] := --- (exists_root : ∀ f : polynomial K, 0 < degree f → ∃ x, is_root f x) - --- section is_algebraically_closed - --- lemma is_algebraically_closed_of_irreducible_has_root --- (h : ∀ f : polynomial K, irreducible f → ∃ x, is_root f x) : --- is_algebraically_closed K := --- ⟨λ f hf0, let ⟨g, hg⟩ := is_noetherian_ring.exists_irreducible_factor --- (show ¬ is_unit f, from λ h, by rw [is_unit_iff_degree_eq_zero] at h; --- rw h at hf0; exact lt_irrefl _ hf0) --- (λ h, by rw ← degree_eq_bot at h; --- rw h at hf0; exact absurd hf0 dec_trivial) in --- let ⟨x, hx⟩ := h g hg.1 in --- let ⟨i, hi⟩ := hg.2 in --- ⟨x, by rw [hi, is_root.def, eval_mul, show _ = _, from hx, zero_mul]⟩⟩ - --- variables (L : Type*) [nonzero_comm_ring L] [decidable_eq L] [is_algebraically_closed L] - --- lemma exists_root (f : polynomial L) (hf : 0 < f.degree) : --- ∃ x, is_root f x := --- is_algebraically_closed.exists_root f hf - --- -- /- An algebraic extension of -/ --- -- lemma equiv_of_algebraic - --- end is_algebraically_closed - --- --move this --- namespace polynomial --- variables (R : Type u) (A : Type v) --- variables [comm_ring R] [comm_ring A] [algebra R A] --- variables [decidable_eq R] (x : A) - --- @[simp] lemma aeval_X : aeval R A x X = x := eval₂_X _ x - --- @[simp] lemma aeval_C (r : R) : aeval R A x (C r) = algebra_map A r := eval₂_C _ x - --- end polynomial - --- namespace algebraic_closure - --- section classical - --- local attribute [instance, priority 1] classical.dec - --- /-- The `big_type` with cardinality strictly larger than any algebraic extension -/ --- private def big_type (K : Type u) [discrete_field K] := set (ℕ × polynomial K) - --- private def algebraic_embedding_aux {L : Type*} [discrete_field L] [algebra K L] --- (h : ∀ l : L, is_integral K l) (x : L) : ℕ × polynomial K := --- let f := classical.some (h x) in --- ⟨list.index_of x (quotient.out ((f.map (algebra_map L)).roots.1)), f⟩ - --- private lemma algebraic_embedding_aux_injective --- {L : Type*} [discrete_field L] [algebra K L] --- (h : ∀ l : L, is_integral K l) : injective (algebraic_embedding_aux h) := --- λ x y hxy, --- let f := classical.some (h x) in --- let g := classical.some (h y) in --- have hf : monic f ∧ aeval K L x f = 0, from classical.some_spec (h x), --- have hg : monic g ∧ aeval K L y g = 0, from classical.some_spec (h y), --- have hfg : f = g, from (prod.ext_iff.1 hxy).2, --- have hfg' : list.index_of x (quotient.out ((f.map (algebra_map L)).roots.1)) = --- list.index_of y (quotient.out ((f.map (algebra_map L)).roots.1)), --- from (prod.ext_iff.1 hxy).1.trans (hfg.symm ▸ rfl), --- have hx : x ∈ quotient.out ((f.map (algebra_map L)).roots.1), --- from multiset.mem_coe.1 begin --- show x ∈ quotient.mk _, --- rw [quotient.out_eq, ← finset.mem_def, mem_roots (mt (map_eq_zero (algebra_map L)).1 --- (ne_zero_of_monic hf.1)), is_root.def, eval_map, ← aeval_def, hf.2], --- end, --- have hy : y ∈ quotient.out ((g.map (algebra_map L)).roots.1), --- from multiset.mem_coe.1 begin --- show y ∈ quotient.mk _, --- rw [quotient.out_eq, ← finset.mem_def, mem_roots (mt (map_eq_zero (algebra_map L)).1 --- (ne_zero_of_monic hg.1)), is_root.def, eval_map, ← aeval_def, hg.2], --- end, --- (list.index_of_inj hx (by rwa hfg)).1 hfg' - --- private def algebraic_embedding_big_type {L : Type*} [discrete_field L] [algebra K L] --- (h : ∀ l : L, is_integral K l) : L ↪ big_type K := --- ⟨_, injective_comp injective_eq $ algebraic_embedding_aux_injective h⟩ - --- private def algebraic_embedding {L : Type*} [discrete_field L] [algebra K L] --- (h : ∀ l : L, is_integral K l) : L ↪ ℕ × polynomial K := --- ⟨_, algebraic_embedding_aux_injective h⟩ - --- private def bembedding (K : Type u) [discrete_field K] : K ↪ big_type K := --- ⟨λ a, show set _, from {(0, X - C a)}, λ a b, by simp [C_inj]⟩ - --- instance range_bembedding.discrete_field : discrete_field (set.range (bembedding K)) := --- equiv.discrete_field (equiv.set.range _ (bembedding K).2).symm - --- private structure extension (K : Type u) [discrete_field K] : Type u := --- (carrier : set (big_type K)) --- [field : discrete_field ↥carrier] --- [algebra : algebra K ↥carrier] --- (algebraic : ∀ x : carrier, is_integral K x) - --- attribute [instance] extension.field extension.algebra - --- private def base_extension (K : Type u) [discrete_field K] : extension K := --- { carrier := set.range (bembedding K), --- algebra := algebra.of_ring_hom (equiv.set.range _ (bembedding K).2).symm.symm --- (by apply_instance), --- algebraic := --- begin --- rintro ⟨_, x, rfl⟩, --- refine ⟨X + C (-x), monic_X_add_C (-x), _⟩, --- rw [alg_hom.map_add, C_neg, alg_hom.map_neg, polynomial.aeval_X, polynomial.aeval_C], --- exact add_neg_self _ --- end } - --- -- /-- not used but might help with sorries -/ --- -- private def extension.of_algebraic {L : Type v} [discrete_field L] [algebra K L] --- -- (hL : ∀ x : L, is_integral K x) : extension K := --- -- { carrier := set.range (algebraic_embedding_big_type hL), --- -- field := equiv.discrete_field (equiv.set.range _ (algebraic_embedding_big_type hL).2).symm, --- -- algebra := sorry, -- a field isomorphic to an algebra is an algebra --- -- algebraic := sorry -- a field isomorphic to an algebraic extension is algebraic --- -- } - --- @[simp] lemma inclusion_refl {α : Type*} {s : set α} (h : s ⊆ s) : inclusion h = id := --- funext $ λ x, by { cases x, refl } - --- @[simp] lemma inclusion_trans {α : Type*} {s t u : set α} (hst : s ⊆ t) (htu : t ⊆ u) : --- inclusion (set.subset.trans hst htu) = inclusion htu ∘ inclusion hst := --- funext $ λ x, by { cases x, refl } - --- instance : preorder (extension K) := --- { le := λ L M, ∃ hLM : L.carrier ⊆ M.carrier, (is_ring_hom (inclusion hLM) ∧ --- (inclusion hLM ∘ (algebra_map L.carrier : K → L.carrier) = algebra_map M.carrier)), --- le_refl := λ L, --- begin --- use set.subset.refl L.carrier, --- rw inclusion_refl, --- exact ⟨is_ring_hom.id, comp.left_id _⟩ --- end, --- le_trans := λ L M N ⟨hLM₁, hLM₂, hLM₃⟩ ⟨hMN₁, hMN₂, hMN₃⟩, --- begin --- use set.subset.trans hLM₁ hMN₁, --- rw inclusion_trans, resetI, --- refine ⟨is_ring_hom.comp _ _, _⟩, --- rw [comp.assoc, hLM₃, hMN₃] --- end } - --- lemma le_def {L M : extension K} : --- L ≤ M ↔ ∃ hLM : L.carrier ⊆ M.carrier, (is_ring_hom (inclusion hLM) ∧ --- (inclusion hLM ∘ (algebra_map L.carrier : K → L.carrier) = algebra_map M.carrier)) := iff.rfl - --- lemma subset_of_le {L M : extension K} (h : L ≤ M) : L.carrier ⊆ M.carrier := --- by { rw le_def at h, rcases h with ⟨_,_,_⟩, assumption } - --- lemma ring_hom_of_le {L M : extension K} (h : L ≤ M) : --- is_ring_hom (inclusion $ subset_of_le h) := --- by { rw le_def at h, rcases h with ⟨_,_,_⟩, assumption } - --- lemma compat {L M : extension K} (h : L ≤ M) : --- inclusion (subset_of_le h) ∘ (algebra_map L.carrier : K → L.carrier) = algebra_map M.carrier := --- by { rw le_def at h, rcases h with ⟨_,_,_⟩, assumption } - --- local attribute [instance] ring_hom_of_le - --- private structure chain' (c : set (extension K)) : Prop := --- (chain : chain (≤) c) - --- local attribute [class] chain' - --- private lemma is_chain (c : set (extension K)) [chain' c]: chain (≤) c := --- chain'.chain (by apply_instance) - --- section chain - --- variables (c : set (extension K)) [hcn : nonempty c] --- include c hcn - --- variable [hcn' : chain' c] --- include hcn' - --- instance chain_directed_order : directed_order c := --- ⟨λ ⟨i, hi⟩ ⟨j, hj⟩, let ⟨k, hkc, hk⟩ := chain.directed_on --- (is_chain c) i hi j hj in ⟨⟨k, hkc⟩, hk⟩⟩ - --- private def chain_map (i j : c) (hij : i ≤ j) : i.1.carrier → j.1.carrier := --- inclusion (exists.elim hij (λ h _, h)) - --- instance chain_field_hom (i j : c) (hij : i ≤ j) : is_field_hom (chain_map c i j hij) := --- ring_hom_of_le hij - --- instance chain_directed_system : directed_system (λ i : c, i.1.carrier) (chain_map c) := --- begin --- split; intros; simp only [chain_map], --- { exact congr_fun (inclusion_refl _) x }, --- { exact (congr_fun (inclusion_trans _ _) x).symm } --- end - --- private def chain_limit : Type u := ring.direct_limit (λ i : c, i.1.carrier) (chain_map c) - --- private lemma of_eq_of (x : big_type K) (i j : c) (hi : x ∈ i.1.carrier) (hj : x ∈ j.1.carrier) : --- ring.direct_limit.of (λ i : c, i.1.carrier) (chain_map c) i ⟨x, hi⟩ = --- ring.direct_limit.of (λ i : c, i.1.carrier) (chain_map c) j ⟨x, hj⟩ := --- have hij : i ≤ j ∨ j ≤ i, --- from show i.1 ≤ j.1 ∨ j.1 ≤ i.1, from chain.total (is_chain c) i.2 j.2, --- hij.elim --- (λ hij, begin --- rw ← @ring.direct_limit.of_f c _ _ _ (λ i : c, i.1.carrier) _ _ (chain_map c) _ --- _ _ _ hij, --- simp [chain_map, inclusion] --- end) --- (λ hij, begin --- rw ← @ring.direct_limit.of_f c _ _ _ (λ i : c, i.1.carrier) _ _ (chain_map c) _ --- _ _ _ hij, --- simp [chain_map, inclusion] --- end) - --- private lemma injective_aux (i j : c) --- (x y : ⋃ i : c, i.1.carrier) (hx : x.1 ∈ i.1.carrier) (hy : y.1 ∈ j.1.carrier) : --- ring.direct_limit.of (λ i : c, i.1.carrier) (chain_map c) i ⟨x, hx⟩ = --- ring.direct_limit.of (λ i : c, i.1.carrier) (chain_map c) j ⟨y, hy⟩ → --- x = y := --- have hij : i ≤ j ∨ j ≤ i, --- from show i.1 ≤ j.1 ∨ j.1 ≤ i.1, from chain.total (is_chain c) i.2 j.2, --- have hinj : ∀ (i j : c) (hij : i ≤ j), injective (chain_map c i j hij), --- from λ _ _ _, is_field_hom.injective _, --- hij.elim --- (λ hij h, begin --- rw ← @ring.direct_limit.of_f c _ _ _ (λ i : c, i.1.carrier) _ _ (chain_map c) _ --- _ _ _ hij at h, --- simpa [chain_map, inclusion, subtype.coe_ext.symm] using ring.direct_limit.of_inj hinj j h, --- end) --- (λ hji h, begin --- rw ← @ring.direct_limit.of_f c _ _ _ (λ i : c, i.1.carrier) _ _ (chain_map c) _ --- _ _ _ hji at h, --- simpa [chain_map, inclusion, subtype.coe_ext.symm] using ring.direct_limit.of_inj hinj i h, --- end) - --- private def equiv_direct_limit : (⋃ (i : c), i.1.carrier) ≃ --- ring.direct_limit (λ i : c, i.1.carrier) (chain_map c) := --- @equiv.of_bijective (⋃ i : c, i.1.carrier) --- (ring.direct_limit (λ i : c, i.1.carrier) (chain_map c)) --- (λ x, ring.direct_limit.of _ _ (classical.some (set.mem_Union.1 x.2)) --- ⟨_, classical.some_spec (set.mem_Union.1 x.2)⟩) --- ⟨λ x y, injective_aux _ _ _ _ _ _ _, --- λ x, let ⟨i, ⟨y, hy⟩, hy'⟩ := ring.direct_limit.exists_of x in --- ⟨⟨y, _, ⟨i, rfl⟩, hy⟩, begin --- convert hy', --- exact of_eq_of _ _ _ _ _ _ --- end⟩⟩ - --- instance Union_field : discrete_field (⋃ i : c, i.1.carrier) := --- @equiv.discrete_field _ _ (equiv_direct_limit c) --- (field.direct_limit.discrete_field _ _) - --- set_option class.instance_max_depth 50 - --- instance is_field_hom_Union (i : c) : is_field_hom --- (inclusion (set.subset_Union (λ i : c, i.1.carrier) i)) := --- suffices inclusion (set.subset_Union (λ i : c, i.1.carrier) i) = --- ((equiv_direct_limit c).symm ∘ --- ring.direct_limit.of (λ i : c, i.1.carrier) (chain_map c) i), --- by rw this; exact is_ring_hom.comp _ _, --- funext $ λ ⟨_, _⟩, --- (equiv_direct_limit c).injective $ --- by rw [function.comp_app, equiv.apply_symm_apply]; --- exact of_eq_of _ _ _ _ _ _ - --- def Union_algebra (L : c) : algebra K (⋃ i : c, i.1.carrier) := --- algebra.of_ring_hom ((inclusion (set.subset_Union (λ i : c, i.1.carrier) L)) ∘ algebra_map _) --- (by { refine @is_ring_hom.comp _ _ _ _ _ _ _ _ _ _ }) - --- lemma Union_compat (L : c) (M : c) : --- (inclusion (set.subset_Union (λ i : c, i.1.carrier) M)) ∘ --- (algebra_map (M.val.carrier) : K → M.val.carrier) = --- by haveI := Union_algebra c L; exact algebra_map (↥⋃ (i : ↥c), (i.val).carrier) := --- begin --- rcases chain.directed_on (is_chain c) L.1 L.2 M.1 M.2 with ⟨N, hN, hLN, hMN⟩, --- rw show (inclusion (set.subset_Union (λ i : c, i.1.carrier) M)) = --- ((inclusion (set.subset_Union (λ i : c, i.1.carrier) ⟨N, hN⟩)) ∘ --- inclusion (subset_of_le hMN)), --- { funext x, refl }, --- rw comp.assoc, --- rw show inclusion (subset_of_le hMN) ∘ (algebra_map _ : K → (M.val).carrier) = --- inclusion (subset_of_le hLN) ∘ algebra_map _, --- { rw [compat, ← compat] }, --- rw ← comp.assoc, --- rw ← show (inclusion (set.subset_Union (λ i : c, i.1.carrier) L)) = --- ((inclusion (set.subset_Union (λ i : c, i.1.carrier) ⟨N, hN⟩)) ∘ --- inclusion (subset_of_le hLN)), --- { funext x, refl }, --- refl --- end - --- end chain - --- private def maximal_extension_chain (c : set (extension K)) (hc : chain (≤) c) : --- { ub : extension K // ∀ L, L ∈ c → L ≤ ub } := --- if h : nonempty c --- then --- let L := classical.some (classical.exists_true_of_nonempty h) in --- by letI : chain' c := ⟨hc⟩; exact --- ⟨{ carrier := ⋃ (i : c), i.1.carrier, --- algebra := Union_algebra c L, --- algebraic := --- begin --- rintro ⟨x, hx⟩, --- rw mem_Union at hx, --- cases hx with L' hx, --- rcases (L'.val).algebraic ⟨x, hx⟩ with ⟨p, pmonic, hp⟩, --- use [p, pmonic], --- rw aeval_def at hp ⊢, --- replace hp := congr_arg (inclusion (set.subset_Union (λ i : c, i.1.carrier) L')) hp, --- convert hp using 1; symmetry, --- { rw p.hom_eval₂ _ (inclusion _) _, --- congr' 1, --- { exact Union_compat c L L' }, --- { apply_instance, }, --- { apply is_ring_hom.is_semiring_hom, } }, --- { refine is_ring_hom.map_zero _ }, --- end }, --- λ e he, ⟨by convert subset_Union _ (⟨e, he⟩ : c); refl, --- algebraic_closure.is_field_hom_Union c ⟨e, he⟩, Union_compat c L ⟨e, he⟩⟩⟩ --- else ⟨base_extension K, λ a ha, (h ⟨⟨a, ha⟩⟩).elim⟩ - --- section adjoin_root --- variables {L : extension K} (f : polynomial L.carrier) [hif : irreducible f] --- include hif - --- open algebra - --- instance adjoin_root_algebraic_closure.field : --- discrete_field (adjoin_root f) := adjoin_root.field - --- instance adjoin_root_algebraic_closure.is_ring_hom : --- is_ring_hom (@adjoin_root.of _ _ _ f) := adjoin_root.is_ring_hom - --- private def adjoin_root.of_embedding : L.carrier ↪ adjoin_root f := --- ⟨adjoin_root.of, @is_field_hom.injective _ _ _ _ _ $ adjoin_root_algebraic_closure.is_ring_hom f⟩ - --- variable (K) - --- def aux_instance : algebra K (adjoin_root f) := --- algebra.of_ring_hom (adjoin_root.of ∘ algebra_map _) (is_ring_hom.comp _ _) - --- local attribute [instance] aux_instance - --- lemma adjoin_root.algebraic (x : adjoin_root f) : is_integral K x := --- is_integral_trans L.algebraic x $ adjoin_root.is_integral hif.ne_zero x - --- private def adjoin_root_extension_map : adjoin_root f ↪ big_type K := --- thing (adjoin_root.of_embedding f) ⟨subtype.val, subtype.val_injective⟩ --- (λ i, let e : big_type K ↪ ℕ × polynomial K := i.trans --- (algebraic_embedding (adjoin_root.algebraic K f)) in --- cantor_injective e.1 e.2) - --- private lemma adjoin_root_extension_map_apply (x : L.carrier) : --- (adjoin_root_extension_map K f) (@adjoin_root.of _ _ _ f x) = x.val := --- thing_commutes _ _ _ _ - --- instance range_adjoin_root_extension_map.discrete_field : --- discrete_field (set.range (@adjoin_root_extension_map K _ _ f _)) := --- equiv.discrete_field (equiv.set.range _ (embedding.inj _)).symm - --- private def adjoin_root_extension : extension K := --- { carrier := set.range (@adjoin_root_extension_map K _ _ f _), --- algebra := algebra.of_ring_hom --- ((equiv.set.range _ (embedding.inj' (adjoin_root_extension_map K f))).symm.symm ∘ --- algebra_map _) (is_ring_hom.comp _ _), --- algebraic := --- begin --- rintro ⟨_, x, rfl⟩, --- rcases adjoin_root.algebraic K f x with ⟨p, pmonic, hp⟩, --- use [p, pmonic], --- rw [aeval_def] at hp ⊢, --- replace hp := congr_arg --- ((equiv.set.range _ (embedding.inj' (adjoin_root_extension_map K f))).symm.symm) hp, --- convert hp using 1, --- symmetry, --- convert polynomial.hom_eval₂ _ _ _ _, --- all_goals {apply_instance} --- end } - --- variable {L} --- private lemma subset_adjoin_root_extension : L.carrier ⊆ (adjoin_root_extension K f).carrier := --- λ x h, ⟨adjoin_root.of_embedding f ⟨x, h⟩, thing_commutes _ _ _ _⟩ - --- private lemma adjoin_root_inclusion_eq : inclusion (subset_adjoin_root_extension K f) = --- ((equiv.set.range _ (adjoin_root_extension_map K f).2).symm.symm ∘ adjoin_root.of_embedding f) := --- funext $ λ ⟨_, _⟩, subtype.eq $ eq.symm $ adjoin_root_extension_map_apply _ _ _ - --- private lemma le_adjoin_root_extension : L ≤ adjoin_root_extension K f := --- ⟨subset_adjoin_root_extension K f, --- begin --- rw [adjoin_root_inclusion_eq]; dsimp [adjoin_root.of_embedding], --- exact ⟨is_ring_hom.comp _ _, rfl⟩ --- end⟩ - --- private def equiv_adjoin_root_of_le (h : adjoin_root_extension K f ≤ L) : --- L.carrier ≃+* adjoin_root f := --- have left_inv : left_inverse (inclusion h.fst ∘ (equiv.set.range _ --- (adjoin_root_extension_map K f).2)) adjoin_root.of, --- from λ _, by simp [adjoin_root_extension_map_apply, inclusion], --- have hom : is_ring_hom (coe : L.carrier → adjoin_root f), by apply_instance, --- { to_fun := coe, --- inv_fun := inclusion h.fst ∘ (equiv.set.range _ (adjoin_root_extension_map K f).2), --- left_inv := left_inv, --- right_inv := right_inverse_of_injective_of_left_inverse --- (injective_comp (inclusion_injective _) (equiv.injective _)) --- left_inv, --- map_add' := hom.map_add, --- map_mul' := hom.map_mul } - --- private def adjoin_root_equiv_adjoin_root_extension : --- adjoin_root f ≃+* (adjoin_root_extension K f).carrier := --- (equiv.set.range _ (adjoin_root_extension_map K f).2).symm.ring_equiv.symm - --- end adjoin_root - --- private lemma exists_algebraic_closure (K : Type u) [discrete_field K] : --- ∃ m : extension K, ∀ a, m ≤ a → a ≤ m := --- zorn (λ c hc, (maximal_extension_chain c hc).exists_of_subtype) (λ _ _ _, le_trans) - --- private def closed_extension (K : Type u) [discrete_field K] := --- classical.some (exists_algebraic_closure K) - --- def algebraic_closure (K : Type u) [discrete_field K] : Type u := --- ((classical.some (exists_algebraic_closure K))).carrier - --- end classical - --- section is_algebraically_closed --- /- In this section we prove the algebraic closure is algebraically closed -/ - --- local attribute [reducible] algebraic_closure - --- variables (f : polynomial (algebraic_closure K)) [hif : irreducible f] --- include hif - --- variable (K) - --- def algebraic_closure_equiv_adjoin_root : algebraic_closure K ≃+* adjoin_root f := --- equiv_adjoin_root_of_le K f $ --- classical.some_spec (exists_algebraic_closure K) _ (le_adjoin_root_extension _ _) - --- instance ring_equiv.is_semiring_hom {α β : Type*} [ring α] [ring β] (e : α ≃+* β) : --- is_semiring_hom (e : α → β) := --- is_ring_hom.is_semiring_hom _ - --- omit hif - --- private def is_algebraically_closed_aux : is_algebraically_closed (algebraic_closure K) := --- is_algebraically_closed_of_irreducible_has_root $ --- λ f hf, let e := by exactI algebraic_closure_equiv_adjoin_root K f in --- ⟨_, exists_root_of_equiv e (adjoin_root.eval₂_root f)⟩ - --- end is_algebraically_closed - --- /- To avoid diamonds, the `decidable_eq` instance is set to `classical.dec_eq`, --- as opposed to the (noncomputable, but not def-eq to `classical.dec_eq`) instance given by --- `(closed_extension K).field` -/ --- instance : discrete_field (algebraic_closure K) := --- { has_decidable_eq := classical.dec_eq _, --- ..(closed_extension K).field } - --- instance : algebra K (algebraic_closure K) := (closed_extension K).algebra - --- instance : is_algebraically_closed (algebraic_closure K) := --- by convert is_algebraically_closed_aux K - --- protected def is_integral : ∀ x : algebraic_closure K, is_integral K x := --- (closed_extension K).algebraic - --- attribute [irreducible] algebraic_closure closed_extension algebraic_closure.algebra - --- namespace lift --- /- In this section, the homomorphism from any algebraic extension into an algebraically --- closed extension is proven to exist. The assumption that M is algebraically closed could probably --- easily be switched to an assumption that M contains all the roots of polynomials in K -/ --- variables {L : Type v} {M : Type w} [discrete_field L] [algebra K L] --- [discrete_field M] [algebra K M] [is_algebraically_closed M] (hL : ∀ x : L, is_integral K x) - --- variables (K L M) --- include hL - --- /-- This structure is used to prove the existence of a homomorphism from any algebraic extension --- into an algebraic closure -/ --- private structure subfield_with_hom := --- (carrier : subalgebra K L) --- (emb : (@alg_hom K _ M _ _ _ (subalgebra.algebra carrier) _)) - --- variables {K L M hL} - --- namespace subfield_with_hom --- variables {E₁ E₂ E₃ : subfield_with_hom K L M hL} - --- instance : has_le (subfield_with_hom K L M hL) := --- { le := λ E₁ E₂, ∃ h : E₁.carrier ≤ E₂.carrier, E₂.emb ∘ inclusion h = E₁.emb } - --- lemma le_def : E₁ ≤ E₂ ↔ ∃ h : E₁.carrier ≤ E₂.carrier, E₂.emb ∘ inclusion h = E₁.emb := iff.rfl - --- lemma subalgebra_le_of_le (h : E₁ ≤ E₂) : E₁.carrier ≤ E₂.carrier := --- by { rw le_def at h, cases h, assumption } - --- lemma compat (h : E₁ ≤ E₂) : E₂.emb ∘ inclusion (subalgebra_le_of_le h) = E₁.emb := --- by { rw le_def at h, cases h, assumption } - --- instance : preorder (subfield_with_hom K L M hL) := --- { le := λ E₁ E₂, ∃ h : E₁.carrier ≤ E₂.carrier, E₂.emb ∘ inclusion h = E₁.emb, --- le_refl := λ E, ⟨le_refl _, by rw [inclusion_refl, comp.right_id]⟩, --- le_trans := λ E₁ E₂ E₃ h₁₂ h₂₃, ⟨le_trans (subalgebra_le_of_le h₁₂) (subalgebra_le_of_le h₂₃), --- begin --- erw inclusion_trans (subalgebra_le_of_le h₁₂) (subalgebra_le_of_le h₂₃), --- rw [← comp.assoc, compat, compat] --- end⟩ } - --- end subfield_with_hom - --- open lattice - --- def maximal_subfield_with_hom_chain (c : set (subfield_with_hom K L M hL)) (hc : chain (≤) c) : --- ∃ ub : subfield_with_hom K L M hL, ∀ N, N ∈ c → N ≤ ub := --- let ub : subfield_with_hom K L M hL := --- { carrier := Sup (subfield_with_hom.carrier '' c), --- emb := sorry } in --- ⟨ub, λ N hN, --- begin --- refine ⟨lattice.le_Sup ⟨N, hN, rfl⟩, _⟩, --- sorry --- end⟩ - --- variables (hL M) - --- lemma exists_maximal_subfield_with_hom : ∃ E : subfield_with_hom K L M hL, --- ∀ N, E ≤ N → N ≤ E := --- zorn (maximal_subfield_with_hom_chain) (λ _ _ _, le_trans) - --- def maximal_subfield_with_hom : subfield_with_hom K L M hL := --- classical.some (exists_maximal_subfield_with_hom M hL) - --- lemma maximal_subfield_with_hom_is_maximal : --- ∀ (N : subfield_with_hom K L M hL), --- (maximal_subfield_with_hom M hL) ≤ N → N ≤ (maximal_subfield_with_hom M hL) := --- classical.some_spec (exists_maximal_subfield_with_hom M hL) - --- lemma emb_injective (E : subfield_with_hom K L M hL) : --- injective E.emb := --- begin --- sorry --- -- let tmpa : _ := _, let tmpb : _ := _, --- -- refine @is_field_hom.injective _ M tmpa _ _ tmpb, --- -- swap, --- -- { sorry }, --- -- { exact { .. E.emb.hom } } --- end - --- lemma maximal_subfield_with_hom_surj : --- surjective (maximal_subfield_with_hom M hL).carrier.val := --- begin --- intros x, refine ⟨⟨x, _⟩, rfl⟩, --- replace hx := (maximal_subfield_with_hom M hL).carrier.is_integral x (hL x), --- let p := minimal_polynomial hx, --- have H := exists_root M (p.map (maximal_subfield_with_hom M hL).emb) _, --- swap, --- { calc 0 < degree p : --- begin --- sorry -- minimal_polynomial.degree_pos gives diamonds --- end --- ... = degree (p.map (maximal_subfield_with_hom M hL).emb) : --- begin --- symmetry, --- -- refine @polynomial.degree_map' _ _ _ _ _ _ p _ (by exact is_ring_hom.is_semiring_hom _) _, --- sorry, --- end }, --- let y := classical.some H, --- let f := algebra.adjoin_singleton_desc x hx --- (maximal_subfield_with_hom M hL).emb y (classical.some_spec H), --- let tmpa : subalgebra _ L := algebra.adjoin _ ({x} : set L), --- let tmpb : _ := _, --- let N : subfield_with_hom K L M hL := --- { carrier := subalgebra.under (maximal_subfield_with_hom M hL).carrier tmpa, --- emb := --- { to_fun := f, --- hom := algebra.adjoin_singleton_desc.is_ring_hom _ _ _ _ _, --- commutes' := tmpb } }, --- have hN : x ∈ N.carrier := algebra.subset_adjoin (mem_singleton x), --- refine subfield_with_hom.subalgebra_le_of_le (maximal_subfield_with_hom_is_maximal M hL N _) hN, --- { refine ⟨λ l hl, ring.subset_closure (mem_union_left _ _), _⟩, --- { rw mem_range, refine ⟨⟨l, hl⟩, rfl⟩ }, --- { sorry } }, --- { sorry } --- end - --- end lift - --- variables {L : Type v} {M : Type w} [discrete_field L] [algebra K L] --- [discrete_field M] [algebra K M] [is_algebraically_closed M] (hL : ∀ x : L, is_integral K x) - --- variables (K L M) --- include hL - --- /-- A (random) hom from an algebraic extension of K into an algebraic closure -/ --- def lift : L →ₐ[K] M := --- (lift.maximal_subfield_with_hom M hL).emb.comp $ alg_equiv.to_alg_hom $ alg_equiv.symm --- { ..(lift.maximal_subfield_with_hom M hL).carrier.val, --- ..equiv.of_bijective --- ⟨subtype.val_injective, lift.maximal_subfield_with_hom_surj _ _⟩ } +end is_alg_closed diff --git a/src/ring_theory/algebraic.lean b/src/ring_theory/algebraic.lean index 2e530e70ac41e..4c72da8e1182e 100644 --- a/src/ring_theory/algebraic.lean +++ b/src/ring_theory/algebraic.lean @@ -116,6 +116,11 @@ begin exact is_integral_trans L_alg A_alg, end +/-- If A is an algebraic algebra over K, then A is algebraic over L when L is an extension of K -/ +lemma is_algebraic_of_larger_field (A_alg : is_algebraic K A) : is_algebraic L A := +λ x, let ⟨p, hp⟩ := A_alg x in +⟨p.map (algebra_map _ _), map_ne_zero hp.1, by simp [hp.2]⟩ + /-- A field extension is algebraic if it is finite. -/ lemma is_algebraic_of_finite [finite : finite_dimensional K L] : is_algebraic K L := λ x, (is_algebraic_iff_is_integral _).mpr (is_integral_of_submodule_noetherian ⊤ From cc7699e7c590b910c0ec92d4ee35311f0f6b3c85 Mon Sep 17 00:00:00 2001 From: Chris Hughes Date: Thu, 9 Sep 2021 09:46:40 +0100 Subject: [PATCH 35/51] update module docstring --- src/field_theory/algebraic_closure.lean | 5 ++++- 1 file changed, 4 insertions(+), 1 deletion(-) diff --git a/src/field_theory/algebraic_closure.lean b/src/field_theory/algebraic_closure.lean index a7e015e8bc9bb..08681496fc7c5 100644 --- a/src/field_theory/algebraic_closure.lean +++ b/src/field_theory/algebraic_closure.lean @@ -26,9 +26,12 @@ polynomial in `k` splits. out by a maximal ideal containing every `f(x_f)`, and then repeating this step countably many times. See Exercise 1.13 in Atiyah--Macdonald. +- `is_alg_closed.lift` is a map from an algebraic extension `L` of `K`, into any algebraically + closed extension of `K`. + ## TODO -Show that any algebraic extension embeds into any algebraically closed extension (via Zorn's lemma). +Show that any two algebraic closures are isomorphic ## Tags From 7313b6ce246350ee31f97897a191888f27a4050b Mon Sep 17 00:00:00 2001 From: Chris Hughes Date: Thu, 9 Sep 2021 09:52:31 +0100 Subject: [PATCH 36/51] delete field_theory/algebraic --- src/field_theory/algebraic.lean | 248 -------------------------------- 1 file changed, 248 deletions(-) delete mode 100644 src/field_theory/algebraic.lean diff --git a/src/field_theory/algebraic.lean b/src/field_theory/algebraic.lean deleted file mode 100644 index 8f358e32512ad..0000000000000 --- a/src/field_theory/algebraic.lean +++ /dev/null @@ -1,248 +0,0 @@ -import ring_theory.adjoin_root -import ring_theory.integral_closure -import field_theory.minimal_polynomial - -universe variables u v w - -namespace polynomial -variables {α : Type*} {β : Type*} {γ : Type*} - -noncomputable theory -open_locale classical - -section coeffs -variables [comm_semiring α] (p : polynomial α) - -def nonzero_coeffs : finset α := p.support.image p.coeff - -lemma mem_nonzero_coeffs (i : ℕ) (hi : p.coeff i ≠ 0) : p.coeff i ∈ p.nonzero_coeffs := -finset.mem_image.mpr ⟨i, finsupp.mem_support_iff.mpr hi, rfl⟩ - -def coeffs : finset α := p.nonzero_coeffs ∪ {0} - -lemma mem_coeffs (i : ℕ) : p.coeff i ∈ p.coeffs := -if hi : p.coeff i = 0 -then finset.mem_union_right _ (by simp [hi]) -else finset.mem_union_left _ $ mem_nonzero_coeffs p i hi - -end coeffs - -section -variables [decidable_eq α] [comm_semiring α] - -lemma coeff_sum' (s : finset β) (f : β → polynomial α) (n : ℕ) : - coeff (s.sum f) n = s.sum (λ x, coeff (f x) n) := -(@finset.sum_hom _ _ _ s f _ _ (λ q, coeff q n) _).symm - -end - -section injective -open function -variables [decidable_eq α] [decidable_eq β] [comm_semiring α] [comm_semiring β] -variables {f : α → β} [is_semiring_hom f] - -lemma map_injective (hf : function.injective f) : - function.injective (map f : polynomial α → polynomial β) := -λ p q h, ext $ λ m, hf $ -begin - rw ext_iff at h, - specialize h m, - rw [coeff_map f, coeff_map f] at h, - exact h -end - -lemma leading_coeff_of_injective (hf : injective f) (p : polynomial α) : - leading_coeff (p.map f) = f (leading_coeff p) := -begin - delta leading_coeff, - rw [coeff_map f, nat_degree_map' p hf] -end - -lemma monic_of_injective (hf : injective f) {p : polynomial α} (hp : (p.map f).monic) : p.monic := -begin - apply hf, - rw [← leading_coeff_of_injective hf, hp.leading_coeff, is_semiring_hom.map_one f] -end - -end injective -section lift -open function -variables [decidable_eq α] [comm_ring α] (s : set α) [is_subring s] -variables (p : polynomial α) (hp : ∀ i, p.coeff i ∈ s) - -def lift : polynomial s := -(finset.range (p.nat_degree + 1)).sum (λ i, C ⟨(p.coeff i), hp i⟩ * X^i) - -lemma map_coe_lift : map coe (p.lift s hp) = p := -begin - conv_rhs {rw p.as_sum}, - rw ext_iff, intro n, - rw [coeff_map, lift, coeff_sum', coeff_sum', ← finset.sum_hom coe], - all_goals { try {apply_instance} }, - apply finset.sum_congr rfl, - intros i hi, - rw [coeff_C_mul, coeff_C_mul, is_ring_hom.map_mul coe, ← coeff_map coe, map_pow coe, map_X coe], - { refl }, - all_goals { apply_instance } -end - -end lift - -end polynomial - -namespace subalgebra -open polynomial -variables {R : Type*} {A : Type*} -variables [comm_ring R] [comm_ring A] [algebra R A] - -lemma zero_mem (S : subalgebra R A) : (0 : A) ∈ S := -submodule.zero_mem (S : submodule R A) - -variables [decidable_eq R] [decidable_eq A] - -protected lemma is_integral (S : subalgebra R A) (x : A) (hx : is_integral R x) : - is_integral S x := -begin - rcases hx with ⟨p, pmonic, hp⟩, - use p.map (algebra_map S), - split, - { exact monic_map _ pmonic }, - { rwa [aeval_def, eval₂_map] } -end - -end subalgebra - -section subfield -variables {K : Type u} [discrete_field K] -variables (s : set K) [is_subring s] - -def subfield_of_inv_mem (h : ∀ x ∈ s, x⁻¹ ∈ s) : discrete_field s := -{ inv := λ x, ⟨x⁻¹, h x.val x.property⟩, - zero_ne_one := λ h, zero_ne_one $ congr_arg subtype.val h, - mul_inv_cancel := λ x hx, subtype.val_injective $ mul_inv_cancel - $ λ hz, hx $ subtype.val_injective hz, - inv_mul_cancel := λ x hx, subtype.val_injective $ inv_mul_cancel - $ λ hz, hx $ subtype.val_injective hz, - inv_zero := subtype.val_injective $ inv_zero, - has_decidable_eq := subtype.decidable_eq, - .. @subtype.comm_ring K _ s _ } - -end subfield - -namespace subalgebra -open polynomial -variables {K : Type u} {L : Type v} [discrete_field K] [discrete_field L] [algebra K L] - -noncomputable def inv_poly (x : L) (hx : is_integral K x) : polynomial K := -let p := minimal_polynomial hx in -((p - C (p.coeff 0)) /ₘ X) * C (p.coeff 0)⁻¹ - -variables (L_alg : ∀ x:L, is_integral K x) -include L_alg - -instance (E : subalgebra K L) : discrete_field (subtype E.carrier) := -subfield_of_inv_mem _ $ λ x (hx : x ∈ E), sorry - -end subalgebra - -namespace algebra -open set polynomial -variables {R : Type*} {A : Type*} {B : Type*} -variables [decidable_eq R] [decidable_eq A] [decidable_eq B] -variables [comm_ring R] [comm_ring A] [algebra R A] [comm_ring B] - -def adjoin_singleton_desc (x : A) (hx : is_integral R x) - (f : R → B) [is_ring_hom f] (y : B) (hy : is_root ((minimal_polynomial hx).map f) y) : -(adjoin R ({x} : set A) : Type _) → B := -sorry - -instance adjoin_singleton_desc.is_ring_hom (x : A) (hx : is_integral R x) - (f : R → B) [is_ring_hom f] (y : B) (hy : is_root ((minimal_polynomial hx).map f) y) : - is_ring_hom (adjoin_singleton_desc x hx f y hy) := -sorry - -end algebra - --- namespace subalgebra --- open set lattice --- variables {R : Type*} {A : Type*} {B : Type*} --- variables [comm_ring R] [comm_ring A] [algebra R A] [comm_ring B] [algebra R B] - --- def Sup.desc (Ss : set (subalgebra R A)) (f : Π S : Ss, (S : subalgebra R A) →ₐ[R] B) --- (hf : ∀ S₁ S₂ : Ss, ∃ h : (S₁ : subalgebra R A) ≤ S₂, (f S₂) ∘ inclusion h = f S₁) : --- (Sup Ss : subalgebra R A) →ₐ[R] B := --- sorry - --- end subalgebra - -open function algebra polynomial -variables {R : Type*} {A : Type*} {B : Type*} -variables [decidable_eq R] [decidable_eq A] [decidable_eq B] -variables [comm_ring R] [comm_ring A] [comm_ring B] -variables [algebra R A] [algebra A B] -variables (A_alg : ∀ x : A, is_integral R x) - -include A_alg - -set_option class.instance_max_depth 50 - -lemma is_integral_trans_aux (x : B) {p : polynomial A} (pmonic : monic p) (hp : aeval A B x p = 0) - (S : set (comap R A B)) - (hS : S = (↑((finset.range (p.nat_degree + 1)).image - (λ i, to_comap R A B (p.coeff i))) : set (comap R A B))) : - is_integral (adjoin R S) (comap.to_comap R A B x) := -begin - have coeffs_mem : ∀ i, coeff (map (to_comap R A B) p) i ∈ adjoin R S, - { intro i, - by_cases hi : i ∈ finset.range (p.nat_degree + 1), - { apply algebra.subset_adjoin, subst S, - rw [finset.mem_coe, finset.mem_image, coeff_map], - exact ⟨i, hi, rfl⟩ }, - { rw [finset.mem_range, not_lt] at hi, - rw [coeff_map, coeff_eq_zero_of_nat_degree_lt hi, alg_hom.map_zero], - exact subalgebra.zero_mem _ } }, - let q : polynomial (adjoin R S) := polynomial.lift _ (p.map $ to_comap R A B) coeffs_mem, - have hq : (q.map (algebra_map (comap R A B))) = (p.map $ to_comap R A B) := - map_coe_lift _ (p.map $ to_comap R A B) coeffs_mem, - use q, - split, - { suffices h : (q.map (algebra_map (comap R A B))).monic, - { refine @monic_of_injective _ _ _ _ _ _ _ - (by exact is_ring_hom.is_semiring_hom _) _ _ h, - exact subtype.val_injective }, - { erw map_coe_lift, exact monic_map _ pmonic } }, - { convert hp using 1, - replace hq := congr_arg (eval (comap.to_comap R A B x)) hq, - convert hq using 1; symmetry, swap, - exact eval_map _ _, - refine @eval_map _ _ _ _ _ _ (by exact is_ring_hom.is_semiring_hom _) _ }, -end - -lemma is_integral_trans (x : B) (hx : is_integral A x) : - is_integral R (comap.to_comap R A B x) := -begin - rcases hx with ⟨p, pmonic, hp⟩, - let S : set (comap R A B) := - (↑((finset.range (p.nat_degree + 1)).image - (λ i, to_comap R A B (p.coeff i))) : set (comap R A B)), - refine is_integral_of_mem_of_fg (adjoin R (S ∪ {comap.to_comap R A B x})) _ _ _, - swap, { apply subset_adjoin, simp }, - apply fg_trans, - { apply fg_adjoin_of_finite, { apply finset.finite_to_set }, - intros x hx, - rw [finset.mem_coe, finset.mem_image] at hx, - rcases hx with ⟨i, hi, rfl⟩, - rcases A_alg (p.coeff i) with ⟨q, qmonic, hq⟩, - use [q, qmonic], - replace hq := congr_arg (to_comap R A B : A → (comap R A B)) hq, - rw alg_hom.map_zero at hq, - convert hq using 1, - symmetry, exact polynomial.hom_eval₂ _ _ _ _ }, - { apply fg_adjoin_singleton_of_integral, - exact is_integral_trans_aux A_alg _ pmonic hp _ rfl } -end -. - -lemma algebraic_trans (B_alg : ∀ x : B, is_integral A x) : - ∀ x : comap R A B, is_integral R x := -λ x, is_integral_trans A_alg x (B_alg x) From 2c34044a83c5b0bba7681a1af1b7bbc7f6e5e408 Mon Sep 17 00:00:00 2001 From: Chris Hughes Date: Thu, 9 Sep 2021 09:53:19 +0100 Subject: [PATCH 37/51] delete nonsense comments --- src/field_theory/algebraic_closure.lean | 15 --------------- 1 file changed, 15 deletions(-) diff --git a/src/field_theory/algebraic_closure.lean b/src/field_theory/algebraic_closure.lean index 08681496fc7c5..91c23c93d5b1e 100644 --- a/src/field_theory/algebraic_closure.lean +++ b/src/field_theory/algebraic_closure.lean @@ -43,21 +43,6 @@ universes u v w noncomputable theory open_locale classical big_operators open polynomial --- #print ideal.quotient.lift_mk --- open ideal - --- @[simp] lemma ideal.quotient.lift_comp_mk {R : Type*} [comm_ring R] --- {S : Type v} [comm_ring S] {I : ideal R} (f : R →+* S) (H : I ≤ f.ker) : --- (ideal.quotient.lift I f H).comp I^.quotient.mk = f := --- by ext; simp - --- lemma ker_quotient_lift {R : Type} [comm_ring R] --- {S : Type v} [comm_ring S] {I : ideal R} (f : R →+* S) (H : I ≤ f.ker) : --- (ideal.quotient.lift I f H).ker = (f.ker).map I^.quotient.mk := --- by rwa [← (comap_injective_of_surjective _ quotient.mk_surjective).eq_iff, --- ring_hom.ker_eq_comap_bot, comap_comap, ← ring_hom.ker_eq_comap_bot, --- comap_map_of_surjective _ quotient.mk_surjective, --- ideal.quotient.lift_comp_mk, ← ring_hom.ker_eq_comap_bot, mk_ker, left_eq_sup] variables (k : Type u) [field k] From 1b1453deef1ae54156de04e2cf7e2d7ccc12cd03 Mon Sep 17 00:00:00 2001 From: Chris Hughes <33847686+ChrisHughes24@users.noreply.github.com> Date: Thu, 9 Sep 2021 12:07:34 +0100 Subject: [PATCH 38/51] Update src/algebra/algebra/subalgebra.lean Co-authored-by: Eric Wieser --- src/algebra/algebra/subalgebra.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/src/algebra/algebra/subalgebra.lean b/src/algebra/algebra/subalgebra.lean index 188da6289ccb2..4403f148589ab 100644 --- a/src/algebra/algebra/subalgebra.lean +++ b/src/algebra/algebra/subalgebra.lean @@ -754,7 +754,7 @@ let K : subalgebra R A := let ⟨j, hj⟩ := set.mem_Union.1 hy in let ⟨k, hik, hjk⟩ := dir i j in set.mem_Union.2 ⟨k, subalgebra.add_mem (S k) (hik hi) (hjk hj)⟩, - algebra_map_mem' := λ r, let i := @nonempty.some ι infer_instance in + algebra_map_mem' := λ r, let ⟨i⟩ := ‹nonempty ι› in set.mem_Union.2 ⟨i, subalgebra.algebra_map_mem _ _⟩ } in have supr S = K, from le_antisymm (supr_le (λ i, set.subset_Union (λ i, ↑(S i)) i)) From c9342f9d91e637f2b72cedc8fa4652b38e9cf0f8 Mon Sep 17 00:00:00 2001 From: Chris Hughes Date: Thu, 9 Sep 2021 12:15:30 +0100 Subject: [PATCH 39/51] minor changes --- src/algebra/algebra/subalgebra.lean | 5 ++--- 1 file changed, 2 insertions(+), 3 deletions(-) diff --git a/src/algebra/algebra/subalgebra.lean b/src/algebra/algebra/subalgebra.lean index 1d5ab47eb7053..69ab8bd158149 100644 --- a/src/algebra/algebra/subalgebra.lean +++ b/src/algebra/algebra/subalgebra.lean @@ -736,8 +736,7 @@ section supr_lift variables {ι : Type*} lemma coe_supr_of_directed [nonempty ι] {S : ι → subalgebra R A} - (dir : directed (≤) S) : - ((supr S : subalgebra R A) : set A) = ⋃ i, (S i : set A) := + (dir : directed (≤) S) : ↑(supr S) = ⋃ i, (S i : set A) := let K : subalgebra R A := { carrier := ⋃ i, (S i), mul_mem' := λ x y hx hy, @@ -765,7 +764,7 @@ noncomputable def supr_lift [nonempty ι] (dir : directed (≤) K) (f : Π i, K i →ₐ[R] B) (hf : ∀ (i j : ι) (h : K i ≤ K j), f i = (f j).comp (inclusion h)) : - (supr K : subalgebra R A) →ₐ[R] B := + ↥(supr K) →ₐ[R] B := { to_fun := set.lift_of_eq_Union (λ i, ↑(K i)) (λ i x, f i x) (λ i j x hxi hxj, let ⟨k, hik, hjk⟩ := dir i j in From 21ed26b56d8094806b394854b6c3edc65a5f727e Mon Sep 17 00:00:00 2001 From: Chris Hughes Date: Thu, 9 Sep 2021 13:52:00 +0100 Subject: [PATCH 40/51] mention glue in module docstring --- src/data/set/Union_lift.lean | 6 +++--- 1 file changed, 3 insertions(+), 3 deletions(-) diff --git a/src/data/set/Union_lift.lean b/src/data/set/Union_lift.lean index 053e918ccd57d..dc0f4ffeda899 100644 --- a/src/data/set/Union_lift.lean +++ b/src/data/set/Union_lift.lean @@ -7,8 +7,8 @@ import data.set.lattice import order.directed /-! # Union lift -This file defines `set.Union_lift` to lift function defined on each of a collection of sets -to the Union of those sets. +This file defines `set.Union_lift` to glue together functions defined on each of a collection of +sets to make a function on the Union of those sets. ## Main definitions @@ -37,7 +37,7 @@ or binary functions are preserved. These lemmas are: ## Tags -directed union, directed supremum +directed union, directed supremum, glue, gluing -/ variables {α ι β : Type*} From 773938b8eb1f9c936ec2e23478817edb016addbe Mon Sep 17 00:00:00 2001 From: Chris Hughes Date: Fri, 10 Sep 2021 21:21:56 +0100 Subject: [PATCH 41/51] start on minor changes --- src/data/set/Union_lift.lean | 55 +++++++++++++++--------------------- 1 file changed, 23 insertions(+), 32 deletions(-) diff --git a/src/data/set/Union_lift.lean b/src/data/set/Union_lift.lean index dc0f4ffeda899..f19da915d7f10 100644 --- a/src/data/set/Union_lift.lean +++ b/src/data/set/Union_lift.lean @@ -53,32 +53,29 @@ it on each component, and proving that it agrees on the intersections. -/ noncomputable def Union_lift (S : ι → set α) (f : Π i (x : S i), β) (hf : ∀ i j (x : α) (hxi : x ∈ S i) (hxj : x ∈ S j), f i ⟨x, hxi⟩ = f j ⟨x, hxj⟩) - (x : Union S) : β := -let i := classical.indefinite_description _ (mem_Union.1 x.2) in + (T : set α) (hT : T ⊆ Union S) (x : T) : β := +let i := classical.indefinite_description _ (mem_Union.1 (hT x.prop)) in f i ⟨x, i.prop⟩ -@[simp] lemma Union_lift_inclusion {S : ι → set α} +variables + {S : ι → set α} {f : Π i (x : S i), β} - {hf : ∀ i j (x : α) (hxi : x ∈ S i) (hxj : x ∈ S j), f i ⟨x, hxi⟩ = f j ⟨x, hxj⟩ } - {i : ι} (x : S i) (h : S i ⊆ Union S := set.subset_Union S i) : - Union_lift S f hf (set.inclusion h x) = f i x := -let j := classical.indefinite_description _ (mem_Union.1 (h x.2)) in -by cases x with x hx; exact hf j i x j.2 hx + {hf : ∀ i j (x : α) (hxi : x ∈ S i) (hxj : x ∈ S j), f i ⟨x, hxi⟩ = f j ⟨x, hxj⟩} + {T : set α} {hT : T ⊆ Union S} @[simp] lemma Union_lift_mk - {S : ι → set α} - {f : Π i (x : S i), β} - {hf : ∀ i j (x : α) (hxi : x ∈ S i) (hxj : x ∈ S j), f i ⟨x, hxi⟩ = f j ⟨x, hxj⟩ } - {i : ι} (x : S i) (hx : (x : α) ∈ Union S := set.subset_Union S i x.2) : - Union_lift S f hf ⟨x, hx⟩ = f i x := -Union_lift_inclusion x + {i : ι} (x : S i) (hx : (x : α) ∈ T) : + Union_lift S f hf T hT ⟨x, hx⟩ = f i x := +let j := classical.indefinite_description _ (mem_Union.1 (hT hx)) in +by cases x with x hx; exact hf j i x j.2 _ + +@[simp] lemma Union_lift_inclusion {i : ι} (x : S i) + (h : S i ⊆ T) : Union_lift S f hf T hT (set.inclusion h x) = f i x := +Union_lift_mk x _ lemma Union_lift_of_mem - {S : ι → set α} - {f : Π i (x : S i), β} - {hf : ∀ i j (x : α) (hxi : x ∈ S i) (hxj : x ∈ S j), f i ⟨x, hxi⟩ = f j ⟨x, hxj⟩ } - (x : Union S) {i : ι} (hx : (x : α) ∈ S i) : - Union_lift S f hf x = f i ⟨x, hx⟩ := + (x : T) {i : ι} (hx : (x : α) ∈ S i) : + Union_lift S f hf T hT x = f i ⟨x, hx⟩ := by cases x with x hx; exact hf _ _ _ _ _ /-- `Union_lift_const` is useful for proving that `Union_lift` is a homomorphism @@ -87,16 +84,13 @@ by cases x with x hx; exact hf _ _ _ _ _ of group homomorphisms on a union of subgroups preserves `1`. See also `lift_of_eq_Union_const` -/ lemma Union_lift_const - {S : ι → set α} - {f : Π i (x : S i), β} - {hf : ∀ i j (x : α) (hxi : x ∈ S i) (hxj : x ∈ S j), f i ⟨x, hxi⟩ = f j ⟨x, hxj⟩} - (c : Union S) + (c : T) (ci : Π i, S i) (hci : ∀ i, (ci i : α) = c) (cβ : β) (h : ∀ i, f i (ci i) = cβ) : - Union_lift S f hf c = cβ := -let ⟨i, hi⟩ := set.mem_Union.1 c.prop in + Union_lift S f hf T hT c = cβ := +let ⟨i, hi⟩ := set.mem_Union.1 (hT c.prop) in have (ci i) = ⟨c, hi⟩, from subtype.ext (hci i), by rw [Union_lift_of_mem _ hi, ← this, h] @@ -106,17 +100,14 @@ by rw [Union_lift_of_mem _ hi, ← this, h] of linear_maps on a union of submodules preserves scalar multiplication. See also `lift_of_eq_Union_unary` -/ lemma Union_lift_unary - {S : ι → set α} - {f : Π i (x : S i), β} - {hf : ∀ i j (x : α) (hxi : x ∈ S i) (hxj : x ∈ S j), f i ⟨x, hxi⟩ = f j ⟨x, hxj⟩} - (u : (Union S) → (Union S)) + (u : T → T) (ui : Π i, S i → S i) - (hui : ∀ i x, set.inclusion (set.subset_Union S i) (ui i x) = - u (set.inclusion (set.subset_Union S i) x)) + (hui : ∀ i x (hx : ↑x ∈ T), set.inclusion (set.subset_Union S i) (ui i x) = + u ⟨x, hx⟩) (uβ : β → β) (h : ∀ i (x : S i), (f i (ui i x)) = uβ (f i x)) (x : Union S) : - Union_lift S f hf (u x) = uβ (Union_lift S f hf x) := + Union_lift S f hf T hT (u x) = uβ (Union_lift S f hf T hT x) := begin cases set.mem_Union.1 x.prop with i hi, rw [Union_lift_of_mem x hi, ← h i], From a6d4fd362260b63566e75598340f59b3d67a070c Mon Sep 17 00:00:00 2001 From: Chris Hughes Date: Sun, 12 Sep 2021 16:11:25 +0100 Subject: [PATCH 42/51] refactor Union_lift file --- src/data/set/Union_lift.lean | 185 +++++++++-------------------------- 1 file changed, 44 insertions(+), 141 deletions(-) diff --git a/src/data/set/Union_lift.lean b/src/data/set/Union_lift.lean index f19da915d7f10..c1bb0dea79d2e 100644 --- a/src/data/set/Union_lift.lean +++ b/src/data/set/Union_lift.lean @@ -12,28 +12,23 @@ sets to make a function on the Union of those sets. ## Main definitions -* `set.Union_lift` - Given a Union of sets `Union S`, define a function on the Union by defining -it on each component, and proving that it agrees on the intersections. -* `set.lift_of_eq_Union` Version of `Union_lift` for a set that is propositionally equal to `Union S`. - This is mainly motivated by the need to define functions on directed supremums of e.g. subgroups, - where the supremum is propositionally, but not definitionally equal to the Union. +* `set.Union_lift` - Given a Union of sets `Union S`, define a function on any subset of the Union + by defining it on each component, and proving that it agrees on the intersections. +* `set.lift_cover` - Version of `set.Union_lift` for the special case that the sets cover the + entire type. ## Main statements There are proofs of the obvious properties of `Union_lift`, i.e. what it does to elements of each of the sets in the `Union`, stated in different ways. -There are also three lemmas about `Union_lift` and three lemmas about `lift_of_eq_Union` intended -to aid with proving that `Union_lift` of `lift_of_eq_Union` is a homomorphism when defined -on a Union of substructures. There is one lemma each to show that constants, unary functions, -or binary functions are preserved. These lemmas are: +There are also three lemmas about `Union_lift` intended to aid with proving that `Union_lift` is a +homomorphism when defined on a Union of substructures. There is one lemma each to show that +constants, unary functions, or binary functions are preserved. These lemmas are: *`set.Union_lift_const` *`set.Union_lift_unary` *`set.Union_lift_binary` -*`set.lift_of_eq_Union_const` -*`set.lift_of_eq_Union_unary` -*`set.lift_of_eq_Union_binary` ## Tags @@ -44,6 +39,8 @@ variables {α ι β : Type*} namespace set +section Union_lift + /- The unused argument `hf` is left in the definition so that the `simp` lemmas `Union_lift_inclusion` will work without the user having to provide `hf` explicitly to simplify terms involving `Union_lift`. -/ @@ -61,7 +58,7 @@ variables {S : ι → set α} {f : Π i (x : S i), β} {hf : ∀ i j (x : α) (hxi : x ∈ S i) (hxj : x ∈ S j), f i ⟨x, hxi⟩ = f j ⟨x, hxj⟩} - {T : set α} {hT : T ⊆ Union S} + {T : set α} {hT : T ⊆ Union S} {hT' : T = Union S} @[simp] lemma Union_lift_mk {i : ι} (x : S i) (hx : (x : α) ∈ T) : @@ -81,8 +78,7 @@ by cases x with x hx; exact hf _ _ _ _ _ /-- `Union_lift_const` is useful for proving that `Union_lift` is a homomorphism of algebraic structures when defined on the Union of algebraic subobjects. For example, it could be used to prove that the lift of a collection - of group homomorphisms on a union of subgroups preserves `1`. See also - `lift_of_eq_Union_const` -/ + of group homomorphisms on a union of subgroups preserves `1`. -/ lemma Union_lift_const (c : T) (ci : Π i, S i) @@ -97,47 +93,44 @@ by rw [Union_lift_of_mem _ hi, ← this, h] /-- `Union_lift_unary` is useful for proving that `Union_lift` is a homomorphism of algebraic structures when defined on the Union of algebraic subobjects. For example, it could be used to prove that the lift of a collection - of linear_maps on a union of submodules preserves scalar multiplication. See also - `lift_of_eq_Union_unary` -/ + of linear_maps on a union of submodules preserves scalar multiplication. -/ lemma Union_lift_unary (u : T → T) (ui : Π i, S i → S i) - (hui : ∀ i x (hx : ↑x ∈ T), set.inclusion (set.subset_Union S i) (ui i x) = - u ⟨x, hx⟩) + (hui : ∀ i (x : S i), u (set.inclusion (show S i ⊆ T, from hT'.symm ▸ set.subset_Union S i) x) + = set.inclusion (show S i ⊆ T, from hT'.symm ▸ set.subset_Union S i) (ui i x)) (uβ : β → β) (h : ∀ i (x : S i), (f i (ui i x)) = uβ (f i x)) - (x : Union S) : - Union_lift S f hf T hT (u x) = uβ (Union_lift S f hf T hT x) := + (x : T) : + Union_lift S f hf T (le_of_eq hT') (u x) = uβ (Union_lift S f hf T (le_of_eq hT') x) := begin + subst hT', cases set.mem_Union.1 x.prop with i hi, rw [Union_lift_of_mem x hi, ← h i], have : x = (set.inclusion (set.subset_Union S i) ⟨x, hi⟩), { cases x, refl }, have hx' : (set.inclusion (set.subset_Union S i) (ui i ⟨x, hi⟩) : α) ∈ S i, - from (ui i ⟨x, hi⟩).prop, - conv_lhs { rw [this, ← hui, Union_lift_of_mem _ hx'] }, - simp only [coe_inclusion, subtype.coe_eta] + from (ui i ⟨x, hi⟩).prop, + conv_lhs { rw [this, hui, Union_lift_inclusion] } end /-- `Union_lift_binary` is useful for proving that `Union_lift` is a homomorphism of algebraic structures when defined on the Union of algebraic subobjects. For example, it could be used to prove that the lift of a collection - of group homomorphisms on a union of subgroups preserves `*`. See also - `lift_of_eq_Union_binary` -/ + of group homomorphisms on a union of subgroups preserves `*`. -/ lemma Union_lift_binary - {S : ι → set α} - {f : Π i (x : S i), β} - {hf : ∀ i j (x : α) (hxi : x ∈ S i) (hxj : x ∈ S j), f i ⟨x, hxi⟩ = f j ⟨x, hxj⟩} (dir: directed (≤) S) - (op : (Union S) → (Union S) → (Union S)) + (op : T → T → T) (opi : Π i, S i → S i → S i) - (hopi : ∀ i x y, set.inclusion (set.subset_Union S i) (opi i x y) = - op (set.inclusion (set.subset_Union S i) x) - (set.inclusion (set.subset_Union S i) y)) + (hopi : ∀ i x y, set.inclusion (show S i ⊆ T, from hT'.symm ▸ set.subset_Union S i) (opi i x y) = + op (set.inclusion (show S i ⊆ T, from hT'.symm ▸ set.subset_Union S i) x) + (set.inclusion (show S i ⊆ T, from hT'.symm ▸ set.subset_Union S i) y)) (opβ : β → β → β) (h : ∀ i (x y : S i), (f i (opi i x y)) = opβ (f i x) (f i y)) - (x y : Union S) : - Union_lift S f hf (op x y) = opβ (Union_lift S f hf x) (Union_lift S f hf y) := + (x y : T) : + Union_lift S f hf T (le_of_eq hT') (op x y) = + opβ (Union_lift S f hf T (le_of_eq hT') x) (Union_lift S f hf T (le_of_eq hT') y) := begin + subst hT', cases set.mem_Union.1 x.prop with i hi, cases set.mem_Union.1 y.prop with j hj, rcases dir i j with ⟨k, hik, hjk⟩, @@ -150,118 +143,28 @@ begin simp only [coe_inclusion, subtype.coe_eta] end -/-- Version of `Union_lift` for sets that are propositionally equal to the Union. - Define a map on a set equal to the `Union` of sets by defining it on each of the sets - in the `Union`. -/ -noncomputable def lift_of_eq_Union - (S : ι → set α) - (f : Π i (x : S i), β) - (hf : ∀ i j (x : α) (hxi : x ∈ S i) (hxj : x ∈ S j), f i ⟨x, hxi⟩ = f j ⟨x, hxj⟩) - (T : set α) (hT : T = ⋃ i, S i) - (x : T) : β := -by subst hT; exact Union_lift S f hf x - -@[simp] lemma lift_of_eq_Union_inclusion {S : ι → set α} - {f : Π i (x : S i), β} - {hf : ∀ i j (x : α) (hxi : x ∈ S i) (hxj : x ∈ S j), f i ⟨x, hxi⟩ = f j ⟨x, hxj⟩ } - {T : set α} {hT : T = Union S} - {i : ι} (x : S i) (h : S i ⊆ T := hT.symm ▸ set.subset_Union S i) : - lift_of_eq_Union S f hf T hT (set.inclusion h x) = f i x := -begin - subst hT, - delta lift_of_eq_Union, - exact Union_lift_inclusion x h -end - -@[simp] lemma lift_of_eq_Union_mk - {S : ι → set α} - {f : Π i (x : S i), β} - {hf : ∀ i j (x : α) (hxi : x ∈ S i) (hxj : x ∈ S j), f i ⟨x, hxi⟩ = f j ⟨x, hxj⟩ } - {T : set α} {hT : T = Union S} - {i : ι} (x : S i) - (hx : (x : α) ∈ T := hT.symm ▸ set.subset_Union S i x.prop) : - lift_of_eq_Union S f hf T hT ⟨x, hx⟩ = f i x := -begin - subst hT, - delta lift_of_eq_Union, - exact Union_lift_mk x hx -end +end Union_lift -lemma lift_of_eq_Union_of_mem - {S : ι → set α} - {f : Π i (x : S i), β} - {hf : ∀ i j (x : α) (hxi : x ∈ S i) (hxj : x ∈ S j), f i ⟨x, hxi⟩ = f j ⟨x, hxj⟩ } - {T : set α} {hT : T = Union S} - (x : T) {i : ι} (hx : (x : α) ∈ S i) : - lift_of_eq_Union S f hf T hT x = f i ⟨x, hx⟩ := -begin - subst hT, - delta lift_of_eq_Union, - exact Union_lift_of_mem x hx -end - -/-- Version of `Union_lift_const` for `lift_of_eq_Union`. -/ -lemma lift_of_eq_Union_const +variables {S : ι → set α} {f : Π i (x : S i), β} {hf : ∀ i j (x : α) (hxi : x ∈ S i) (hxj : x ∈ S j), f i ⟨x, hxi⟩ = f j ⟨x, hxj⟩} - {T : set α} {hT : T = Union S} - (c : T) - (ci : Π i, S i) - (hci : ∀ i, (ci i : α) = c) - (cβ : β) - (h : ∀ i, f i (ci i) = cβ) : - lift_of_eq_Union S f hf T hT c = cβ := -begin - subst T, - delta lift_of_eq_Union, - exact Union_lift_const c ci hci cβ h -end + {hS : Union S = univ} -/-- Version of `Union_lift_unary` for `lift_of_eq_Union`. -/ -lemma lift_of_eq_Union_unary - {S : ι → set α} - {f : Π i (x : S i), β} - {hf : ∀ i j (x : α) (hxi : x ∈ S i) (hxj : x ∈ S j), f i ⟨x, hxi⟩ = f j ⟨x, hxj⟩} - {T : set α} {hT : T = Union S} - (u : T → T) - (ui : Π i, S i → S i) - (hui : ∀ i x, set.inclusion (show S i ⊆ T, from hT.symm ▸ set.subset_Union S i) (ui i x) = - u (set.inclusion (show S i ⊆ T, from hT.symm ▸ set.subset_Union S i) x)) - (uβ : β → β) - (h : ∀ i (x : S i), (f i (ui i x)) = uβ (f i x)) - (x : T) : - lift_of_eq_Union S f hf T hT (u x) = uβ (lift_of_eq_Union S f hf T hT x) := -begin - subst T, - delta lift_of_eq_Union, - exact Union_lift_unary u ui hui uβ h x -end +noncomputable def lift_cover + (S : ι → set α) + (f : Π i (x : S i), β) + (hf : ∀ i j (x : α) (hxi : x ∈ S i) (hxj : x ∈ S j), f i ⟨x, hxi⟩ = f j ⟨x, hxj⟩) + (hS : Union S = univ) (a : α) : β := +Union_lift S f hf univ (hS ▸ set.subset.refl _) ⟨a, trivial⟩ -/-- Version of `Union_lift_binary` for `lift_of_eq_Union`. -/ -lemma lift_of_eq_Union_binary - {S : ι → set α} - {f : Π i (x : S i), β} - {hf : ∀ i j (x : α) (hxi : x ∈ S i) (hxj : x ∈ S j), f i ⟨x, hxi⟩ = f j ⟨x, hxj⟩} - {T : set α} {hT : T = ⋃ i, S i} - (dir: directed (≤) S) - (op : T → T → T) - (opi : Π i, S i → S i → S i) - (hopi : ∀ i x y, set.inclusion (show S i ⊆ T, from hT.symm ▸ set.subset_Union S i) (opi i x y) = - op (set.inclusion (show S i ⊆ T, from hT.symm ▸ set.subset_Union S i) x) - (set.inclusion (show S i ⊆ T, from hT.symm ▸ set.subset_Union S i) y)) - (opβ : β → β → β) - (h : ∀ i (x y : S i), (f i (opi i x y)) = opβ (f i x) (f i y)) - (x y : T) : - lift_of_eq_Union S f hf T hT (op x y) = - opβ (lift_of_eq_Union S f hf T hT x) - (lift_of_eq_Union S f hf T hT y) := -begin - subst T, - delta lift_of_eq_Union, - exact Union_lift_binary dir op opi hopi opβ h x y -end +@[simp] lemma lift_cover_coe {i : ι} (x : S i) : lift_cover S f hf hS x = f i x := +Union_lift_mk x _ + +lemma lift_cover_of_mem {i : ι} {x : α} (hx : (x : α) ∈ S i) : + lift_cover S f hf hS x = f i ⟨x, hx⟩ := +Union_lift_of_mem ⟨x, trivial⟩ hx -attribute [irreducible] lift_of_eq_Union Union_lift +attribute [irreducible] Union_lift lift_cover end set From e1c9aefef8ae9b5da1a23bde28cf2542a13c7e3d Mon Sep 17 00:00:00 2001 From: Chris Hughes Date: Sun, 12 Sep 2021 16:21:41 +0100 Subject: [PATCH 43/51] hopefully fix supr_lift --- src/algebra/algebra/subalgebra.lean | 14 +++++++------- 1 file changed, 7 insertions(+), 7 deletions(-) diff --git a/src/algebra/algebra/subalgebra.lean b/src/algebra/algebra/subalgebra.lean index e48b092d1d54d..43145f262ceb7 100644 --- a/src/algebra/algebra/subalgebra.lean +++ b/src/algebra/algebra/subalgebra.lean @@ -769,19 +769,19 @@ noncomputable def supr_lift [nonempty ι] (f : Π i, K i →ₐ[R] B) (hf : ∀ (i j : ι) (h : K i ≤ K j), f i = (f j).comp (inclusion h)) : ↥(supr K) →ₐ[R] B := -{ to_fun := set.lift_of_eq_Union (λ i, ↑(K i)) (λ i x, f i x) +{ to_fun := set.Union_lift (λ i, ↑(K i)) (λ i x, f i x) (λ i j x hxi hxj, let ⟨k, hik, hjk⟩ := dir i j in begin rw [hf i k hik, hf j k hjk], refl end) _ - (by exactI coe_supr_of_directed dir), - map_one' := set.lift_of_eq_Union_const _ (λ _, 1) (λ _, rfl) _ (by simp), - map_zero' := set.lift_of_eq_Union_const _ (λ _, 0) (λ _, rfl) _ (by simp), - map_mul' := set.lift_of_eq_Union_binary dir _ (λ _, (*)) (λ _ _ _, rfl) _ (by simp), - map_add' := set.lift_of_eq_Union_binary dir _ (λ _, (+)) (λ _ _ _, rfl) _ (by simp), - commutes' := λ r, set.lift_of_eq_Union_const _ (λ _, algebra_map _ _ r) + (le_of_eq $ by exactI coe_supr_of_directed dir), + map_one' := set.Union_lift_const _ (λ _, 1) (λ _, rfl) _ (by simp), + map_zero' := set.Union_lift_const _ (λ _, 0) (λ _, rfl) _ (by simp), + map_mul' := set.Union_lift_binary dir _ (λ _, (*)) (λ _ _ _, rfl) _ (by simp), + map_add' := set.Union_lift_binary dir _ (λ _, (+)) (λ _ _ _, rfl) _ (by simp), + commutes' := λ r, set.Union_lift_const _ (λ _, algebra_map _ _ r) (λ _, rfl) _ (λ i, by erw [alg_hom.commutes (f i)]) } variables [nonempty ι] {K : ι → subalgebra R A} {dir : directed (≤) K} From 1ac0b94b1b90bbe413ec364ed444c91f878aa1e0 Mon Sep 17 00:00:00 2001 From: Chris Hughes Date: Sun, 12 Sep 2021 20:26:05 +0100 Subject: [PATCH 44/51] fix build --- src/algebra/algebra/subalgebra.lean | 16 +++++++++------- src/data/set/Union_lift.lean | 2 +- 2 files changed, 10 insertions(+), 8 deletions(-) diff --git a/src/algebra/algebra/subalgebra.lean b/src/algebra/algebra/subalgebra.lean index 84920aeef4450..17dd23caadb80 100644 --- a/src/algebra/algebra/subalgebra.lean +++ b/src/algebra/algebra/subalgebra.lean @@ -783,12 +783,14 @@ noncomputable def supr_lift [nonempty ι] begin rw [hf i k hik, hf j k hjk], refl - end) _ - (le_of_eq $ by exactI coe_supr_of_directed dir), + end) ↑(supr K) + (by rw coe_supr_of_directed dir; refl), map_one' := set.Union_lift_const _ (λ _, 1) (λ _, rfl) _ (by simp), map_zero' := set.Union_lift_const _ (λ _, 0) (λ _, rfl) _ (by simp), - map_mul' := set.Union_lift_binary dir _ (λ _, (*)) (λ _ _ _, rfl) _ (by simp), - map_add' := set.Union_lift_binary dir _ (λ _, (+)) (λ _ _ _, rfl) _ (by simp), + map_mul' := set.Union_lift_binary (coe_supr_of_directed dir) dir _ + (λ _, (*)) (λ _ _ _, rfl) _ (by simp), + map_add' := set.Union_lift_binary (coe_supr_of_directed dir) dir _ + (λ _, (+)) (λ _ _ _, rfl) _ (by simp), commutes' := λ r, set.Union_lift_const _ (λ _, algebra_map _ _ r) (λ _, rfl) _ (λ i, by erw [alg_hom.commutes (f i)]) } @@ -799,7 +801,7 @@ variables [nonempty ι] {K : ι → subalgebra R A} {dir : directed (≤) K} @[simp] lemma supr_lift_inclusion {i : ι} (x : K i) (h : K i ≤ supr K := le_supr K i) : supr_lift K dir f hf (inclusion h x) = f i x := -set.lift_of_eq_Union_inclusion _ +set.Union_lift_inclusion _ _ @[simp] lemma supr_lift_comp_inclusion {i : ι} (h : K i ≤ supr K := le_supr K i) : @@ -809,11 +811,11 @@ by ext; simp @[simp] lemma supr_lift_mk {i : ι} (x : K i) (hx : (x : A) ∈ supr K := set_like.le_def.2 (le_supr K i) x.prop) : supr_lift K dir f hf ⟨x, hx⟩ = f i x := -set.lift_of_eq_Union_mk x hx +set.Union_lift_mk x hx lemma supr_lift_of_mem {i : ι} (x : supr K) (hx : (x : A) ∈ K i) : supr_lift K dir f hf x = f i ⟨x, hx⟩ := -set.lift_of_eq_Union_of_mem x hx +set.Union_lift_of_mem x hx end supr_lift diff --git a/src/data/set/Union_lift.lean b/src/data/set/Union_lift.lean index c1bb0dea79d2e..f97281e182bd2 100644 --- a/src/data/set/Union_lift.lean +++ b/src/data/set/Union_lift.lean @@ -58,7 +58,7 @@ variables {S : ι → set α} {f : Π i (x : S i), β} {hf : ∀ i j (x : α) (hxi : x ∈ S i) (hxj : x ∈ S j), f i ⟨x, hxi⟩ = f j ⟨x, hxj⟩} - {T : set α} {hT : T ⊆ Union S} {hT' : T = Union S} + {T : set α} {hT : T ⊆ Union S} (hT' : T = Union S) @[simp] lemma Union_lift_mk {i : ι} (x : S i) (hx : (x : α) ∈ T) : From 4547efe01cc133e739d2d32b2bd70d0a7c3fee89 Mon Sep 17 00:00:00 2001 From: Chris Hughes Date: Mon, 13 Sep 2021 10:20:36 +0100 Subject: [PATCH 45/51] docstring --- src/data/set/Union_lift.lean | 2 ++ 1 file changed, 2 insertions(+) diff --git a/src/data/set/Union_lift.lean b/src/data/set/Union_lift.lean index f97281e182bd2..e1017d4fd6260 100644 --- a/src/data/set/Union_lift.lean +++ b/src/data/set/Union_lift.lean @@ -151,6 +151,8 @@ variables {hf : ∀ i j (x : α) (hxi : x ∈ S i) (hxj : x ∈ S j), f i ⟨x, hxi⟩ = f j ⟨x, hxj⟩} {hS : Union S = univ} +/-- Glue together functions defined on each of a collection `S` of sets that cover a type. See + also `set.Union_lift`. -/ noncomputable def lift_cover (S : ι → set α) (f : Π i (x : S i), β) From e7aebce9ca4cdb4d0282b5dc6bf4b82af855a174 Mon Sep 17 00:00:00 2001 From: Chris Hughes Date: Wed, 15 Sep 2021 13:22:33 +0100 Subject: [PATCH 46/51] add T : subalgebra R A argument --- src/algebra/algebra/subalgebra.lean | 32 ++++++++++++++--------------- 1 file changed, 16 insertions(+), 16 deletions(-) diff --git a/src/algebra/algebra/subalgebra.lean b/src/algebra/algebra/subalgebra.lean index 17dd23caadb80..95885a059307b 100644 --- a/src/algebra/algebra/subalgebra.lean +++ b/src/algebra/algebra/subalgebra.lean @@ -775,8 +775,10 @@ noncomputable def supr_lift [nonempty ι] (K : ι → subalgebra R A) (dir : directed (≤) K) (f : Π i, K i →ₐ[R] B) - (hf : ∀ (i j : ι) (h : K i ≤ K j), f i = (f j).comp (inclusion h)) : - ↥(supr K) →ₐ[R] B := + (hf : ∀ (i j : ι) (h : K i ≤ K j), f i = (f j).comp (inclusion h)) + (T : subalgebra R A) (hT : T = supr K) : + ↥T →ₐ[R] B := +by subst hT; exact { to_fun := set.Union_lift (λ i, ↑(K i)) (λ i x, f i x) (λ i j x hxi hxj, let ⟨k, hik, hjk⟩ := dir i j in @@ -797,25 +799,23 @@ noncomputable def supr_lift [nonempty ι] variables [nonempty ι] {K : ι → subalgebra R A} {dir : directed (≤) K} {f : Π i, K i →ₐ[R] B} {hf : ∀ (i j : ι) (h : K i ≤ K j), f i = (f j).comp (inclusion h)} + {T : subalgebra R A} {hT : T = supr K} -@[simp] lemma supr_lift_inclusion {i : ι} (x : K i) - (h : K i ≤ supr K := le_supr K i) : - supr_lift K dir f hf (inclusion h x) = f i x := -set.Union_lift_inclusion _ _ +@[simp] lemma supr_lift_inclusion {i : ι} (x : K i) (h : K i ≤ T) : + supr_lift K dir f hf T hT (inclusion h x) = f i x := +by subst T; exact set.Union_lift_inclusion _ _ -@[simp] lemma supr_lift_comp_inclusion {i : ι} - (h : K i ≤ supr K := le_supr K i) : - (supr_lift K dir f hf).comp (inclusion h) = f i := +@[simp] lemma supr_lift_comp_inclusion {i : ι} (h : K i ≤ T) : + (supr_lift K dir f hf T hT).comp (inclusion h) = f i := by ext; simp -@[simp] lemma supr_lift_mk {i : ι} (x : K i) - (hx : (x : A) ∈ supr K := set_like.le_def.2 (le_supr K i) x.prop) : - supr_lift K dir f hf ⟨x, hx⟩ = f i x := -set.Union_lift_mk x hx +@[simp] lemma supr_lift_mk {i : ι} (x : K i) (hx : (x : A) ∈ T) : + supr_lift K dir f hf T hT ⟨x, hx⟩ = f i x := +by subst hT; exact set.Union_lift_mk x hx -lemma supr_lift_of_mem {i : ι} (x : supr K) (hx : (x : A) ∈ K i) : - supr_lift K dir f hf x = f i ⟨x, hx⟩ := -set.Union_lift_of_mem x hx +lemma supr_lift_of_mem {i : ι} (x : T) (hx : (x : A) ∈ K i) : + supr_lift K dir f hf T hT x = f i ⟨x, hx⟩ := +by subst hT; exact set.Union_lift_of_mem x hx end supr_lift From 07bb91d28c354fc8e1d8488c93101214b9a9de1c Mon Sep 17 00:00:00 2001 From: Chris Hughes <33847686+ChrisHughes24@users.noreply.github.com> Date: Wed, 15 Sep 2021 14:45:46 +0100 Subject: [PATCH 47/51] Update src/data/set/Union_lift.lean Co-authored-by: Johan Commelin --- src/data/set/Union_lift.lean | 7 +------ 1 file changed, 1 insertion(+), 6 deletions(-) diff --git a/src/data/set/Union_lift.lean b/src/data/set/Union_lift.lean index e1017d4fd6260..92c9952569a3b 100644 --- a/src/data/set/Union_lift.lean +++ b/src/data/set/Union_lift.lean @@ -79,12 +79,7 @@ by cases x with x hx; exact hf _ _ _ _ _ of algebraic structures when defined on the Union of algebraic subobjects. For example, it could be used to prove that the lift of a collection of group homomorphisms on a union of subgroups preserves `1`. -/ -lemma Union_lift_const - (c : T) - (ci : Π i, S i) - (hci : ∀ i, (ci i : α) = c) - (cβ : β) - (h : ∀ i, f i (ci i) = cβ) : +lemma Union_lift_const (c : T) (ci : Π i, S i) (hci : ∀ i, (ci i : α) = c) (cβ : β) (h : ∀ i, f i (ci i) = cβ) : Union_lift S f hf T hT c = cβ := let ⟨i, hi⟩ := set.mem_Union.1 (hT c.prop) in have (ci i) = ⟨c, hi⟩, from subtype.ext (hci i), From 22b584b58888c1fbe28f5ce8a37f71d187ec5c64 Mon Sep 17 00:00:00 2001 From: Chris Hughes <33847686+ChrisHughes24@users.noreply.github.com> Date: Wed, 15 Sep 2021 14:46:04 +0100 Subject: [PATCH 48/51] Update src/data/set/Union_lift.lean Co-authored-by: Johan Commelin --- src/data/set/Union_lift.lean | 4 +--- 1 file changed, 1 insertion(+), 3 deletions(-) diff --git a/src/data/set/Union_lift.lean b/src/data/set/Union_lift.lean index 92c9952569a3b..034cd08ec5b14 100644 --- a/src/data/set/Union_lift.lean +++ b/src/data/set/Union_lift.lean @@ -89,9 +89,7 @@ by rw [Union_lift_of_mem _ hi, ← this, h] of algebraic structures when defined on the Union of algebraic subobjects. For example, it could be used to prove that the lift of a collection of linear_maps on a union of submodules preserves scalar multiplication. -/ -lemma Union_lift_unary - (u : T → T) - (ui : Π i, S i → S i) +lemma Union_lift_unary (u : T → T) (ui : Π i, S i → S i) (hui : ∀ i (x : S i), u (set.inclusion (show S i ⊆ T, from hT'.symm ▸ set.subset_Union S i) x) = set.inclusion (show S i ⊆ T, from hT'.symm ▸ set.subset_Union S i) (ui i x)) (uβ : β → β) From 6a16989c00d5ed6671d362ac44969f36473b0c57 Mon Sep 17 00:00:00 2001 From: Chris Hughes <33847686+ChrisHughes24@users.noreply.github.com> Date: Wed, 15 Sep 2021 14:46:17 +0100 Subject: [PATCH 49/51] Update src/data/set/Union_lift.lean Co-authored-by: Johan Commelin --- src/data/set/Union_lift.lean | 5 +---- 1 file changed, 1 insertion(+), 4 deletions(-) diff --git a/src/data/set/Union_lift.lean b/src/data/set/Union_lift.lean index 034cd08ec5b14..dec00731928ca 100644 --- a/src/data/set/Union_lift.lean +++ b/src/data/set/Union_lift.lean @@ -110,10 +110,7 @@ end of algebraic structures when defined on the Union of algebraic subobjects. For example, it could be used to prove that the lift of a collection of group homomorphisms on a union of subgroups preserves `*`. -/ -lemma Union_lift_binary - (dir: directed (≤) S) - (op : T → T → T) - (opi : Π i, S i → S i → S i) +lemma Union_lift_binary (dir: directed (≤) S) (op : T → T → T) (opi : Π i, S i → S i → S i) (hopi : ∀ i x y, set.inclusion (show S i ⊆ T, from hT'.symm ▸ set.subset_Union S i) (opi i x y) = op (set.inclusion (show S i ⊆ T, from hT'.symm ▸ set.subset_Union S i) x) (set.inclusion (show S i ⊆ T, from hT'.symm ▸ set.subset_Union S i) y)) From 9e2ef06a9f97eb4c2b5ca0cb28cd8d941d2a8223 Mon Sep 17 00:00:00 2001 From: Chris Hughes Date: Wed, 15 Sep 2021 16:55:57 +0100 Subject: [PATCH 50/51] lint line length --- src/data/set/Union_lift.lean | 4 ++-- 1 file changed, 2 insertions(+), 2 deletions(-) diff --git a/src/data/set/Union_lift.lean b/src/data/set/Union_lift.lean index dec00731928ca..c65156fc3296e 100644 --- a/src/data/set/Union_lift.lean +++ b/src/data/set/Union_lift.lean @@ -79,8 +79,8 @@ by cases x with x hx; exact hf _ _ _ _ _ of algebraic structures when defined on the Union of algebraic subobjects. For example, it could be used to prove that the lift of a collection of group homomorphisms on a union of subgroups preserves `1`. -/ -lemma Union_lift_const (c : T) (ci : Π i, S i) (hci : ∀ i, (ci i : α) = c) (cβ : β) (h : ∀ i, f i (ci i) = cβ) : - Union_lift S f hf T hT c = cβ := +lemma Union_lift_const (c : T) (ci : Π i, S i) (hci : ∀ i, (ci i : α) = c) (cβ : β) + (h : ∀ i, f i (ci i) = cβ) : Union_lift S f hf T hT c = cβ := let ⟨i, hi⟩ := set.mem_Union.1 (hT c.prop) in have (ci i) = ⟨c, hi⟩, from subtype.ext (hci i), by rw [Union_lift_of_mem _ hi, ← this, h] From dca29292599732da49d65c6e03ac92c218cc1053 Mon Sep 17 00:00:00 2001 From: Chris Hughes Date: Thu, 16 Sep 2021 01:40:10 +0100 Subject: [PATCH 51/51] fix build --- src/field_theory/algebraic_closure.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/src/field_theory/algebraic_closure.lean b/src/field_theory/algebraic_closure.lean index 91c23c93d5b1e..ef126fa043571 100644 --- a/src/field_theory/algebraic_closure.lean +++ b/src/field_theory/algebraic_closure.lean @@ -453,7 +453,7 @@ by haveI : nonempty c := set.nonempty.to_subtype hcn; exact { simp [← hij.snd x] }, { erw [alg_hom.comp_apply, ← hji.snd (inclusion h x), inclusion_inclusion, inclusion_self, alg_hom.id_apply x] } - end } in + end _ rfl } in ⟨ub, λ N hN, ⟨(le_supr (λ i : c, (i : subfield_with_hom K L M hL).carrier) ⟨N, hN⟩ : _), begin intro x,