author | wenzelm |
Thu, 22 Feb 2024 19:58:52 +0100 | |
changeset 79705 | a6dc0d4ffea2 |
parent 70215 | 8371a25ca177 |
child 80914 | d97fdabd9e2b |
permissions | -rw-r--r-- |
70160
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1 |
(* Title: HOL/Algebra/Algebraic_Closure.thy |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
2 |
Author: Paulo EmÃlio de Vilhena |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
3 |
|
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
4 |
With contributions by Martin Baillon. |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
5 |
*) |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
6 |
|
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
7 |
theory Algebraic_Closure |
70212
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
8 |
imports Indexed_Polynomials Polynomial_Divisibility Finite_Extensions |
70160
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
9 |
|
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
10 |
begin |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
11 |
|
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
12 |
section \<open>Algebraic Closure\<close> |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
13 |
|
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
14 |
subsection \<open>Definitions\<close> |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
15 |
|
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
16 |
inductive iso_incl :: "'a ring \<Rightarrow> 'a ring \<Rightarrow> bool" (infixl "\<lesssim>" 65) for A B |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
17 |
where iso_inclI [intro]: "id \<in> ring_hom A B \<Longrightarrow> iso_incl A B" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
18 |
|
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
19 |
definition law_restrict :: "('a, 'b) ring_scheme \<Rightarrow> 'a ring" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
20 |
where "law_restrict R \<equiv> (ring.truncate R) |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
21 |
\<lparr> mult := (\<lambda>a \<in> carrier R. \<lambda>b \<in> carrier R. a \<otimes>\<^bsub>R\<^esub> b), |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
22 |
add := (\<lambda>a \<in> carrier R. \<lambda>b \<in> carrier R. a \<oplus>\<^bsub>R\<^esub> b) \<rparr>" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
23 |
|
70212
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
24 |
definition (in ring) \<sigma> :: "'a list \<Rightarrow> ((('a list \<times> nat) multiset) \<Rightarrow> 'a) list" |
70160
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
25 |
where "\<sigma> P = map indexed_const P" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
26 |
|
70212
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
27 |
definition (in ring) extensions :: "((('a list \<times> nat) multiset) \<Rightarrow> 'a) ring set" |
70160
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
28 |
where "extensions \<equiv> { L \<comment> \<open>such that\<close>. |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
29 |
\<comment> \<open>i\<close> (field L) \<and> |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
30 |
\<comment> \<open>ii\<close> (indexed_const \<in> ring_hom R L) \<and> |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
31 |
\<comment> \<open>iii\<close> (\<forall>\<P> \<in> carrier L. carrier_coeff \<P>) \<and> |
70212
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
32 |
\<comment> \<open>iv\<close> (\<forall>\<P> \<in> carrier L. \<forall>P \<in> carrier (poly_ring R). \<forall>i. |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
33 |
\<not> index_free \<P> (P, i) \<longrightarrow> |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
34 |
\<X>\<^bsub>(P, i)\<^esub> \<in> carrier L \<and> (ring.eval L) (\<sigma> P) \<X>\<^bsub>(P, i)\<^esub> = \<zero>\<^bsub>L\<^esub>) }" |
70160
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
35 |
|
70212
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
36 |
abbreviation (in ring) restrict_extensions :: "((('a list \<times> nat) multiset) \<Rightarrow> 'a) ring set" ("\<S>") |
70160
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
37 |
where "\<S> \<equiv> law_restrict ` extensions" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
38 |
|
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
39 |
|
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
40 |
subsection \<open>Basic Properties\<close> |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
41 |
|
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
42 |
lemma law_restrict_carrier: "carrier (law_restrict R) = carrier R" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
43 |
by (simp add: law_restrict_def ring.defs) |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
44 |
|
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
45 |
lemma law_restrict_one: "one (law_restrict R) = one R" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
46 |
by (simp add: law_restrict_def ring.defs) |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
47 |
|
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
48 |
lemma law_restrict_zero: "zero (law_restrict R) = zero R" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
49 |
by (simp add: law_restrict_def ring.defs) |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
50 |
|
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
51 |
lemma law_restrict_mult: "monoid.mult (law_restrict R) = (\<lambda>a \<in> carrier R. \<lambda>b \<in> carrier R. a \<otimes>\<^bsub>R\<^esub> b)" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
52 |
by (simp add: law_restrict_def ring.defs) |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
53 |
|
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
54 |
lemma law_restrict_add: "add (law_restrict R) = (\<lambda>a \<in> carrier R. \<lambda>b \<in> carrier R. a \<oplus>\<^bsub>R\<^esub> b)" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
55 |
by (simp add: law_restrict_def ring.defs) |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
56 |
|
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
57 |
lemma (in ring) law_restrict_is_ring: "ring (law_restrict R)" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
58 |
by (unfold_locales) (auto simp add: law_restrict_def Units_def ring.defs, |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
59 |
simp_all add: a_assoc a_comm m_assoc l_distr r_distr a_lcomm) |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
60 |
|
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
61 |
lemma (in field) law_restrict_is_field: "field (law_restrict R)" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
62 |
proof - |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
63 |
have "comm_monoid_axioms (law_restrict R)" |
70215
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
paulson <lp15@cam.ac.uk>
parents:
70214
diff
changeset
|
64 |
using m_comm unfolding comm_monoid_axioms_def law_restrict_carrier law_restrict_mult by auto |
70160
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
65 |
then interpret L: cring "law_restrict R" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
66 |
using cring.intro law_restrict_is_ring comm_monoid.intro ring.is_monoid by auto |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
67 |
have "Units R = Units (law_restrict R)" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
68 |
unfolding Units_def law_restrict_carrier law_restrict_mult law_restrict_one by auto |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
69 |
thus ?thesis |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
70 |
using L.cring_fieldI unfolding field_Units law_restrict_carrier law_restrict_zero by simp |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
71 |
qed |
70215
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
paulson <lp15@cam.ac.uk>
parents:
70214
diff
changeset
|
72 |
|
70160
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
73 |
lemma law_restrict_iso_imp_eq: |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
74 |
assumes "id \<in> ring_iso (law_restrict A) (law_restrict B)" and "ring A" and "ring B" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
75 |
shows "law_restrict A = law_restrict B" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
76 |
proof - |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
77 |
have "carrier A = carrier B" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
78 |
using ring_iso_memE(5)[OF assms(1)] unfolding bij_betw_def law_restrict_def by (simp add: ring.defs) |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
79 |
hence mult: "a \<otimes>\<^bsub>law_restrict A\<^esub> b = a \<otimes>\<^bsub>law_restrict B\<^esub> b" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
80 |
and add: "a \<oplus>\<^bsub>law_restrict A\<^esub> b = a \<oplus>\<^bsub>law_restrict B\<^esub> b" for a b |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
81 |
using ring_iso_memE(2-3)[OF assms(1)] unfolding law_restrict_def by (auto simp add: ring.defs) |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
82 |
have "monoid.mult (law_restrict A) = monoid.mult (law_restrict B)" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
83 |
using mult by auto |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
84 |
moreover have "add (law_restrict A) = add (law_restrict B)" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
85 |
using add by auto |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
86 |
moreover from \<open>carrier A = carrier B\<close> have "carrier (law_restrict A) = carrier (law_restrict B)" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
87 |
unfolding law_restrict_def by (simp add: ring.defs) |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
88 |
moreover have "\<zero>\<^bsub>law_restrict A\<^esub> = \<zero>\<^bsub>law_restrict B\<^esub>" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
89 |
using ring_hom_zero[OF _ assms(2-3)[THEN ring.law_restrict_is_ring]] assms(1) |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
90 |
unfolding ring_iso_def by auto |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
91 |
moreover have "\<one>\<^bsub>law_restrict A\<^esub> = \<one>\<^bsub>law_restrict B\<^esub>" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
92 |
using ring_iso_memE(4)[OF assms(1)] by simp |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
93 |
ultimately show ?thesis by simp |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
94 |
qed |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
95 |
|
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
96 |
lemma law_restrict_hom: "h \<in> ring_hom A B \<longleftrightarrow> h \<in> ring_hom (law_restrict A) (law_restrict B)" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
97 |
proof |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
98 |
assume "h \<in> ring_hom A B" thus "h \<in> ring_hom (law_restrict A) (law_restrict B)" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
99 |
by (auto intro!: ring_hom_memI dest: ring_hom_memE simp: law_restrict_def ring.defs) |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
100 |
next |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
101 |
assume h: "h \<in> ring_hom (law_restrict A) (law_restrict B)" show "h \<in> ring_hom A B" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
102 |
using ring_hom_memE[OF h] by (auto intro!: ring_hom_memI simp: law_restrict_def ring.defs) |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
103 |
qed |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
104 |
|
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
105 |
lemma iso_incl_hom: "A \<lesssim> B \<longleftrightarrow> (law_restrict A) \<lesssim> (law_restrict B)" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
106 |
using law_restrict_hom iso_incl.simps by blast |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
107 |
|
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
108 |
|
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
109 |
subsection \<open>Partial Order\<close> |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
110 |
|
70215
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
paulson <lp15@cam.ac.uk>
parents:
70214
diff
changeset
|
111 |
lemma iso_incl_backwards: |
70160
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
112 |
assumes "A \<lesssim> B" shows "id \<in> ring_hom A B" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
113 |
using assms by cases |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
114 |
|
70215
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
paulson <lp15@cam.ac.uk>
parents:
70214
diff
changeset
|
115 |
lemma iso_incl_antisym_aux: |
70160
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
116 |
assumes "A \<lesssim> B" and "B \<lesssim> A" shows "id \<in> ring_iso A B" |
70215
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
paulson <lp15@cam.ac.uk>
parents:
70214
diff
changeset
|
117 |
proof - |
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
paulson <lp15@cam.ac.uk>
parents:
70214
diff
changeset
|
118 |
have hom: "id \<in> ring_hom A B" "id \<in> ring_hom B A" |
70160
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
119 |
using assms(1-2)[THEN iso_incl_backwards] by auto |
70215
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
paulson <lp15@cam.ac.uk>
parents:
70214
diff
changeset
|
120 |
thus ?thesis |
70160
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
121 |
using hom[THEN ring_hom_memE(1)] by (auto simp add: ring_iso_def bij_betw_def inj_on_def) |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
122 |
qed |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
123 |
|
70215
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
paulson <lp15@cam.ac.uk>
parents:
70214
diff
changeset
|
124 |
lemma iso_incl_refl: "A \<lesssim> A" |
70160
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
125 |
by (rule iso_inclI[OF ring_hom_memI], auto) |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
126 |
|
70215
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
paulson <lp15@cam.ac.uk>
parents:
70214
diff
changeset
|
127 |
lemma iso_incl_trans: |
70160
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
128 |
assumes "A \<lesssim> B" and "B \<lesssim> C" shows "A \<lesssim> C" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
129 |
using ring_hom_trans[OF assms[THEN iso_incl_backwards]] by auto |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
130 |
|
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
131 |
lemma (in ring) iso_incl_antisym: |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
132 |
assumes "A \<in> \<S>" "B \<in> \<S>" and "A \<lesssim> B" "B \<lesssim> A" shows "A = B" |
70215
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
paulson <lp15@cam.ac.uk>
parents:
70214
diff
changeset
|
133 |
proof - |
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
paulson <lp15@cam.ac.uk>
parents:
70214
diff
changeset
|
134 |
obtain A' B' :: "(('a list \<times> nat) multiset \<Rightarrow> 'a) ring" |
70160
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
135 |
where A: "A = law_restrict A'" "ring A'" and B: "B = law_restrict B'" "ring B'" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
136 |
using assms(1-2) field.is_ring by (auto simp add: extensions_def) |
70215
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
paulson <lp15@cam.ac.uk>
parents:
70214
diff
changeset
|
137 |
thus ?thesis |
70160
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
138 |
using law_restrict_iso_imp_eq iso_incl_antisym_aux[OF assms(3-4)] by simp |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
139 |
qed |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
140 |
|
70212
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
141 |
lemma (in ring) iso_incl_partial_order: "partial_order_on \<S> (relation_of (\<lesssim>) \<S>)" |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
142 |
using iso_incl_refl iso_incl_trans iso_incl_antisym by (rule partial_order_on_relation_ofI) |
70160
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
143 |
|
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
144 |
lemma iso_inclE: |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
145 |
assumes "ring A" and "ring B" and "A \<lesssim> B" shows "ring_hom_ring A B id" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
146 |
using iso_incl_backwards[OF assms(3)] ring_hom_ring.intro[OF assms(1-2)] |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
147 |
unfolding symmetric[OF ring_hom_ring_axioms_def] by simp |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
148 |
|
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
149 |
lemma iso_incl_imp_same_eval: |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
150 |
assumes "ring A" and "ring B" and "A \<lesssim> B" and "a \<in> carrier A" and "set p \<subseteq> carrier A" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
151 |
shows "(ring.eval A) p a = (ring.eval B) p a" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
152 |
using ring_hom_ring.eval_hom'[OF iso_inclE[OF assms(1-3)] assms(4-5)] by simp |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
153 |
|
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
154 |
|
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
155 |
subsection \<open>Extensions Non Empty\<close> |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
156 |
|
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
157 |
lemma (in ring) indexed_const_is_inj: "inj indexed_const" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
158 |
unfolding indexed_const_def by (rule inj_onI, metis) |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
159 |
|
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
160 |
lemma (in ring) indexed_const_inj_on: "inj_on indexed_const (carrier R)" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
161 |
unfolding indexed_const_def by (rule inj_onI, metis) |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
162 |
|
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
163 |
lemma (in field) extensions_non_empty: "\<S> \<noteq> {}" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
164 |
proof - |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
165 |
have "image_ring indexed_const R \<in> extensions" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
166 |
proof (auto simp add: extensions_def) |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
167 |
show "field (image_ring indexed_const R)" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
168 |
using inj_imp_image_ring_is_field[OF indexed_const_inj_on] . |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
169 |
next |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
170 |
show "indexed_const \<in> ring_hom R (image_ring indexed_const R)" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
171 |
using inj_imp_image_ring_iso[OF indexed_const_inj_on] unfolding ring_iso_def by auto |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
172 |
next |
70212
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
173 |
fix \<P> :: "(('a list \<times> nat) multiset) \<Rightarrow> 'a" and P and i |
70160
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
174 |
assume "\<P> \<in> carrier (image_ring indexed_const R)" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
175 |
then obtain k where "k \<in> carrier R" and "\<P> = indexed_const k" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
176 |
unfolding image_ring_carrier by blast |
70212
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
177 |
hence "index_free \<P> (P, i)" for P i |
70160
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
178 |
unfolding index_free_def indexed_const_def by auto |
70212
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
179 |
thus "\<not> index_free \<P> (P, i) \<Longrightarrow> \<X>\<^bsub>(P, i)\<^esub> \<in> carrier (image_ring indexed_const R)" |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
180 |
and "\<not> index_free \<P> (P, i) \<Longrightarrow> ring.eval (image_ring indexed_const R) (\<sigma> P) \<X>\<^bsub>(P, i)\<^esub> = \<zero>\<^bsub>image_ring indexed_const R\<^esub>" |
70160
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
181 |
by auto |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
182 |
from \<open>k \<in> carrier R\<close> and \<open>\<P> = indexed_const k\<close> show "carrier_coeff \<P>" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
183 |
unfolding indexed_const_def carrier_coeff_def by auto |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
184 |
qed |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
185 |
thus ?thesis |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
186 |
by blast |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
187 |
qed |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
188 |
|
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
189 |
|
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
190 |
subsection \<open>Chains\<close> |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
191 |
|
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
192 |
definition union_ring :: "(('a, 'c) ring_scheme) set \<Rightarrow> 'a ring" |
70215
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
paulson <lp15@cam.ac.uk>
parents:
70214
diff
changeset
|
193 |
where "union_ring C = |
70160
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
194 |
\<lparr> carrier = (\<Union>(carrier ` C)), |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
195 |
monoid.mult = (\<lambda>a b. (monoid.mult (SOME R. R \<in> C \<and> a \<in> carrier R \<and> b \<in> carrier R) a b)), |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
196 |
one = one (SOME R. R \<in> C), |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
197 |
zero = zero (SOME R. R \<in> C), |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
198 |
add = (\<lambda>a b. (add (SOME R. R \<in> C \<and> a \<in> carrier R \<and> b \<in> carrier R) a b)) \<rparr>" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
199 |
|
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
200 |
|
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
201 |
lemma union_ring_carrier: "carrier (union_ring C) = (\<Union>(carrier ` C))" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
202 |
unfolding union_ring_def by simp |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
203 |
|
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
204 |
context |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
205 |
fixes C :: "'a ring set" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
206 |
assumes field_chain: "\<And>R. R \<in> C \<Longrightarrow> field R" and chain: "\<And>R S. \<lbrakk> R \<in> C; S \<in> C \<rbrakk> \<Longrightarrow> R \<lesssim> S \<or> S \<lesssim> R" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
207 |
begin |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
208 |
|
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
209 |
lemma ring_chain: "R \<in> C \<Longrightarrow> ring R" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
210 |
using field.is_ring[OF field_chain] by blast |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
211 |
|
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
212 |
lemma same_one_same_zero: |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
213 |
assumes "R \<in> C" shows "\<one>\<^bsub>union_ring C\<^esub> = \<one>\<^bsub>R\<^esub>" and "\<zero>\<^bsub>union_ring C\<^esub> = \<zero>\<^bsub>R\<^esub>" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
214 |
proof - |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
215 |
have "\<one>\<^bsub>R\<^esub> = \<one>\<^bsub>S\<^esub>" if "R \<in> C" and "S \<in> C" for R S |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
216 |
using ring_hom_one[of id] chain[OF that] unfolding iso_incl.simps by auto |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
217 |
moreover have "\<zero>\<^bsub>R\<^esub> = \<zero>\<^bsub>S\<^esub>" if "R \<in> C" and "S \<in> C" for R S |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
218 |
using chain[OF that] ring_hom_zero[OF _ ring_chain ring_chain] that unfolding iso_incl.simps by auto |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
219 |
ultimately have "one (SOME R. R \<in> C) = \<one>\<^bsub>R\<^esub>" and "zero (SOME R. R \<in> C) = \<zero>\<^bsub>R\<^esub>" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
220 |
using assms by (metis (mono_tags) someI)+ |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
221 |
thus "\<one>\<^bsub>union_ring C\<^esub> = \<one>\<^bsub>R\<^esub>" and "\<zero>\<^bsub>union_ring C\<^esub> = \<zero>\<^bsub>R\<^esub>" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
222 |
unfolding union_ring_def by auto |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
223 |
qed |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
224 |
|
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
225 |
lemma same_laws: |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
226 |
assumes "R \<in> C" and "a \<in> carrier R" and "b \<in> carrier R" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
227 |
shows "a \<otimes>\<^bsub>union_ring C\<^esub> b = a \<otimes>\<^bsub>R\<^esub> b" and "a \<oplus>\<^bsub>union_ring C\<^esub> b = a \<oplus>\<^bsub>R\<^esub> b" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
228 |
proof - |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
229 |
have "a \<otimes>\<^bsub>R\<^esub> b = a \<otimes>\<^bsub>S\<^esub> b" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
230 |
if "R \<in> C" "a \<in> carrier R" "b \<in> carrier R" and "S \<in> C" "a \<in> carrier S" "b \<in> carrier S" for R S |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
231 |
using ring_hom_memE(2)[of id R S] ring_hom_memE(2)[of id S R] that chain[OF that(1,4)] |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
232 |
unfolding iso_incl.simps by auto |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
233 |
moreover have "a \<oplus>\<^bsub>R\<^esub> b = a \<oplus>\<^bsub>S\<^esub> b" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
234 |
if "R \<in> C" "a \<in> carrier R" "b \<in> carrier R" and "S \<in> C" "a \<in> carrier S" "b \<in> carrier S" for R S |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
235 |
using ring_hom_memE(3)[of id R S] ring_hom_memE(3)[of id S R] that chain[OF that(1,4)] |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
236 |
unfolding iso_incl.simps by auto |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
237 |
ultimately |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
238 |
have "monoid.mult (SOME R. R \<in> C \<and> a \<in> carrier R \<and> b \<in> carrier R) a b = a \<otimes>\<^bsub>R\<^esub> b" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
239 |
and "add (SOME R. R \<in> C \<and> a \<in> carrier R \<and> b \<in> carrier R) a b = a \<oplus>\<^bsub>R\<^esub> b" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
240 |
using assms by (metis (mono_tags, lifting) someI)+ |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
241 |
thus "a \<otimes>\<^bsub>union_ring C\<^esub> b = a \<otimes>\<^bsub>R\<^esub> b" and "a \<oplus>\<^bsub>union_ring C\<^esub> b = a \<oplus>\<^bsub>R\<^esub> b" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
242 |
unfolding union_ring_def by auto |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
243 |
qed |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
244 |
|
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
245 |
lemma exists_superset_carrier: |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
246 |
assumes "finite S" and "S \<noteq> {}" and "S \<subseteq> carrier (union_ring C)" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
247 |
shows "\<exists>R \<in> C. S \<subseteq> carrier R" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
248 |
using assms |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
249 |
proof (induction, simp) |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
250 |
case (insert s S) |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
251 |
obtain R where R: "s \<in> carrier R" "R \<in> C" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
252 |
using insert(5) unfolding union_ring_def by auto |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
253 |
show ?case |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
254 |
proof (cases) |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
255 |
assume "S = {}" thus ?thesis |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
256 |
using R by blast |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
257 |
next |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
258 |
assume "S \<noteq> {}" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
259 |
then obtain T where T: "S \<subseteq> carrier T" "T \<in> C" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
260 |
using insert(3,5) by blast |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
261 |
have "carrier R \<subseteq> carrier T \<or> carrier T \<subseteq> carrier R" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
262 |
using ring_hom_memE(1)[of id R] ring_hom_memE(1)[of id T] chain[OF R(2) T(2)] |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
263 |
unfolding iso_incl.simps by auto |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
264 |
thus ?thesis |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
265 |
using R T by auto |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
266 |
qed |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
267 |
qed |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
268 |
|
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
269 |
lemma union_ring_is_monoid: |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
270 |
assumes "C \<noteq> {}" shows "comm_monoid (union_ring C)" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
271 |
proof |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
272 |
fix a b c |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
273 |
assume "a \<in> carrier (union_ring C)" "b \<in> carrier (union_ring C)" "c \<in> carrier (union_ring C)" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
274 |
then obtain R where R: "R \<in> C" "a \<in> carrier R" "b \<in> carrier R" "c \<in> carrier R" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
275 |
using exists_superset_carrier[of "{ a, b, c }"] by auto |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
276 |
then interpret field R |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
277 |
using field_chain by simp |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
278 |
|
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
279 |
show "a \<otimes>\<^bsub>union_ring C\<^esub> b \<in> carrier (union_ring C)" |
70215
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
paulson <lp15@cam.ac.uk>
parents:
70214
diff
changeset
|
280 |
using R(1-3) unfolding same_laws(1)[OF R(1-3)] unfolding union_ring_def by auto |
70160
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
281 |
show "(a \<otimes>\<^bsub>union_ring C\<^esub> b) \<otimes>\<^bsub>union_ring C\<^esub> c = a \<otimes>\<^bsub>union_ring C\<^esub> (b \<otimes>\<^bsub>union_ring C\<^esub> c)" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
282 |
and "a \<otimes>\<^bsub>union_ring C\<^esub> b = b \<otimes>\<^bsub>union_ring C\<^esub> a" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
283 |
and "\<one>\<^bsub>union_ring C\<^esub> \<otimes>\<^bsub>union_ring C\<^esub> a = a" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
284 |
and "a \<otimes>\<^bsub>union_ring C\<^esub> \<one>\<^bsub>union_ring C\<^esub> = a" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
285 |
using same_one_same_zero[OF R(1)] same_laws(1)[OF R(1)] R(2-4) m_assoc m_comm by auto |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
286 |
next |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
287 |
show "\<one>\<^bsub>union_ring C\<^esub> \<in> carrier (union_ring C)" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
288 |
using ring.ring_simprules(6)[OF ring_chain] assms same_one_same_zero(1) |
70215
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
paulson <lp15@cam.ac.uk>
parents:
70214
diff
changeset
|
289 |
unfolding union_ring_carrier by auto |
70160
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
290 |
qed |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
291 |
|
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
292 |
lemma union_ring_is_abelian_group: |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
293 |
assumes "C \<noteq> {}" shows "cring (union_ring C)" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
294 |
proof (rule cringI[OF abelian_groupI union_ring_is_monoid[OF assms]]) |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
295 |
fix a b c |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
296 |
assume "a \<in> carrier (union_ring C)" "b \<in> carrier (union_ring C)" "c \<in> carrier (union_ring C)" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
297 |
then obtain R where R: "R \<in> C" "a \<in> carrier R" "b \<in> carrier R" "c \<in> carrier R" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
298 |
using exists_superset_carrier[of "{ a, b, c }"] by auto |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
299 |
then interpret field R |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
300 |
using field_chain by simp |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
301 |
|
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
302 |
show "a \<oplus>\<^bsub>union_ring C\<^esub> b \<in> carrier (union_ring C)" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
303 |
using R(1-3) unfolding same_laws(2)[OF R(1-3)] unfolding union_ring_def by auto |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
304 |
show "(a \<oplus>\<^bsub>union_ring C\<^esub> b) \<otimes>\<^bsub>union_ring C\<^esub> c = (a \<otimes>\<^bsub>union_ring C\<^esub> c) \<oplus>\<^bsub>union_ring C\<^esub> (b \<otimes>\<^bsub>union_ring C\<^esub> c)" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
305 |
and "(a \<oplus>\<^bsub>union_ring C\<^esub> b) \<oplus>\<^bsub>union_ring C\<^esub> c = a \<oplus>\<^bsub>union_ring C\<^esub> (b \<oplus>\<^bsub>union_ring C\<^esub> c)" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
306 |
and "a \<oplus>\<^bsub>union_ring C\<^esub> b = b \<oplus>\<^bsub>union_ring C\<^esub> a" |
70215
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
paulson <lp15@cam.ac.uk>
parents:
70214
diff
changeset
|
307 |
and "\<zero>\<^bsub>union_ring C\<^esub> \<oplus>\<^bsub>union_ring C\<^esub> a = a" |
70160
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
308 |
using same_one_same_zero[OF R(1)] same_laws[OF R(1)] R(2-4) l_distr a_assoc a_comm by auto |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
309 |
have "\<exists>a' \<in> carrier R. a' \<oplus>\<^bsub>union_ring C\<^esub> a = \<zero>\<^bsub>union_ring C\<^esub>" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
310 |
using same_laws(2)[OF R(1)] R(2) same_one_same_zero[OF R(1)] by simp |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
311 |
with \<open>R \<in> C\<close> show "\<exists>y \<in> carrier (union_ring C). y \<oplus>\<^bsub>union_ring C\<^esub> a = \<zero>\<^bsub>union_ring C\<^esub>" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
312 |
unfolding union_ring_carrier by auto |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
313 |
next |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
314 |
show "\<zero>\<^bsub>union_ring C\<^esub> \<in> carrier (union_ring C)" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
315 |
using ring.ring_simprules(2)[OF ring_chain] assms same_one_same_zero(2) |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
316 |
unfolding union_ring_carrier by auto |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
317 |
qed |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
318 |
|
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
319 |
lemma union_ring_is_field : |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
320 |
assumes "C \<noteq> {}" shows "field (union_ring C)" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
321 |
proof (rule cring.cring_fieldI[OF union_ring_is_abelian_group[OF assms]]) |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
322 |
have "carrier (union_ring C) - { \<zero>\<^bsub>union_ring C\<^esub> } \<subseteq> Units (union_ring C)" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
323 |
proof |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
324 |
fix a assume "a \<in> carrier (union_ring C) - { \<zero>\<^bsub>union_ring C\<^esub> }" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
325 |
hence "a \<in> carrier (union_ring C)" and "a \<noteq> \<zero>\<^bsub>union_ring C\<^esub>" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
326 |
by auto |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
327 |
then obtain R where R: "R \<in> C" "a \<in> carrier R" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
328 |
using exists_superset_carrier[of "{ a }"] by auto |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
329 |
then interpret field R |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
330 |
using field_chain by simp |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
331 |
|
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
332 |
from \<open>a \<in> carrier R\<close> and \<open>a \<noteq> \<zero>\<^bsub>union_ring C\<^esub>\<close> have "a \<in> Units R" |
70215
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
paulson <lp15@cam.ac.uk>
parents:
70214
diff
changeset
|
333 |
unfolding same_one_same_zero[OF R(1)] field_Units by auto |
70160
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
334 |
hence "\<exists>a' \<in> carrier R. a' \<otimes>\<^bsub>union_ring C\<^esub> a = \<one>\<^bsub>union_ring C\<^esub> \<and> a \<otimes>\<^bsub>union_ring C\<^esub> a' = \<one>\<^bsub>union_ring C\<^esub>" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
335 |
using same_laws[OF R(1)] same_one_same_zero[OF R(1)] R(2) unfolding Units_def by auto |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
336 |
with \<open>R \<in> C\<close> and \<open>a \<in> carrier (union_ring C)\<close> show "a \<in> Units (union_ring C)" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
337 |
unfolding Units_def union_ring_carrier by auto |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
338 |
qed |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
339 |
moreover have "\<zero>\<^bsub>union_ring C\<^esub> \<notin> Units (union_ring C)" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
340 |
proof (rule ccontr) |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
341 |
assume "\<not> \<zero>\<^bsub>union_ring C\<^esub> \<notin> Units (union_ring C)" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
342 |
then obtain a where a: "a \<in> carrier (union_ring C)" "a \<otimes>\<^bsub>union_ring C\<^esub> \<zero>\<^bsub>union_ring C\<^esub> = \<one>\<^bsub>union_ring C\<^esub>" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
343 |
unfolding Units_def by auto |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
344 |
then obtain R where R: "R \<in> C" "a \<in> carrier R" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
345 |
using exists_superset_carrier[of "{ a }"] by auto |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
346 |
then interpret field R |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
347 |
using field_chain by simp |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
348 |
have "\<one>\<^bsub>R\<^esub> = \<zero>\<^bsub>R\<^esub>" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
349 |
using a R same_laws(1)[OF R(1)] same_one_same_zero[OF R(1)] by auto |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
350 |
thus False |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
351 |
using one_not_zero by simp |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
352 |
qed |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
353 |
hence "Units (union_ring C) \<subseteq> carrier (union_ring C) - { \<zero>\<^bsub>union_ring C\<^esub> }" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
354 |
unfolding Units_def by auto |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
355 |
ultimately show "Units (union_ring C) = carrier (union_ring C) - { \<zero>\<^bsub>union_ring C\<^esub> }" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
356 |
by simp |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
357 |
qed |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
358 |
|
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
359 |
lemma union_ring_is_upper_bound: |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
360 |
assumes "R \<in> C" shows "R \<lesssim> union_ring C" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
361 |
using ring_hom_memI[of R id "union_ring C"] same_laws[of R] same_one_same_zero[of R] assms |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
362 |
unfolding union_ring_carrier by auto |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
363 |
|
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
364 |
end |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
365 |
|
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
366 |
|
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
367 |
subsection \<open>Zorn\<close> |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
368 |
|
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
369 |
lemma (in ring) exists_core_chain: |
70212
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
370 |
assumes "C \<in> Chains (relation_of (\<lesssim>) \<S>)" obtains C' where "C' \<subseteq> extensions" and "C = law_restrict ` C'" |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
371 |
using Chains_relation_of[OF assms] by (meson subset_image_iff) |
70160
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
372 |
|
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
373 |
lemma (in ring) core_chain_is_chain: |
70212
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
374 |
assumes "law_restrict ` C \<in> Chains (relation_of (\<lesssim>) \<S>)" shows "\<And>R S. \<lbrakk> R \<in> C; S \<in> C \<rbrakk> \<Longrightarrow> R \<lesssim> S \<or> S \<lesssim> R" |
70160
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
375 |
proof - |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
376 |
fix R S assume "R \<in> C" and "S \<in> C" thus "R \<lesssim> S \<or> S \<lesssim> R" |
70212
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
377 |
using assms(1) unfolding iso_incl_hom[of R] iso_incl_hom[of S] Chains_def relation_of_def |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
378 |
by auto |
70160
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
379 |
qed |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
380 |
|
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
381 |
lemma (in field) exists_maximal_extension: |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
382 |
shows "\<exists>M \<in> \<S>. \<forall>L \<in> \<S>. M \<lesssim> L \<longrightarrow> L = M" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
383 |
proof (rule predicate_Zorn[OF iso_incl_partial_order]) |
70212
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
384 |
fix C assume C: "C \<in> Chains (relation_of (\<lesssim>) \<S>)" |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
385 |
show "\<exists>L \<in> \<S>. \<forall>R \<in> C. R \<lesssim> L" |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
386 |
proof (cases) |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
387 |
assume "C = {}" thus ?thesis |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
388 |
using extensions_non_empty by auto |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
389 |
next |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
390 |
assume "C \<noteq> {}" |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
391 |
from \<open>C \<in> Chains (relation_of (\<lesssim>) \<S>)\<close> |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
392 |
obtain C' where C': "C' \<subseteq> extensions" "C = law_restrict ` C'" |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
393 |
using exists_core_chain by auto |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
394 |
with \<open>C \<noteq> {}\<close> obtain S where S: "S \<in> C'" and "C' \<noteq> {}" |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
395 |
by auto |
70160
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
396 |
|
70212
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
397 |
have core_chain: "\<And>R. R \<in> C' \<Longrightarrow> field R" "\<And>R S. \<lbrakk> R \<in> C'; S \<in> C' \<rbrakk> \<Longrightarrow> R \<lesssim> S \<or> S \<lesssim> R" |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
398 |
using core_chain_is_chain[of C'] C' C unfolding extensions_def by auto |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
399 |
from \<open>C' \<noteq> {}\<close> interpret Union: field "union_ring C'" |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
400 |
using union_ring_is_field[OF core_chain] C'(1) by blast |
70160
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
401 |
|
70212
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
402 |
have "union_ring C' \<in> extensions" |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
403 |
proof (auto simp add: extensions_def) |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
404 |
show "field (union_ring C')" |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
405 |
using Union.field_axioms . |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
406 |
next |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
407 |
from \<open>S \<in> C'\<close> have "indexed_const \<in> ring_hom R S" |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
408 |
using C'(1) unfolding extensions_def by auto |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
409 |
thus "indexed_const \<in> ring_hom R (union_ring C')" |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
410 |
using ring_hom_trans[of _ R S id] union_ring_is_upper_bound[OF core_chain S] |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
411 |
unfolding iso_incl.simps by auto |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
412 |
next |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
413 |
show "a \<in> carrier (union_ring C') \<Longrightarrow> carrier_coeff a" for a |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
414 |
using C'(1) unfolding union_ring_carrier extensions_def by auto |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
415 |
next |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
416 |
fix \<P> P i |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
417 |
assume "\<P> \<in> carrier (union_ring C')" |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
418 |
and P: "P \<in> carrier (poly_ring R)" |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
419 |
and not_index_free: "\<not> index_free \<P> (P, i)" |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
420 |
from \<open>\<P> \<in> carrier (union_ring C')\<close> obtain T where T: "T \<in> C'" "\<P> \<in> carrier T" |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
421 |
using exists_superset_carrier[of C' "{ \<P> }"] core_chain by auto |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
422 |
hence "\<X>\<^bsub>(P, i)\<^esub> \<in> carrier T" and "(ring.eval T) (\<sigma> P) \<X>\<^bsub>(P, i)\<^esub> = \<zero>\<^bsub>T\<^esub>" |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
423 |
and field: "field T" and hom: "indexed_const \<in> ring_hom R T" |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
424 |
using P not_index_free C'(1) unfolding extensions_def by auto |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
425 |
with \<open>T \<in> C'\<close> show "\<X>\<^bsub>(P, i)\<^esub> \<in> carrier (union_ring C')" |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
426 |
unfolding union_ring_carrier by auto |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
427 |
have "set P \<subseteq> carrier R" |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
428 |
using P unfolding sym[OF univ_poly_carrier] polynomial_def by auto |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
429 |
hence "set (\<sigma> P) \<subseteq> carrier T" |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
430 |
using ring_hom_memE(1)[OF hom] unfolding \<sigma>_def by (induct P) (auto) |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
431 |
with \<open>\<X>\<^bsub>(P, i)\<^esub> \<in> carrier T\<close> and \<open>(ring.eval T) (\<sigma> P) \<X>\<^bsub>(P, i)\<^esub> = \<zero>\<^bsub>T\<^esub>\<close> |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
432 |
show "(ring.eval (union_ring C')) (\<sigma> P) \<X>\<^bsub>(P, i)\<^esub> = \<zero>\<^bsub>union_ring C'\<^esub>" |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
433 |
using iso_incl_imp_same_eval[OF field.is_ring[OF field] Union.is_ring |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
434 |
union_ring_is_upper_bound[OF core_chain T(1)]] same_one_same_zero(2)[OF core_chain T(1)] |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
435 |
by auto |
70160
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
436 |
qed |
70212
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
437 |
moreover have "R \<lesssim> law_restrict (union_ring C')" if "R \<in> C" for R |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
438 |
using that union_ring_is_upper_bound[OF core_chain] iso_incl_hom unfolding C' by auto |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
439 |
ultimately show ?thesis |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
440 |
by blast |
70160
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
441 |
qed |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
442 |
qed |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
443 |
|
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
444 |
|
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
445 |
subsection \<open>Existence of roots\<close> |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
446 |
|
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
447 |
lemma polynomial_hom: |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
448 |
assumes "h \<in> ring_hom R S" and "field R" and "field S" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
449 |
shows "p \<in> carrier (poly_ring R) \<Longrightarrow> (map h p) \<in> carrier (poly_ring S)" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
450 |
proof - |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
451 |
assume "p \<in> carrier (poly_ring R)" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
452 |
interpret ring_hom_ring R S h |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
453 |
using ring_hom_ringI2[OF assms(2-3)[THEN field.is_ring] assms(1)] . |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
454 |
|
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
455 |
from \<open>p \<in> carrier (poly_ring R)\<close> have "set p \<subseteq> carrier R" and lc: "p \<noteq> [] \<Longrightarrow> lead_coeff p \<noteq> \<zero>\<^bsub>R\<^esub>" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
456 |
unfolding sym[OF univ_poly_carrier] polynomial_def by auto |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
457 |
hence "set (map h p) \<subseteq> carrier S" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
458 |
by (induct p) (auto) |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
459 |
moreover have "h a = \<zero>\<^bsub>S\<^esub> \<Longrightarrow> a = \<zero>\<^bsub>R\<^esub>" if "a \<in> carrier R" for a |
70215
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
paulson <lp15@cam.ac.uk>
parents:
70214
diff
changeset
|
460 |
using non_trivial_field_hom_is_inj[OF assms(1-3)] that unfolding inj_on_def by simp |
70160
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
461 |
with \<open>set p \<subseteq> carrier R\<close> have "lead_coeff (map h p) \<noteq> \<zero>\<^bsub>S\<^esub>" if "p \<noteq> []" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
462 |
using lc[OF that] that by (cases p) (auto) |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
463 |
ultimately show ?thesis |
70215
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
paulson <lp15@cam.ac.uk>
parents:
70214
diff
changeset
|
464 |
unfolding sym[OF univ_poly_carrier] polynomial_def by auto |
70160
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
465 |
qed |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
466 |
|
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
467 |
lemma (in ring_hom_ring) subfield_polynomial_hom: |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
468 |
assumes "subfield K R" and "\<one>\<^bsub>S\<^esub> \<noteq> \<zero>\<^bsub>S\<^esub>" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
469 |
shows "p \<in> carrier (K[X]\<^bsub>R\<^esub>) \<Longrightarrow> (map h p) \<in> carrier ((h ` K)[X]\<^bsub>S\<^esub>)" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
470 |
proof - |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
471 |
assume "p \<in> carrier (K[X]\<^bsub>R\<^esub>)" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
472 |
hence "p \<in> carrier (poly_ring (R \<lparr> carrier := K \<rparr>))" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
473 |
using R.univ_poly_consistent[OF subfieldE(1)[OF assms(1)]] by simp |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
474 |
moreover have "h \<in> ring_hom (R \<lparr> carrier := K \<rparr>) (S \<lparr> carrier := h ` K \<rparr>)" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
475 |
using hom_mult subfieldE(3)[OF assms(1)] unfolding ring_hom_def subset_iff by auto |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
476 |
moreover have "field (R \<lparr> carrier := K \<rparr>)" and "field (S \<lparr> carrier := (h ` K) \<rparr>)" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
477 |
using R.subfield_iff(2)[OF assms(1)] S.subfield_iff(2)[OF img_is_subfield(2)[OF assms]] by simp+ |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
478 |
ultimately have "(map h p) \<in> carrier (poly_ring (S \<lparr> carrier := h ` K \<rparr>))" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
479 |
using polynomial_hom[of h "R \<lparr> carrier := K \<rparr>" "S \<lparr> carrier := h ` K \<rparr>"] by auto |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
480 |
thus ?thesis |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
481 |
using S.univ_poly_consistent[OF subfieldE(1)[OF img_is_subfield(2)[OF assms]]] by simp |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
482 |
qed |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
483 |
|
70212
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
484 |
|
70160
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
485 |
lemma (in field) exists_root: |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
486 |
assumes "M \<in> extensions" and "\<And>L. \<lbrakk> L \<in> extensions; M \<lesssim> L \<rbrakk> \<Longrightarrow> law_restrict L = law_restrict M" |
70212
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
487 |
and "P \<in> carrier (poly_ring R)" |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
488 |
shows "(ring.splitted M) (\<sigma> P)" |
70160
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
489 |
proof (rule ccontr) |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
490 |
from \<open>M \<in> extensions\<close> interpret M: field M + Hom: ring_hom_ring R M "indexed_const" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
491 |
using ring_hom_ringI2[OF ring_axioms field.is_ring] unfolding extensions_def by auto |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
492 |
interpret UP: principal_domain "poly_ring M" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
493 |
using M.univ_poly_is_principal[OF M.carrier_is_subfield] . |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
494 |
|
70212
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
495 |
assume not_splitted: "\<not> (ring.splitted M) (\<sigma> P)" |
70160
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
496 |
have "(\<sigma> P) \<in> carrier (poly_ring M)" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
497 |
using polynomial_hom[OF Hom.homh field_axioms M.field_axioms assms(3)] unfolding \<sigma>_def by simp |
70212
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
498 |
then obtain Q |
70160
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
499 |
where Q: "Q \<in> carrier (poly_ring M)" "pirreducible\<^bsub>M\<^esub> (carrier M) Q" "Q pdivides\<^bsub>M\<^esub> (\<sigma> P)" |
70212
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
500 |
and degree_gt: "degree Q > 1" |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
501 |
using M.trivial_factors_imp_splitted[of "\<sigma> P"] not_splitted by force |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
502 |
|
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
503 |
from \<open>(\<sigma> P) \<in> carrier (poly_ring M)\<close> have "(\<sigma> P) \<noteq> []" |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
504 |
using M.degree_zero_imp_splitted[of "\<sigma> P"] not_splitted unfolding \<sigma>_def by auto |
70160
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
505 |
|
70212
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
506 |
have "\<exists>i. \<forall>\<P> \<in> carrier M. index_free \<P> (P, i)" |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
507 |
proof (rule ccontr) |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
508 |
assume "\<nexists>i. \<forall>\<P> \<in> carrier M. index_free \<P> (P, i)" |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
509 |
then have "\<X>\<^bsub>(P, i)\<^esub> \<in> carrier M" and "(ring.eval M) (\<sigma> P) \<X>\<^bsub>(P, i)\<^esub> = \<zero>\<^bsub>M\<^esub>" for i |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
510 |
using assms(1,3) unfolding extensions_def by blast+ |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
511 |
with \<open>(\<sigma> P) \<noteq> []\<close> have "((\<lambda>i :: nat. \<X>\<^bsub>(P, i)\<^esub>) ` UNIV) \<subseteq> { a. (ring.is_root M) (\<sigma> P) a }" |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
512 |
unfolding M.is_root_def by auto |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
513 |
moreover have "inj (\<lambda>i :: nat. \<X>\<^bsub>(P, i)\<^esub>)" |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
514 |
unfolding indexed_var_def indexed_const_def indexed_pmult_def inj_def |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
515 |
by (metis (no_types, lifting) add_mset_eq_singleton_iff diff_single_eq_union |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
516 |
multi_member_last prod.inject zero_not_one) |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
517 |
hence "infinite ((\<lambda>i :: nat. \<X>\<^bsub>(P, i)\<^esub>) ` UNIV)" |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
518 |
unfolding infinite_iff_countable_subset by auto |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
519 |
ultimately have "infinite { a. (ring.is_root M) (\<sigma> P) a }" |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
520 |
using finite_subset by auto |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
521 |
with \<open>(\<sigma> P) \<in> carrier (poly_ring M)\<close> show False |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
522 |
using M.finite_number_of_roots by simp |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
523 |
qed |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
524 |
then obtain i :: nat where "\<forall>\<P> \<in> carrier M. index_free \<P> (P, i)" |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
525 |
by blast |
70215
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
paulson <lp15@cam.ac.uk>
parents:
70214
diff
changeset
|
526 |
|
70212
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
527 |
then have hyps: |
70160
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
528 |
\<comment> \<open>i\<close> "field M" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
529 |
\<comment> \<open>ii\<close> "\<And>\<P>. \<P> \<in> carrier M \<Longrightarrow> carrier_coeff \<P>" |
70212
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
530 |
\<comment> \<open>iii\<close> "\<And>\<P>. \<P> \<in> carrier M \<Longrightarrow> index_free \<P> (P, i)" |
70160
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
531 |
\<comment> \<open>iv\<close> "\<zero>\<^bsub>M\<^esub> = indexed_const \<zero>" |
70212
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
532 |
using assms(1,3) unfolding extensions_def by auto |
70160
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
533 |
|
70212
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
534 |
define image_poly where "image_poly = image_ring (eval_pmod M (P, i) Q) (poly_ring M)" |
70160
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
535 |
with \<open>degree Q > 1\<close> have "M \<lesssim> image_poly" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
536 |
using image_poly_iso_incl[OF hyps Q(1)] by auto |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
537 |
moreover have is_field: "field image_poly" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
538 |
using image_poly_is_field[OF hyps Q(1-2)] unfolding image_poly_def by simp |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
539 |
moreover have "image_poly \<in> extensions" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
540 |
proof (auto simp add: extensions_def is_field) |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
541 |
fix \<P> assume "\<P> \<in> carrier image_poly" |
70212
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
542 |
then obtain R where \<P>: "\<P> = eval_pmod M (P, i) Q R" and "R \<in> carrier (poly_ring M)" |
70160
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
543 |
unfolding image_poly_def image_ring_carrier by auto |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
544 |
hence "M.pmod R Q \<in> carrier (poly_ring M)" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
545 |
using M.long_division_closed(2)[OF M.carrier_is_subfield _ Q(1)] by simp |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
546 |
hence "list_all carrier_coeff (M.pmod R Q)" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
547 |
using hyps(2) unfolding sym[OF univ_poly_carrier] list_all_iff polynomial_def by auto |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
548 |
thus "carrier_coeff \<P>" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
549 |
using indexed_eval_in_carrier[of "M.pmod R Q"] unfolding \<P> by simp |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
550 |
next |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
551 |
from \<open>M \<lesssim> image_poly\<close> show "indexed_const \<in> ring_hom R image_poly" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
552 |
using ring_hom_trans[OF Hom.homh, of id] unfolding iso_incl.simps by simp |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
553 |
next |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
554 |
from \<open>M \<lesssim> image_poly\<close> interpret Id: ring_hom_ring M image_poly id |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
555 |
using iso_inclE[OF M.ring_axioms field.is_ring[OF is_field]] by simp |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
556 |
|
70212
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
557 |
fix \<P> S j |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
558 |
assume A: "\<P> \<in> carrier image_poly" "\<not> index_free \<P> (S, j)" "S \<in> carrier (poly_ring R)" |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
559 |
have "\<X>\<^bsub>(S, j)\<^esub> \<in> carrier image_poly \<and> Id.eval (\<sigma> S) \<X>\<^bsub>(S, j)\<^esub> = \<zero>\<^bsub>image_poly\<^esub>" |
70160
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
560 |
proof (cases) |
70212
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
561 |
assume "(P, i) \<noteq> (S, j)" |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
562 |
then obtain Q' where "Q' \<in> carrier M" and "\<not> index_free Q' (S, j)" |
70160
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
563 |
using A(1) image_poly_index_free[OF hyps Q(1) _ A(2)] unfolding image_poly_def by auto |
70212
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
564 |
hence "\<X>\<^bsub>(S, j)\<^esub> \<in> carrier M" and "M.eval (\<sigma> S) \<X>\<^bsub>(S, j)\<^esub> = \<zero>\<^bsub>M\<^esub>" |
70160
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
565 |
using assms(1) A(3) unfolding extensions_def by auto |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
566 |
moreover have "\<sigma> S \<in> carrier (poly_ring M)" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
567 |
using polynomial_hom[OF Hom.homh field_axioms M.field_axioms A(3)] unfolding \<sigma>_def . |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
568 |
ultimately show ?thesis |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
569 |
using Id.eval_hom[OF M.carrier_is_subring] Id.hom_closed Id.hom_zero by auto |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
570 |
next |
70212
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
571 |
assume "\<not> (P, i) \<noteq> (S, j)" hence S: "(P, i) = (S, j)" |
70160
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
572 |
by simp |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
573 |
have poly_hom: "R \<in> carrier (poly_ring image_poly)" if "R \<in> carrier (poly_ring M)" for R |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
574 |
using polynomial_hom[OF Id.homh M.field_axioms is_field that] by simp |
70212
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
575 |
have "\<X>\<^bsub>(S, j)\<^esub> \<in> carrier image_poly" |
70160
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
576 |
using eval_pmod_var(2)[OF hyps Hom.homh Q(1) degree_gt] unfolding image_poly_def S by simp |
70212
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
577 |
moreover have "Id.eval Q \<X>\<^bsub>(S, j)\<^esub> = \<zero>\<^bsub>image_poly\<^esub>" |
70160
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
578 |
using image_poly_eval_indexed_var[OF hyps Hom.homh Q(1) degree_gt Q(2)] unfolding image_poly_def S by simp |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
579 |
moreover have "Q pdivides\<^bsub>image_poly\<^esub> (\<sigma> S)" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
580 |
proof - |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
581 |
obtain R where R: "R \<in> carrier (poly_ring M)" "\<sigma> S = Q \<otimes>\<^bsub>poly_ring M\<^esub> R" |
70212
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
582 |
using Q(3) S unfolding pdivides_def by auto |
70160
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
583 |
moreover have "set Q \<subseteq> carrier M" and "set R \<subseteq> carrier M" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
584 |
using Q(1) R(1) unfolding sym[OF univ_poly_carrier] polynomial_def by auto |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
585 |
ultimately have "Id.normalize (\<sigma> S) = Q \<otimes>\<^bsub>poly_ring image_poly\<^esub> R" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
586 |
using Id.poly_mult_hom'[of Q R] unfolding univ_poly_mult by simp |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
587 |
moreover have "\<sigma> S \<in> carrier (poly_ring M)" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
588 |
using polynomial_hom[OF Hom.homh field_axioms M.field_axioms A(3)] unfolding \<sigma>_def . |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
589 |
hence "\<sigma> S \<in> carrier (poly_ring image_poly)" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
590 |
using polynomial_hom[OF Id.homh M.field_axioms is_field] by simp |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
591 |
hence "Id.normalize (\<sigma> S) = \<sigma> S" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
592 |
using Id.normalize_polynomial unfolding sym[OF univ_poly_carrier] by simp |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
593 |
ultimately show ?thesis |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
594 |
using poly_hom[OF Q(1)] poly_hom[OF R(1)] |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
595 |
unfolding pdivides_def factor_def univ_poly_mult by auto |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
596 |
qed |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
597 |
moreover have "Q \<in> carrier (poly_ring (image_poly))" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
598 |
using poly_hom[OF Q(1)] by simp |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
599 |
ultimately show ?thesis |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
600 |
using domain.pdivides_imp_root_sharing[OF field.axioms(1)[OF is_field], of Q] by auto |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
601 |
qed |
70212
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
602 |
thus "\<X>\<^bsub>(S, j)\<^esub> \<in> carrier image_poly" and "Id.eval (\<sigma> S) \<X>\<^bsub>(S, j)\<^esub> = \<zero>\<^bsub>image_poly\<^esub>" |
70160
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
603 |
by auto |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
604 |
qed |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
605 |
ultimately have "law_restrict M = law_restrict image_poly" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
606 |
using assms(2) by simp |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
607 |
hence "carrier M = carrier image_poly" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
608 |
unfolding law_restrict_def by (simp add:ring.defs) |
70212
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
609 |
moreover have "\<X>\<^bsub>(P, i)\<^esub> \<in> carrier image_poly" |
70160
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
610 |
using eval_pmod_var(2)[OF hyps Hom.homh Q(1) degree_gt] unfolding image_poly_def by simp |
70212
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
611 |
moreover have "\<X>\<^bsub>(P, i)\<^esub> \<notin> carrier M" |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
612 |
using indexed_var_not_index_free[of "(P, i)"] hyps(3) by blast |
70160
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
613 |
ultimately show False by simp |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
614 |
qed |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
615 |
|
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
616 |
lemma (in field) exists_extension_with_roots: |
70212
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
617 |
shows "\<exists>L \<in> extensions. \<forall>P \<in> carrier (poly_ring R). (ring.splitted L) (\<sigma> P)" |
70160
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
618 |
proof - |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
619 |
obtain M where "M \<in> extensions" and "\<forall>L \<in> extensions. M \<lesssim> L \<longrightarrow> law_restrict L = law_restrict M" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
620 |
using exists_maximal_extension iso_incl_hom by blast |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
621 |
thus ?thesis |
70215
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
paulson <lp15@cam.ac.uk>
parents:
70214
diff
changeset
|
622 |
using exists_root[of M] by auto |
70160
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
623 |
qed |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
624 |
|
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
625 |
|
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
626 |
subsection \<open>Existence of Algebraic Closure\<close> |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
627 |
|
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
628 |
locale algebraic_closure = field L + subfield K L for L (structure) and K + |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
629 |
assumes algebraic_extension: "x \<in> carrier L \<Longrightarrow> (algebraic over K) x" |
70212
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
630 |
and roots_over_subfield: "P \<in> carrier (K[X]) \<Longrightarrow> splitted P" |
70160
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
631 |
|
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
632 |
locale algebraically_closed = field L for L (structure) + |
70212
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
633 |
assumes roots_over_carrier: "P \<in> carrier (poly_ring L) \<Longrightarrow> splitted P" |
70160
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
634 |
|
70215
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
paulson <lp15@cam.ac.uk>
parents:
70214
diff
changeset
|
635 |
definition (in field) alg_closure :: "(('a list \<times> nat) multiset \<Rightarrow> 'a) ring" |
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
paulson <lp15@cam.ac.uk>
parents:
70214
diff
changeset
|
636 |
where "alg_closure = (SOME L \<comment> \<open>such that\<close>. |
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
paulson <lp15@cam.ac.uk>
parents:
70214
diff
changeset
|
637 |
\<comment> \<open>i\<close> algebraic_closure L (indexed_const ` (carrier R)) \<and> |
70160
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
638 |
\<comment> \<open>ii\<close> indexed_const \<in> ring_hom R L)" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
639 |
|
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
640 |
lemma algebraic_hom: |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
641 |
assumes "h \<in> ring_hom R S" and "field R" and "field S" and "subfield K R" and "x \<in> carrier R" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
642 |
shows "((ring.algebraic R) over K) x \<Longrightarrow> ((ring.algebraic S) over (h ` K)) (h x)" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
643 |
proof - |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
644 |
interpret Hom: ring_hom_ring R S h |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
645 |
using ring_hom_ringI2[OF assms(2-3)[THEN field.is_ring] assms(1)] . |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
646 |
assume "(Hom.R.algebraic over K) x" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
647 |
then obtain p where p: "p \<in> carrier (K[X]\<^bsub>R\<^esub>)" and "p \<noteq> []" and eval: "Hom.R.eval p x = \<zero>\<^bsub>R\<^esub>" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
648 |
using domain.algebraicE[OF field.axioms(1) subfieldE(1), of R K x] assms(2,4-5) by auto |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
649 |
hence "(map h p) \<in> carrier ((h ` K)[X]\<^bsub>S\<^esub>)" and "(map h p) \<noteq> []" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
650 |
using Hom.subfield_polynomial_hom[OF assms(4) one_not_zero[OF assms(3)]] by auto |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
651 |
moreover have "Hom.S.eval (map h p) (h x) = \<zero>\<^bsub>S\<^esub>" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
652 |
using Hom.eval_hom[OF subfieldE(1)[OF assms(4)] assms(5) p] unfolding eval by simp |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
653 |
ultimately show ?thesis |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
654 |
using Hom.S.non_trivial_ker_imp_algebraic[of "h ` K" "h x"] unfolding a_kernel_def' by auto |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
655 |
qed |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
656 |
|
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
657 |
lemma (in field) exists_closure: |
70212
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
658 |
obtains L :: "((('a list \<times> nat) multiset) \<Rightarrow> 'a) ring" |
70160
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
659 |
where "algebraic_closure L (indexed_const ` (carrier R))" and "indexed_const \<in> ring_hom R L" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
660 |
proof - |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
661 |
obtain L where "L \<in> extensions" |
70212
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
662 |
and roots: "\<And>P. P \<in> carrier (poly_ring R) \<Longrightarrow> (ring.splitted L) (\<sigma> P)" |
70160
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
663 |
using exists_extension_with_roots by auto |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
664 |
|
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
665 |
let ?K = "indexed_const ` (carrier R)" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
666 |
let ?set_of_algs = "{ x \<in> carrier L. ((ring.algebraic L) over ?K) x }" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
667 |
let ?M = "L \<lparr> carrier := ?set_of_algs \<rparr>" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
668 |
|
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
669 |
from \<open>L \<in> extensions\<close> |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
670 |
have L: "field L" and hom: "ring_hom_ring R L indexed_const" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
671 |
using ring_hom_ringI2[OF ring_axioms field.is_ring] unfolding extensions_def by auto |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
672 |
have "subfield ?K L" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
673 |
using ring_hom_ring.img_is_subfield(2)[OF hom carrier_is_subfield |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
674 |
domain.one_not_zero[OF field.axioms(1)[OF L]]] by auto |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
675 |
hence set_of_algs: "subfield ?set_of_algs L" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
676 |
using field.subfield_of_algebraics[OF L, of ?K] by simp |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
677 |
have M: "field ?M" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
678 |
using ring.subfield_iff(2)[OF field.is_ring[OF L] set_of_algs] by simp |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
679 |
|
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
680 |
interpret Id: ring_hom_ring ?M L id |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
681 |
using ring_hom_ringI[OF field.is_ring[OF M] field.is_ring[OF L]] by auto |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
682 |
|
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
683 |
have is_subfield: "subfield ?K ?M" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
684 |
proof (intro ring.subfield_iff(1)[OF field.is_ring[OF M]]) |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
685 |
have "L \<lparr> carrier := ?K \<rparr> = ?M \<lparr> carrier := ?K \<rparr>" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
686 |
by simp |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
687 |
moreover from \<open>subfield ?K L\<close> have "field (L \<lparr> carrier := ?K \<rparr>)" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
688 |
using ring.subfield_iff(2)[OF field.is_ring[OF L]] by simp |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
689 |
ultimately show "field (?M \<lparr> carrier := ?K \<rparr>)" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
690 |
by simp |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
691 |
next |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
692 |
show "?K \<subseteq> carrier ?M" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
693 |
proof |
70212
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
694 |
fix x :: "(('a list \<times> nat) multiset) \<Rightarrow> 'a" |
70160
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
695 |
assume "x \<in> ?K" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
696 |
hence "x \<in> carrier L" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
697 |
using ring_hom_memE(1)[OF ring_hom_ring.homh[OF hom]] by auto |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
698 |
moreover from \<open>subfield ?K L\<close> and \<open>x \<in> ?K\<close> have "(Id.S.algebraic over ?K) x" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
699 |
using domain.algebraic_self[OF field.axioms(1)[OF L] subfieldE(1)] by auto |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
700 |
ultimately show "x \<in> carrier ?M" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
701 |
by auto |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
702 |
qed |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
703 |
qed |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
704 |
|
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
705 |
have "algebraic_closure ?M ?K" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
706 |
proof (intro algebraic_closure.intro[OF M is_subfield]) |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
707 |
have "(Id.R.algebraic over ?K) x" if "x \<in> carrier ?M" for x |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
708 |
using that Id.S.algebraic_consistent[OF subfieldE(1)[OF set_of_algs]] by simp |
70212
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
709 |
moreover have "Id.R.splitted P" if "P \<in> carrier (?K[X]\<^bsub>?M\<^esub>)" for P |
70160
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
710 |
proof - |
70212
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
711 |
from \<open>P \<in> carrier (?K[X]\<^bsub>?M\<^esub>)\<close> have "P \<in> carrier (poly_ring ?M)" |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
712 |
using Id.R.carrier_polynomial_shell[OF subfieldE(1)[OF is_subfield]] by simp |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
713 |
show ?thesis |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
714 |
proof (cases "degree P = 0") |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
715 |
case True with \<open>P \<in> carrier (poly_ring ?M)\<close> show ?thesis |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
716 |
using domain.degree_zero_imp_splitted[OF field.axioms(1)[OF M]] |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
717 |
by fastforce |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
718 |
next |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
719 |
case False then have "degree P > 0" |
70160
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
720 |
by simp |
70212
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
721 |
from \<open>P \<in> carrier (?K[X]\<^bsub>?M\<^esub>)\<close> have "P \<in> carrier (?K[X]\<^bsub>L\<^esub>)" |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
722 |
unfolding Id.S.univ_poly_consistent[OF subfieldE(1)[OF set_of_algs]] . |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
723 |
hence "set P \<subseteq> ?K" |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
724 |
unfolding sym[OF univ_poly_carrier] polynomial_def by auto |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
725 |
hence "\<exists>Q. set Q \<subseteq> carrier R \<and> P = \<sigma> Q" |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
726 |
proof (induct P, simp add: \<sigma>_def) |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
727 |
case (Cons p P) |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
728 |
then obtain q Q where "q \<in> carrier R" "set Q \<subseteq> carrier R" |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
729 |
and "\<sigma> Q = P" "indexed_const q = p" |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
730 |
unfolding \<sigma>_def by auto |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
731 |
hence "set (q # Q) \<subseteq> carrier R" and "\<sigma> (q # Q) = (p # P)" |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
732 |
unfolding \<sigma>_def by auto |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
733 |
thus ?case |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
734 |
by metis |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
735 |
qed |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
736 |
then obtain Q where "set Q \<subseteq> carrier R" and "\<sigma> Q = P" |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
737 |
by auto |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
738 |
moreover have "lead_coeff Q \<noteq> \<zero>" |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
739 |
proof (rule ccontr) |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
740 |
assume "\<not> lead_coeff Q \<noteq> \<zero>" then have "lead_coeff Q = \<zero>" |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
741 |
by simp |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
742 |
with \<open>\<sigma> Q = P\<close> and \<open>degree P > 0\<close> have "lead_coeff P = indexed_const \<zero>" |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
743 |
unfolding \<sigma>_def by (metis diff_0_eq_0 length_map less_irrefl_nat list.map_sel(1) list.size(3)) |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
744 |
hence "lead_coeff P = \<zero>\<^bsub>L\<^esub>" |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
745 |
using ring_hom_zero[OF ring_hom_ring.homh ring_hom_ring.axioms(1-2)] hom by auto |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
746 |
with \<open>degree P > 0\<close> have "\<not> P \<in> carrier (?K[X]\<^bsub>?M\<^esub>)" |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
747 |
unfolding sym[OF univ_poly_carrier] polynomial_def by auto |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
748 |
with \<open>P \<in> carrier (?K[X]\<^bsub>?M\<^esub>)\<close> show False |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
749 |
by simp |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
750 |
qed |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
751 |
ultimately have "Q \<in> carrier (poly_ring R)" |
70160
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
752 |
unfolding sym[OF univ_poly_carrier] polynomial_def by auto |
70212
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
753 |
with \<open>\<sigma> Q = P\<close> have "Id.S.splitted P" |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
754 |
using roots[of Q] by simp |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
755 |
|
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
756 |
from \<open>P \<in> carrier (poly_ring ?M)\<close> show ?thesis |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
757 |
proof (rule field.trivial_factors_imp_splitted[OF M]) |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
758 |
fix R |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
759 |
assume R: "R \<in> carrier (poly_ring ?M)" "pirreducible\<^bsub>?M\<^esub> (carrier ?M) R" and "R pdivides\<^bsub>?M\<^esub> P" |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
760 |
|
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
761 |
from \<open>P \<in> carrier (poly_ring ?M)\<close> and \<open>R \<in> carrier (poly_ring ?M)\<close> |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
762 |
have "P \<in> carrier ((?set_of_algs)[X]\<^bsub>L\<^esub>)" and "R \<in> carrier ((?set_of_algs)[X]\<^bsub>L\<^esub>)" |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
763 |
unfolding Id.S.univ_poly_consistent[OF subfieldE(1)[OF set_of_algs]] by auto |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
764 |
hence in_carrier: "P \<in> carrier (poly_ring L)" "R \<in> carrier (poly_ring L)" |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
765 |
using Id.S.carrier_polynomial_shell[OF subfieldE(1)[OF set_of_algs]] by auto |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
766 |
|
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
767 |
from \<open>R pdivides\<^bsub>?M\<^esub> P\<close> have "R divides\<^bsub>((?set_of_algs)[X]\<^bsub>L\<^esub>)\<^esub> P" |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
768 |
unfolding pdivides_def Id.S.univ_poly_consistent[OF subfieldE(1)[OF set_of_algs]] |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
769 |
by simp |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
770 |
with \<open>P \<in> carrier ((?set_of_algs)[X]\<^bsub>L\<^esub>)\<close> and \<open>R \<in> carrier ((?set_of_algs)[X]\<^bsub>L\<^esub>)\<close> |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
771 |
have "R pdivides\<^bsub>L\<^esub> P" |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
772 |
using domain.pdivides_iff_shell[OF field.axioms(1)[OF L] set_of_algs, of R P] by simp |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
773 |
with \<open>Id.S.splitted P\<close> and \<open>degree P \<noteq> 0\<close> have "Id.S.splitted R" |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
774 |
using field.pdivides_imp_splitted[OF L in_carrier(2,1)] by fastforce |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
775 |
show "degree R \<le> 1" |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
776 |
proof (cases "Id.S.roots R = {#}") |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
777 |
case True with \<open>Id.S.splitted R\<close> show ?thesis |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
778 |
unfolding Id.S.splitted_def by simp |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
779 |
next |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
780 |
case False with \<open>R \<in> carrier (poly_ring L)\<close> |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
781 |
obtain a where "a \<in> carrier L" and "a \<in># Id.S.roots R" |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
782 |
and "[ \<one>\<^bsub>L\<^esub>, \<ominus>\<^bsub>L\<^esub> a ] \<in> carrier (poly_ring L)" and pdiv: "[ \<one>\<^bsub>L\<^esub>, \<ominus>\<^bsub>L\<^esub> a ] pdivides\<^bsub>L\<^esub> R" |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
783 |
using domain.not_empty_rootsE[OF field.axioms(1)[OF L], of R] by blast |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
784 |
|
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
785 |
from \<open>P \<in> carrier (?K[X]\<^bsub>L\<^esub>)\<close> |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
786 |
have "(Id.S.algebraic over ?K) a" |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
787 |
proof (rule Id.S.algebraicI) |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
788 |
from \<open>degree P \<noteq> 0\<close> show "P \<noteq> []" |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
789 |
by auto |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
790 |
next |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
791 |
from \<open>a \<in># Id.S.roots R\<close> and \<open>R \<in> carrier (poly_ring L)\<close> |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
792 |
have "Id.S.eval R a = \<zero>\<^bsub>L\<^esub>" |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
793 |
using domain.roots_mem_iff_is_root[OF field.axioms(1)[OF L]] |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
794 |
unfolding Id.S.is_root_def by auto |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
795 |
with \<open>R pdivides\<^bsub>L\<^esub> P\<close> and \<open>a \<in> carrier L\<close> show "Id.S.eval P a = \<zero>\<^bsub>L\<^esub>" |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
796 |
using domain.pdivides_imp_root_sharing[OF field.axioms(1)[OF L] in_carrier(2)] by simp |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
797 |
qed |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
798 |
with \<open>a \<in> carrier L\<close> have "a \<in> ?set_of_algs" |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
799 |
by simp |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
800 |
hence "[ \<one>\<^bsub>L\<^esub>, \<ominus>\<^bsub>L\<^esub> a ] \<in> carrier ((?set_of_algs)[X]\<^bsub>L\<^esub>)" |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
801 |
using subringE(3,5)[of ?set_of_algs L] subfieldE(1,6)[OF set_of_algs] |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
802 |
unfolding sym[OF univ_poly_carrier] polynomial_def by simp |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
803 |
hence "[ \<one>\<^bsub>L\<^esub>, \<ominus>\<^bsub>L\<^esub> a ] \<in> carrier (poly_ring ?M)" |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
804 |
unfolding Id.S.univ_poly_consistent[OF subfieldE(1)[OF set_of_algs]] by simp |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
805 |
|
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
806 |
from \<open>[ \<one>\<^bsub>L\<^esub>, \<ominus>\<^bsub>L\<^esub> a ] \<in> carrier ((?set_of_algs)[X]\<^bsub>L\<^esub>)\<close> |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
807 |
and \<open>R \<in> carrier ((?set_of_algs)[X]\<^bsub>L\<^esub>)\<close> |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
808 |
have "[ \<one>\<^bsub>L\<^esub>, \<ominus>\<^bsub>L\<^esub> a ] divides\<^bsub>(?set_of_algs)[X]\<^bsub>L\<^esub>\<^esub> R" |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
809 |
using pdiv domain.pdivides_iff_shell[OF field.axioms(1)[OF L] set_of_algs] by simp |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
810 |
hence "[ \<one>\<^bsub>L\<^esub>, \<ominus>\<^bsub>L\<^esub> a ] divides\<^bsub>poly_ring ?M\<^esub> R" |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
811 |
unfolding pdivides_def Id.S.univ_poly_consistent[OF subfieldE(1)[OF set_of_algs]] |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
812 |
by simp |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
813 |
|
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
814 |
have "[ \<one>\<^bsub>L\<^esub>, \<ominus>\<^bsub>L\<^esub> a ] \<notin> Units (poly_ring ?M)" |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
815 |
using Id.R.univ_poly_units[OF field.carrier_is_subfield[OF M]] by force |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
816 |
with \<open>[ \<one>\<^bsub>L\<^esub>, \<ominus>\<^bsub>L\<^esub> a ] \<in> carrier (poly_ring ?M)\<close> and \<open>R \<in> carrier (poly_ring ?M)\<close> |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
817 |
and \<open>[ \<one>\<^bsub>L\<^esub>, \<ominus>\<^bsub>L\<^esub> a ] divides\<^bsub>poly_ring ?M\<^esub> R\<close> |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
818 |
have "[ \<one>\<^bsub>L\<^esub>, \<ominus>\<^bsub>L\<^esub> a ] \<sim>\<^bsub>poly_ring ?M\<^esub> R" |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
819 |
using Id.R.divides_pirreducible_condition[OF R(2)] by auto |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
820 |
with \<open>[ \<one>\<^bsub>L\<^esub>, \<ominus>\<^bsub>L\<^esub> a ] \<in> carrier (poly_ring ?M)\<close> and \<open>R \<in> carrier (poly_ring ?M)\<close> |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
821 |
have "degree R = 1" |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
822 |
using domain.associated_polynomials_imp_same_length[OF field.axioms(1)[OF M] |
70215
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
paulson <lp15@cam.ac.uk>
parents:
70214
diff
changeset
|
823 |
Id.R.carrier_is_subring, of "[ \<one>\<^bsub>L\<^esub>, \<ominus>\<^bsub>L\<^esub> a ]" R] by force |
70212
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
824 |
thus ?thesis |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
825 |
by simp |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
826 |
qed |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
827 |
qed |
70160
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
828 |
qed |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
829 |
qed |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
830 |
ultimately show "algebraic_closure_axioms ?M ?K" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
831 |
unfolding algebraic_closure_axioms_def by auto |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
832 |
qed |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
833 |
moreover have "indexed_const \<in> ring_hom R ?M" |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
834 |
using ring_hom_ring.homh[OF hom] subfieldE(3)[OF is_subfield] |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
835 |
unfolding subset_iff ring_hom_def by auto |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
836 |
ultimately show thesis |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
837 |
using that by auto |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
838 |
qed |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
839 |
|
70215
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
paulson <lp15@cam.ac.uk>
parents:
70214
diff
changeset
|
840 |
lemma (in field) alg_closureE: |
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
paulson <lp15@cam.ac.uk>
parents:
70214
diff
changeset
|
841 |
shows "algebraic_closure alg_closure (indexed_const ` (carrier R))" |
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
paulson <lp15@cam.ac.uk>
parents:
70214
diff
changeset
|
842 |
and "indexed_const \<in> ring_hom R alg_closure" |
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
paulson <lp15@cam.ac.uk>
parents:
70214
diff
changeset
|
843 |
using exists_closure unfolding alg_closure_def |
70160
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
844 |
by (metis (mono_tags, lifting) someI2)+ |
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
845 |
|
70215
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
paulson <lp15@cam.ac.uk>
parents:
70214
diff
changeset
|
846 |
lemma (in field) algebraically_closedI': |
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
paulson <lp15@cam.ac.uk>
parents:
70214
diff
changeset
|
847 |
assumes "\<And>p. \<lbrakk> p \<in> carrier (poly_ring R); degree p > 1 \<rbrakk> \<Longrightarrow> splitted p" |
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
paulson <lp15@cam.ac.uk>
parents:
70214
diff
changeset
|
848 |
shows "algebraically_closed R" |
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
paulson <lp15@cam.ac.uk>
parents:
70214
diff
changeset
|
849 |
proof |
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
paulson <lp15@cam.ac.uk>
parents:
70214
diff
changeset
|
850 |
fix p assume "p \<in> carrier (poly_ring R)" show "splitted p" |
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
paulson <lp15@cam.ac.uk>
parents:
70214
diff
changeset
|
851 |
proof (cases "degree p \<le> 1") |
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
paulson <lp15@cam.ac.uk>
parents:
70214
diff
changeset
|
852 |
case True with \<open>p \<in> carrier (poly_ring R)\<close> show ?thesis |
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
paulson <lp15@cam.ac.uk>
parents:
70214
diff
changeset
|
853 |
using degree_zero_imp_splitted degree_one_imp_splitted by fastforce |
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
paulson <lp15@cam.ac.uk>
parents:
70214
diff
changeset
|
854 |
next |
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
paulson <lp15@cam.ac.uk>
parents:
70214
diff
changeset
|
855 |
case False with \<open>p \<in> carrier (poly_ring R)\<close> show ?thesis |
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
paulson <lp15@cam.ac.uk>
parents:
70214
diff
changeset
|
856 |
using assms by fastforce |
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
paulson <lp15@cam.ac.uk>
parents:
70214
diff
changeset
|
857 |
qed |
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
paulson <lp15@cam.ac.uk>
parents:
70214
diff
changeset
|
858 |
qed |
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
paulson <lp15@cam.ac.uk>
parents:
70214
diff
changeset
|
859 |
|
70212
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
860 |
lemma (in field) algebraically_closedI: |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
861 |
assumes "\<And>p. \<lbrakk> p \<in> carrier (poly_ring R); degree p > 1 \<rbrakk> \<Longrightarrow> \<exists>x \<in> carrier R. eval p x = \<zero>" |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
862 |
shows "algebraically_closed R" |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
863 |
proof |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
864 |
fix p assume "p \<in> carrier (poly_ring R)" thus "splitted p" |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
865 |
proof (induction "degree p" arbitrary: p rule: less_induct) |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
866 |
case less show ?case |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
867 |
proof (cases "degree p \<le> 1") |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
868 |
case True with \<open>p \<in> carrier (poly_ring R)\<close> show ?thesis |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
869 |
using degree_zero_imp_splitted degree_one_imp_splitted by fastforce |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
870 |
next |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
871 |
case False then have "degree p > 1" |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
872 |
by simp |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
873 |
with \<open>p \<in> carrier (poly_ring R)\<close> have "roots p \<noteq> {#}" |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
874 |
using assms[of p] roots_mem_iff_is_root[of p] unfolding is_root_def by force |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
875 |
then obtain a where a: "a \<in> carrier R" "a \<in># roots p" |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
876 |
and pdiv: "[ \<one>, \<ominus> a ] pdivides p" and in_carrier: "[ \<one>, \<ominus> a ] \<in> carrier (poly_ring R)" |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
877 |
using less(2) by blast |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
878 |
then obtain q where q: "q \<in> carrier (poly_ring R)" and p: "p = [ \<one>, \<ominus> a ] \<otimes>\<^bsub>poly_ring R\<^esub> q" |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
879 |
unfolding pdivides_def by blast |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
880 |
with \<open>degree p > 1\<close> have not_zero: "q \<noteq> []" and "p \<noteq> []" |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
881 |
using domain.integral_iff[OF univ_poly_is_domain[OF carrier_is_subring] in_carrier, of q] |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
882 |
by (auto simp add: univ_poly_zero[of R "carrier R"]) |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
883 |
hence deg: "degree p = Suc (degree q)" |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
884 |
using poly_mult_degree_eq[OF carrier_is_subring] in_carrier q p |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
885 |
unfolding univ_poly_carrier sym[OF univ_poly_mult[of R "carrier R"]] by auto |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
886 |
hence "splitted q" |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
887 |
using less(1)[OF _ q] by simp |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
888 |
moreover have "roots p = add_mset a (roots q)" |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
889 |
using poly_mult_degree_one_monic_imp_same_roots[OF a(1) q not_zero] p by simp |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
890 |
ultimately show ?thesis |
70215
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
paulson <lp15@cam.ac.uk>
parents:
70214
diff
changeset
|
891 |
unfolding splitted_def deg by simp |
70212
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
892 |
qed |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
893 |
qed |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
894 |
qed |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
895 |
|
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
896 |
sublocale algebraic_closure \<subseteq> algebraically_closed |
70215
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
paulson <lp15@cam.ac.uk>
parents:
70214
diff
changeset
|
897 |
proof (rule algebraically_closedI') |
70212
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
898 |
fix P assume in_carrier: "P \<in> carrier (poly_ring L)" and gt_one: "degree P > 1" |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
899 |
then have gt_zero: "degree P > 0" |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
900 |
by simp |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
901 |
|
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
902 |
define A where "A = finite_extension K P" |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
903 |
|
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
904 |
from \<open>P \<in> carrier (poly_ring L)\<close> have "set P \<subseteq> carrier L" |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
905 |
by (simp add: polynomial_incl univ_poly_carrier) |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
906 |
hence A: "subfield A L" and P: "P \<in> carrier (A[X])" |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
907 |
using finite_extension_mem[OF subfieldE(1)[OF subfield_axioms], of P] in_carrier |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
908 |
algebraic_extension finite_extension_is_subfield[OF subfield_axioms, of P] |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
909 |
unfolding sym[OF A_def] sym[OF univ_poly_carrier] polynomial_def by auto |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
910 |
from \<open>set P \<subseteq> carrier L\<close> have incl: "K \<subseteq> A" |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
911 |
using finite_extension_incl[OF subfieldE(3)[OF subfield_axioms]] unfolding A_def by simp |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
912 |
|
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
913 |
interpret UP_K: domain "K[X]" |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
914 |
using univ_poly_is_domain[OF subfieldE(1)[OF subfield_axioms]] . |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
915 |
interpret UP_A: domain "A[X]" |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
916 |
using univ_poly_is_domain[OF subfieldE(1)[OF A]] . |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
917 |
interpret Rupt: ring "Rupt A P" |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
918 |
unfolding rupture_def using ideal.quotient_is_ring[OF UP_A.cgenideal_ideal[OF P]] . |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
919 |
interpret Hom: ring_hom_ring "L \<lparr> carrier := A \<rparr>" "Rupt A P" "rupture_surj A P \<circ> poly_of_const" |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
920 |
using ring_hom_ringI2[OF subring_is_ring[OF subfieldE(1)] Rupt.ring_axioms |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
921 |
rupture_surj_norm_is_hom[OF subfieldE(1) P]] A by simp |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
922 |
let ?h = "rupture_surj A P \<circ> poly_of_const" |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
923 |
|
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
924 |
have h_simp: "rupture_surj A P ` poly_of_const ` E = ?h ` E" for E |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
925 |
by auto |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
926 |
hence aux_lemmas: |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
927 |
"subfield (rupture_surj A P ` poly_of_const ` K) (Rupt A P)" |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
928 |
"subfield (rupture_surj A P ` poly_of_const ` A) (Rupt A P)" |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
929 |
using Hom.img_is_subfield(2)[OF _ rupture_one_not_zero[OF A P gt_zero]] |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
930 |
ring.subfield_iff(1)[OF subring_is_ring[OF subfieldE(1)[OF A]]] |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
931 |
subfield_iff(2)[OF subfield_axioms] subfield_iff(2)[OF A] incl |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
932 |
by auto |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
933 |
|
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
934 |
have "carrier (K[X]) \<subseteq> carrier (A[X])" |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
935 |
using subsetI[of "carrier (K[X])" "carrier (A[X])"] incl |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
936 |
unfolding sym[OF univ_poly_carrier] polynomial_def by auto |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
937 |
hence "id \<in> ring_hom (K[X]) (A[X])" |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
938 |
unfolding ring_hom_def unfolding univ_poly_mult univ_poly_add univ_poly_one by (simp add: subsetD) |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
939 |
hence "rupture_surj A P \<in> ring_hom (K[X]) (Rupt A P)" |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
940 |
using ring_hom_trans[OF _ rupture_surj_hom(1)[OF subfieldE(1)[OF A] P], of id] by simp |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
941 |
then interpret Hom': ring_hom_ring "K[X]" "Rupt A P" "rupture_surj A P" |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
942 |
using ring_hom_ringI2[OF UP_K.ring_axioms Rupt.ring_axioms] by simp |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
943 |
|
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
944 |
from \<open>id \<in> ring_hom (K[X]) (A[X])\<close> have Id: "ring_hom_ring (K[X]) (A[X]) id" |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
945 |
using ring_hom_ringI2[OF UP_K.ring_axioms UP_A.ring_axioms] by simp |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
946 |
hence "subalgebra (poly_of_const ` K) (carrier (K[X])) (A[X])" |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
947 |
using ring_hom_ring.img_is_subalgebra[OF Id _ UP_K.carrier_is_subalgebra[OF subfieldE(3)]] |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
948 |
univ_poly_subfield_of_consts[OF subfield_axioms] by auto |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
949 |
|
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
950 |
moreover from \<open>carrier (K[X]) \<subseteq> carrier (A[X])\<close> have "poly_of_const ` K \<subseteq> carrier (A[X])" |
70215
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
paulson <lp15@cam.ac.uk>
parents:
70214
diff
changeset
|
951 |
using subfieldE(3)[OF univ_poly_subfield_of_consts[OF subfield_axioms]] by simp |
70212
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
952 |
|
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
953 |
ultimately |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
954 |
have "subalgebra (rupture_surj A P ` poly_of_const ` K) (rupture_surj A P ` carrier (K[X])) (Rupt A P)" |
70215
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
paulson <lp15@cam.ac.uk>
parents:
70214
diff
changeset
|
955 |
using ring_hom_ring.img_is_subalgebra[OF rupture_surj_hom(2)[OF subfieldE(1)[OF A] P]] by simp |
70212
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
956 |
|
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
957 |
moreover have "Rupt.finite_dimension (rupture_surj A P ` poly_of_const ` K) (carrier (Rupt A P))" |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
958 |
proof (intro Rupt.telescopic_base_dim(1)[where |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
959 |
?K = "rupture_surj A P ` poly_of_const ` K" and |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
960 |
?F = "rupture_surj A P ` poly_of_const ` A" and |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
961 |
?E = "carrier (Rupt A P)", OF aux_lemmas]) |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
962 |
show "Rupt.finite_dimension (rupture_surj A P ` poly_of_const ` A) (carrier (Rupt A P))" |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
963 |
using Rupt.finite_dimensionI[OF rupture_dimension[OF A P gt_zero]] . |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
964 |
next |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
965 |
let ?h = "rupture_surj A P \<circ> poly_of_const" |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
966 |
|
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
967 |
from \<open>set P \<subseteq> carrier L\<close> have "finite_dimension K A" |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
968 |
using finite_extension_finite_dimension(1)[OF subfield_axioms, of P] algebraic_extension |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
969 |
unfolding A_def by auto |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
970 |
then obtain Us where Us: "set Us \<subseteq> carrier L" "A = Span K Us" |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
971 |
using exists_base subfield_axioms by blast |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
972 |
hence "?h ` A = Rupt.Span (?h ` K) (map ?h Us)" |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
973 |
using Hom.Span_hom[of K Us] incl Span_base_incl[OF subfield_axioms, of Us] |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
974 |
unfolding Span_consistent[OF subfieldE(1)[OF A]] by simp |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
975 |
moreover have "set (map ?h Us) \<subseteq> carrier (Rupt A P)" |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
976 |
using Span_base_incl[OF subfield_axioms Us(1)] ring_hom_memE(1)[OF Hom.homh] |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
977 |
unfolding sym[OF Us(2)] by auto |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
978 |
ultimately |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
979 |
show "Rupt.finite_dimension (rupture_surj A P ` poly_of_const ` K) (rupture_surj A P ` poly_of_const ` A)" |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
980 |
using Rupt.Span_finite_dimension[OF aux_lemmas(1)] unfolding h_simp by simp |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
981 |
qed |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
982 |
|
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
983 |
moreover have "rupture_surj A P ` carrier (A[X]) = carrier (Rupt A P)" |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
984 |
unfolding rupture_def FactRing_def A_RCOSETS_def' by auto |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
985 |
with \<open>carrier (K[X]) \<subseteq> carrier (A[X])\<close> have "rupture_surj A P ` carrier (K[X]) \<subseteq> carrier (Rupt A P)" |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
986 |
by auto |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
987 |
|
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
988 |
ultimately |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
989 |
have "Rupt.finite_dimension (rupture_surj A P ` poly_of_const ` K) (rupture_surj A P ` carrier (K[X]))" |
70215
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
paulson <lp15@cam.ac.uk>
parents:
70214
diff
changeset
|
990 |
using Rupt.subalbegra_incl_imp_finite_dimension[OF aux_lemmas(1)] by simp |
70212
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
991 |
|
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
992 |
hence "\<not> inj_on (rupture_surj A P) (carrier (K[X]))" |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
993 |
using Hom'.infinite_dimension_hom[OF _ rupture_one_not_zero[OF A P gt_zero] _ |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
994 |
UP_K.carrier_is_subalgebra[OF subfieldE(3)] univ_poly_infinite_dimension[OF subfield_axioms]] |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
995 |
univ_poly_subfield_of_consts[OF subfield_axioms] |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
996 |
by auto |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
997 |
then obtain Q where Q: "Q \<in> carrier (K[X])" "Q \<noteq> []" and "rupture_surj A P Q = \<zero>\<^bsub>Rupt A P\<^esub>" |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
998 |
using Hom'.trivial_ker_imp_inj Hom'.hom_zero unfolding a_kernel_def' univ_poly_zero by blast |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
999 |
with \<open>carrier (K[X]) \<subseteq> carrier (A[X])\<close> have "Q \<in> PIdl\<^bsub>A[X]\<^esub> P" |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
1000 |
using ideal.rcos_const_imp_mem[OF UP_A.cgenideal_ideal[OF P]] |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
1001 |
unfolding rupture_def FactRing_def by auto |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
1002 |
then obtain R where "R \<in> carrier (A[X])" and "Q = R \<otimes>\<^bsub>A[X]\<^esub> P" |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
1003 |
unfolding cgenideal_def by blast |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
1004 |
with \<open>P \<in> carrier (A[X])\<close> have "P pdivides Q" |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
1005 |
using dividesI[of _ "A[X]"] UP_A.m_comm pdivides_iff_shell[OF A] by simp |
70215
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
paulson <lp15@cam.ac.uk>
parents:
70214
diff
changeset
|
1006 |
thus "splitted P" |
70212
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
1007 |
using pdivides_imp_splitted[OF in_carrier |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
1008 |
carrier_polynomial_shell[OF subfieldE(1)[OF subfield_axioms] Q(1)] Q(2) |
70215
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
paulson <lp15@cam.ac.uk>
parents:
70214
diff
changeset
|
1009 |
roots_over_subfield[OF Q(1)]] Q |
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
paulson <lp15@cam.ac.uk>
parents:
70214
diff
changeset
|
1010 |
by simp |
70212
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
1011 |
qed |
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
1012 |
|
70160
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff
changeset
|
1013 |
end |
70212
e886b58bf698
full proof of algebraic closure, by Paulo de Vilhena
paulson <lp15@cam.ac.uk>
parents:
70160
diff
changeset
|
1014 |