| author | wenzelm | 
| Tue, 09 Nov 2021 17:20:04 +0100 | |
| changeset 74739 | a06652d397a7 | 
| parent 73932 | fd21b4a93043 | 
| child 79889 | b187c1b9d6a9 | 
| 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/Polynomial_Divisibility.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  | 
|
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
5  | 
theory Polynomial_Divisibility  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
6  | 
imports Polynomials Embedded_Algebras "HOL-Library.Multiset"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
7  | 
|
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
8  | 
begin  | 
| 
 
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  | 
section \<open>Divisibility of Polynomials\<close>  | 
| 
 
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  | 
subsection \<open>Definitions\<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  | 
abbreviation poly_ring :: "_ \<Rightarrow> ('a  list) ring"
 | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
15  | 
where "poly_ring R \<equiv> univ_poly R (carrier R)"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
16  | 
|
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
17  | 
abbreviation pirreducible :: "_ \<Rightarrow> 'a set \<Rightarrow> 'a list \<Rightarrow> bool" ("pirreducible\<index>")
 | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
18  | 
where "pirreducible\<^bsub>R\<^esub> K p \<equiv> ring_irreducible\<^bsub>(univ_poly R K)\<^esub> p"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
19  | 
|
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
20  | 
abbreviation pprime :: "_ \<Rightarrow> 'a set \<Rightarrow> 'a list \<Rightarrow> bool" ("pprime\<index>")
 | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
21  | 
where "pprime\<^bsub>R\<^esub> K p \<equiv> ring_prime\<^bsub>(univ_poly R K)\<^esub> p"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
22  | 
|
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
23  | 
definition pdivides :: "_ \<Rightarrow> 'a list \<Rightarrow> 'a list \<Rightarrow> bool" (infix "pdivides\<index>" 65)  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
24  | 
where "p pdivides\<^bsub>R\<^esub> q = p divides\<^bsub>(univ_poly R (carrier R))\<^esub> q"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
25  | 
|
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
26  | 
definition rupture :: "_ \<Rightarrow> 'a set \<Rightarrow> 'a list \<Rightarrow> (('a list) set) ring" ("Rupt\<index>")
 | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
27  | 
where "Rupt\<^bsub>R\<^esub> K p = (K[X]\<^bsub>R\<^esub>) Quot (PIdl\<^bsub>K[X]\<^bsub>R\<^esub>\<^esub> p)"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
28  | 
|
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
29  | 
abbreviation (in ring) rupture_surj :: "'a set \<Rightarrow> 'a list \<Rightarrow> 'a list \<Rightarrow> ('a list) set"
 | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
30  | 
where "rupture_surj K p \<equiv> (\<lambda>q. (PIdl\<^bsub>K[X]\<^esub> p) +>\<^bsub>K[X]\<^esub> q)"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
31  | 
|
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
32  | 
|
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
33  | 
subsection \<open>Basic Properties\<close>  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
34  | 
|
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
35  | 
lemma (in ring) carrier_polynomial_shell [intro]:  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
36  | 
assumes "subring K R" and "p \<in> carrier (K[X])" shows "p \<in> carrier (poly_ring R)"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
37  | 
using carrier_polynomial[OF assms(1), of p] assms(2) 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
 | 
38  | 
|
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
39  | 
lemma (in domain) pdivides_zero:  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
40  | 
assumes "subring K R" and "p \<in> carrier (K[X])" shows "p pdivides []"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
41  | 
using ring.divides_zero[OF univ_poly_is_ring[OF carrier_is_subring]  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
42  | 
carrier_polynomial_shell[OF assms]]  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
43  | 
unfolding univ_poly_zero pdivides_def .  | 
| 
 
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 (in domain) zero_pdivides_zero: "[] pdivides []"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
46  | 
using pdivides_zero[OF carrier_is_subring] univ_poly_carrier by blast  | 
| 
 
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 (in domain) zero_pdivides:  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
49  | 
shows "[] pdivides p \<longleftrightarrow> p = []"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
50  | 
using ring.zero_divides[OF univ_poly_is_ring[OF carrier_is_subring]]  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
51  | 
unfolding univ_poly_zero pdivides_def .  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
52  | 
|
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
53  | 
lemma (in domain) pprime_iff_pirreducible:  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
54  | 
assumes "subfield K R" and "p \<in> carrier (K[X])"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
55  | 
shows "pprime K p \<longleftrightarrow> pirreducible K p"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
56  | 
using principal_domain.primeness_condition[OF univ_poly_is_principal] assms by simp  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
57  | 
|
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
58  | 
lemma (in domain) pirreducibleE:  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
59  | 
assumes "subring K R" "p \<in> carrier (K[X])" "pirreducible K p"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
60  | 
shows "p \<noteq> []" "p \<notin> Units (K[X])"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
61  | 
and "\<And>q r. \<lbrakk> q \<in> carrier (K[X]); r \<in> carrier (K[X])\<rbrakk> \<Longrightarrow>  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
62  | 
p = q \<otimes>\<^bsub>K[X]\<^esub> r \<Longrightarrow> q \<in> Units (K[X]) \<or> r \<in> Units (K[X])"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
63  | 
using domain.ring_irreducibleE[OF univ_poly_is_domain[OF assms(1)] _ assms(3)] assms(2)  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
64  | 
by (auto simp add: univ_poly_zero)  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
65  | 
|
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
66  | 
lemma (in domain) pirreducibleI:  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
67  | 
assumes "subring K R" "p \<in> carrier (K[X])" "p \<noteq> []" "p \<notin> Units (K[X])"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
68  | 
and "\<And>q r. \<lbrakk> q \<in> carrier (K[X]); r \<in> carrier (K[X])\<rbrakk> \<Longrightarrow>  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
69  | 
p = q \<otimes>\<^bsub>K[X]\<^esub> r \<Longrightarrow> q \<in> Units (K[X]) \<or> r \<in> Units (K[X])"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
70  | 
shows "pirreducible K p"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
71  | 
using domain.ring_irreducibleI[OF univ_poly_is_domain[OF assms(1)] _ assms(4)] assms(2-3,5)  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
72  | 
by (auto simp add: univ_poly_zero)  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
73  | 
|
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
74  | 
lemma (in domain) univ_poly_carrier_units_incl:  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
75  | 
  shows "Units ((carrier R) [X]) \<subseteq> { [ k ] | k. k \<in> carrier R - { \<zero> } }"
 | 
| 
 
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  | 
fix p assume "p \<in> Units ((carrier R) [X])"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
78  | 
then obtain q  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
79  | 
where p: "polynomial (carrier R) p" and q: "polynomial (carrier R) q" and pq: "poly_mult p q = [ \<one> ]"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
80  | 
unfolding Units_def univ_poly_def by auto  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
81  | 
hence not_nil: "p \<noteq> []" and "q \<noteq> []"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
82  | 
using poly_mult_integral[OF carrier_is_subring p q] poly_mult_zero[OF polynomial_incl[OF p]] by auto  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
83  | 
hence "degree p = 0"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
84  | 
using poly_mult_degree_eq[OF carrier_is_subring p q] unfolding pq by simp  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
85  | 
hence "length p = 1"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
86  | 
using not_nil by (metis One_nat_def Suc_pred length_greater_0_conv)  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
87  | 
then obtain k where k: "p = [ k ]"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
88  | 
by (metis One_nat_def length_0_conv length_Suc_conv)  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
89  | 
  hence "k \<in> carrier R - { \<zero> }"
 | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
90  | 
using p unfolding polynomial_def by auto  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
91  | 
  thus "p \<in> { [ k ] | k. k \<in> carrier R - { \<zero> } }"
 | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
92  | 
unfolding k by blast  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
93  | 
qed  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
94  | 
|
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
95  | 
lemma (in field) univ_poly_carrier_units:  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
96  | 
  "Units ((carrier R) [X]) = { [ k ] | k. k \<in> carrier R - { \<zero> } }"
 | 
| 
 
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  | 
  show "Units ((carrier R) [X]) \<subseteq> { [ k ] | k. k \<in> carrier R - { \<zero> } }"
 | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
99  | 
using univ_poly_carrier_units_incl by simp  | 
| 
 
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  | 
  show "{ [ k ] | k. k \<in> carrier R - { \<zero> } } \<subseteq> Units ((carrier R) [X])"
 | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
102  | 
proof (auto)  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
103  | 
fix k assume k: "k \<in> carrier R" "k \<noteq> \<zero>"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
104  | 
hence inv_k: "inv k \<in> carrier R" "inv k \<noteq> \<zero>" and "k \<otimes> inv k = \<one>" "inv k \<otimes> k = \<one>"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
105  | 
using subfield_m_inv[OF carrier_is_subfield, of k] by auto  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
106  | 
hence "poly_mult [ k ] [ inv k ] = [ \<one> ]" and "poly_mult [ inv k ] [ k ] = [ \<one> ]"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
107  | 
by (auto simp add: k)  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
108  | 
moreover have "polynomial (carrier R) [ k ]" and "polynomial (carrier R) [ inv k ]"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
109  | 
using const_is_polynomial k inv_k by auto  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
110  | 
ultimately show "[ k ] \<in> Units ((carrier R) [X])"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
111  | 
unfolding Units_def univ_poly_def by (auto simp del: poly_mult.simps)  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
112  | 
qed  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
113  | 
qed  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
114  | 
|
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
115  | 
lemma (in domain) univ_poly_units_incl:  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
116  | 
  assumes "subring K R" shows "Units (K[X]) \<subseteq> { [ k ] | k. k \<in> K - { \<zero> } }"
 | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
117  | 
using domain.univ_poly_carrier_units_incl[OF subring_is_domain[OF assms]]  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
118  | 
univ_poly_consistent[OF assms] by auto  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
119  | 
|
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
120  | 
lemma (in ring) univ_poly_units:  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
121  | 
  assumes "subfield K R" shows "Units (K[X]) = { [ k ] | k. k \<in> K - { \<zero> } }"
 | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
122  | 
using field.univ_poly_carrier_units[OF subfield_iff(2)[OF assms]]  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
123  | 
univ_poly_consistent[OF subfieldE(1)[OF assms]] by auto  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
124  | 
|
| 
70215
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
125  | 
lemma (in domain) univ_poly_units':  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
126  | 
assumes "subfield K R" shows "p \<in> Units (K[X]) \<longleftrightarrow> p \<in> carrier (K[X]) \<and> p \<noteq> [] \<and> degree p = 0"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
127  | 
unfolding univ_poly_units[OF assms] sym[OF univ_poly_carrier] polynomial_def  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
128  | 
by (auto, metis hd_in_set le_0_eq le_Suc_eq length_0_conv length_Suc_conv list.sel(1) subsetD)  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
129  | 
|
| 
70160
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
130  | 
corollary (in domain) rupture_one_not_zero:  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
131  | 
assumes "subfield K R" and "p \<in> carrier (K[X])" and "degree p > 0"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
132  | 
shows "\<one>\<^bsub>Rupt K p\<^esub> \<noteq> \<zero>\<^bsub>Rupt K p\<^esub>"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
133  | 
proof (rule ccontr)  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
134  | 
interpret UP: principal_domain "K[X]"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
135  | 
using univ_poly_is_principal[OF assms(1)] .  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
136  | 
|
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
137  | 
assume "\<not> \<one>\<^bsub>Rupt K p\<^esub> \<noteq> \<zero>\<^bsub>Rupt K p\<^esub>"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
138  | 
then have "PIdl\<^bsub>K[X]\<^esub> p +>\<^bsub>K[X]\<^esub> \<one>\<^bsub>K[X]\<^esub> = PIdl\<^bsub>K[X]\<^esub> p"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
139  | 
unfolding rupture_def FactRing_def by simp  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
140  | 
hence "\<one>\<^bsub>K[X]\<^esub> \<in> PIdl\<^bsub>K[X]\<^esub> p"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
141  | 
using ideal.rcos_const_imp_mem[OF UP.cgenideal_ideal[OF assms(2)]] by auto  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
142  | 
then obtain q where "q \<in> carrier (K[X])" and "\<one>\<^bsub>K[X]\<^esub> = q \<otimes>\<^bsub>K[X]\<^esub> p"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
143  | 
using assms(2) unfolding cgenideal_def by auto  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
144  | 
hence "p \<in> Units (K[X])"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
145  | 
unfolding Units_def using assms(2) UP.m_comm by auto  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
146  | 
hence "degree p = 0"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
147  | 
unfolding univ_poly_units[OF assms(1)] by auto  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
148  | 
with \<open>degree p > 0\<close> show False  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
149  | 
by simp  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
150  | 
qed  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
151  | 
|
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
152  | 
corollary (in ring) pirreducible_degree:  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
153  | 
assumes "subfield K R" "p \<in> carrier (K[X])" "pirreducible K p"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
154  | 
shows "degree p \<ge> 1"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
155  | 
proof (rule ccontr)  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
156  | 
assume "\<not> degree p \<ge> 1" then have "length p \<le> 1"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
157  | 
by simp  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
158  | 
moreover have "p \<noteq> []" and "p \<notin> Units (K[X])"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
159  | 
using assms(3) by (auto simp add: ring_irreducible_def irreducible_def univ_poly_zero)  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
160  | 
ultimately obtain k where k: "p = [ k ]"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
161  | 
by (metis append_butlast_last_id butlast_take diff_is_0_eq le_refl self_append_conv2 take0 take_all)  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
162  | 
hence "k \<in> K" and "k \<noteq> \<zero>"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
163  | 
using assms(2) by (auto simp add: polynomial_def univ_poly_def)  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
164  | 
hence "p \<in> Units (K[X])"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
165  | 
using univ_poly_units[OF assms(1)] unfolding k by auto  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
166  | 
from \<open>p \<in> Units (K[X])\<close> and \<open>p \<notin> Units (K[X])\<close> show False by simp  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
167  | 
qed  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
168  | 
|
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
169  | 
corollary (in domain) univ_poly_not_field:  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
170  | 
assumes "subring K R" shows "\<not> field (K[X])"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
171  | 
proof -  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
172  | 
  have "X \<in> carrier (K[X]) - { \<zero>\<^bsub>(K[X])\<^esub> }" and "X \<notin> { [ k ] | k. k \<in> K - { \<zero> } }"
 | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
173  | 
using var_closed(1)[OF assms] unfolding univ_poly_zero var_def by auto  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
174  | 
thus ?thesis  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
175  | 
using field.field_Units[of "K[X]"] univ_poly_units_incl[OF assms] by blast  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
176  | 
qed  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
177  | 
|
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
178  | 
lemma (in domain) rupture_is_field_iff_pirreducible:  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
179  | 
assumes "subfield K R" and "p \<in> carrier (K[X])"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
180  | 
shows "field (Rupt K p) \<longleftrightarrow> pirreducible K p"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
181  | 
proof  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
182  | 
assume "pirreducible K p" thus "field (Rupt K p)"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
183  | 
using principal_domain.field_iff_prime[OF univ_poly_is_principal[OF assms(1)]] assms(2)  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
184  | 
pprime_iff_pirreducible[OF assms] pirreducibleE(1)[OF subfieldE(1)[OF assms(1)]]  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
185  | 
by (simp add: univ_poly_zero rupture_def)  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
186  | 
next  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
187  | 
interpret UP: principal_domain "K[X]"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
188  | 
using univ_poly_is_principal[OF assms(1)] .  | 
| 
 
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  | 
assume field: "field (Rupt K p)"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
191  | 
have "p \<noteq> []"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
192  | 
proof (rule ccontr)  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
193  | 
assume "\<not> p \<noteq> []" then have p: "p = []"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
194  | 
by simp  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
195  | 
hence "Rupt K p \<simeq> (K[X])"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
196  | 
using UP.FactRing_zeroideal(1) UP.genideal_zero  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
197  | 
UP.cgenideal_eq_genideal[OF UP.zero_closed]  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
198  | 
by (simp add: rupture_def univ_poly_zero)  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
199  | 
then obtain h where h: "h \<in> ring_iso (Rupt K p) (K[X])"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
200  | 
unfolding is_ring_iso_def by blast  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
201  | 
moreover have "ring (Rupt K p)"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
202  | 
using field by (simp add: cring_def domain_def field_def)  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
203  | 
ultimately interpret R: ring_hom_ring "Rupt K p" "K[X]" h  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
204  | 
unfolding ring_hom_ring_def ring_hom_ring_axioms_def ring_iso_def  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
205  | 
using UP.ring_axioms by simp  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
206  | 
have "field (K[X])"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
207  | 
using field.ring_iso_imp_img_field[OF field h] by simp  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
208  | 
thus False  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
209  | 
using univ_poly_not_field[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
 | 
210  | 
qed  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
211  | 
thus "pirreducible K p"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
212  | 
using UP.field_iff_prime pprime_iff_pirreducible[OF assms] assms(2) field  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
213  | 
by (simp add: univ_poly_zero rupture_def)  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
214  | 
qed  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
215  | 
|
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
216  | 
lemma (in domain) rupture_surj_hom:  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
217  | 
assumes "subring K R" and "p \<in> carrier (K[X])"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
218  | 
shows "(rupture_surj K p) \<in> ring_hom (K[X]) (Rupt K p)"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
219  | 
and "ring_hom_ring (K[X]) (Rupt K p) (rupture_surj K p)"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
220  | 
proof -  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
221  | 
interpret UP: domain "K[X]"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
222  | 
using univ_poly_is_domain[OF assms(1)] .  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
223  | 
interpret I: ideal "PIdl\<^bsub>K[X]\<^esub> p" "K[X]"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
224  | 
using UP.cgenideal_ideal[OF assms(2)] .  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
225  | 
show "(rupture_surj K p) \<in> ring_hom (K[X]) (Rupt K p)"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
226  | 
and "ring_hom_ring (K[X]) (Rupt K p) (rupture_surj K p)"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
227  | 
using ring_hom_ring.intro[OF UP.ring_axioms I.quotient_is_ring] I.rcos_ring_hom  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
228  | 
unfolding symmetric[OF ring_hom_ring_axioms_def] rupture_def by auto  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
229  | 
qed  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
230  | 
|
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
231  | 
corollary (in domain) rupture_surj_norm_is_hom:  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
232  | 
assumes "subring K R" and "p \<in> carrier (K[X])"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
233  | 
shows "((rupture_surj K p) \<circ> poly_of_const) \<in> ring_hom (R \<lparr> carrier := K \<rparr>) (Rupt K p)"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
234  | 
using ring_hom_trans[OF canonical_embedding_is_hom[OF assms(1)] rupture_surj_hom(1)[OF assms]] .  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
235  | 
|
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
236  | 
lemma (in domain) norm_map_in_poly_ring_carrier:  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
237  | 
assumes "p \<in> carrier (poly_ring R)" and "\<And>a. a \<in> carrier R \<Longrightarrow> f a \<in> carrier (poly_ring R)"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
238  | 
shows "ring.normalize (poly_ring R) (map f p) \<in> carrier (poly_ring (poly_ring R))"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
239  | 
proof -  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
240  | 
have "set p \<subseteq> carrier R"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
241  | 
using assms(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
 | 
242  | 
hence "set (map f p) \<subseteq> carrier (poly_ring R)"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
243  | 
using assms(2) by auto  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
244  | 
thus ?thesis  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
245  | 
using ring.normalize_gives_polynomial[OF univ_poly_is_ring[OF carrier_is_subring]]  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
246  | 
unfolding univ_poly_carrier by simp  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
247  | 
qed  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
248  | 
|
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
249  | 
lemma (in domain) map_in_poly_ring_carrier:  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
250  | 
assumes "p \<in> carrier (poly_ring R)" and "\<And>a. a \<in> carrier R \<Longrightarrow> f a \<in> carrier (poly_ring R)"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
251  | 
and "\<And>a. a \<noteq> \<zero> \<Longrightarrow> f a \<noteq> []"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
252  | 
shows "map f p \<in> carrier (poly_ring (poly_ring R))"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
253  | 
proof -  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
254  | 
interpret UP: ring "poly_ring R"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
255  | 
using univ_poly_is_ring[OF carrier_is_subring] .  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
256  | 
have "lead_coeff p \<noteq> \<zero>" if "p \<noteq> []"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
257  | 
using that assms(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
 | 
258  | 
hence "ring.normalize (poly_ring R) (map f p) = map f p"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
259  | 
by (cases p) (simp_all add: assms(3) univ_poly_zero)  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
260  | 
thus ?thesis  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
261  | 
using norm_map_in_poly_ring_carrier[of p f] assms(1-2) by simp  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
262  | 
qed  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
263  | 
|
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
264  | 
lemma (in domain) map_norm_in_poly_ring_carrier:  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
265  | 
assumes "subring K R" and "p \<in> carrier (K[X])"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
266  | 
shows "map poly_of_const p \<in> carrier (poly_ring (K[X]))"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
267  | 
using domain.map_in_poly_ring_carrier[OF subring_is_domain[OF assms(1)]]  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
268  | 
proof -  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
269  | 
have "\<And>a. a \<in> K \<Longrightarrow> poly_of_const a \<in> carrier (K[X])"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
270  | 
and "\<And>a. a \<noteq> \<zero> \<Longrightarrow> poly_of_const a \<noteq> []"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
271  | 
using ring_hom_memE(1)[OF canonical_embedding_is_hom[OF assms(1)]]  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
272  | 
by (auto simp: poly_of_const_def)  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
273  | 
thus ?thesis  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
274  | 
using domain.map_in_poly_ring_carrier[OF subring_is_domain[OF assms(1)]] assms(2)  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
275  | 
unfolding univ_poly_consistent[OF assms(1)] by simp  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
276  | 
qed  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
277  | 
|
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
278  | 
lemma (in domain) polynomial_rupture:  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
279  | 
assumes "subring K R" and "p \<in> carrier (K[X])"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
280  | 
shows "(ring.eval (Rupt K p)) (map ((rupture_surj K p) \<circ> poly_of_const) p) (rupture_surj K p X) = \<zero>\<^bsub>Rupt K p\<^esub>"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
281  | 
proof -  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
282  | 
let ?surj = "rupture_surj K p"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
283  | 
|
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
284  | 
interpret UP: domain "K[X]"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
285  | 
using univ_poly_is_domain[OF assms(1)] .  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
286  | 
interpret Hom: ring_hom_ring "K[X]" "Rupt K p" ?surj  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
287  | 
using rupture_surj_hom(2)[OF assms] .  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
288  | 
|
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
289  | 
have "(Hom.S.eval) (map (?surj \<circ> poly_of_const) p) (?surj X) = ?surj ((UP.eval) (map poly_of_const p) X)"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
290  | 
using Hom.eval_hom[OF UP.carrier_is_subring var_closed(1)[OF assms(1)]  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
291  | 
map_norm_in_poly_ring_carrier[OF assms]] by simp  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
292  | 
also have " ... = ?surj p"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
293  | 
unfolding sym[OF eval_rewrite[OF assms]] ..  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
294  | 
also have " ... = \<zero>\<^bsub>Rupt K p\<^esub>"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
295  | 
using UP.a_rcos_zero[OF UP.cgenideal_ideal[OF assms(2)] UP.cgenideal_self[OF assms(2)]]  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
296  | 
unfolding rupture_def FactRing_def by simp  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
297  | 
finally show ?thesis .  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
298  | 
qed  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
299  | 
|
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
300  | 
|
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
301  | 
subsection \<open>Division\<close>  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
302  | 
|
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
303  | 
definition (in ring) long_divides :: "'a list \<Rightarrow> 'a list \<Rightarrow> ('a list \<times> 'a list) \<Rightarrow> bool"
 | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
304  | 
where "long_divides p q t \<longleftrightarrow>  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
305  | 
\<comment> \<open>i\<close> (t \<in> carrier (poly_ring R) \<times> carrier (poly_ring R)) \<and>  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
306  | 
\<comment> \<open>ii\<close> (p = (q \<otimes>\<^bsub>poly_ring R\<^esub> (fst t)) \<oplus>\<^bsub>poly_ring R\<^esub> (snd t)) \<and>  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
307  | 
\<comment> \<open>iii\<close> (snd t = [] \<or> degree (snd t) < degree q)"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
308  | 
|
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
309  | 
definition (in ring) long_division :: "'a list \<Rightarrow> 'a list \<Rightarrow> ('a list \<times> 'a list)"
 | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
310  | 
where "long_division p q = (THE t. long_divides p q t)"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
311  | 
|
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
312  | 
definition (in ring) pdiv :: "'a list \<Rightarrow> 'a list \<Rightarrow> 'a list" (infixl "pdiv" 65)  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
313  | 
where "p pdiv q = (if q = [] then [] else fst (long_division p q))"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
314  | 
|
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
315  | 
definition (in ring) pmod :: "'a list \<Rightarrow> 'a list \<Rightarrow> 'a list" (infixl "pmod" 65)  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
316  | 
where "p pmod q = (if q = [] then p else snd (long_division p q))"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
317  | 
|
| 
 
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 (in ring) long_dividesI:  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
320  | 
assumes "b \<in> carrier (poly_ring R)" and "r \<in> carrier (poly_ring R)"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
321  | 
and "p = (q \<otimes>\<^bsub>poly_ring R\<^esub> b) \<oplus>\<^bsub>poly_ring R\<^esub> r" and "r = [] \<or> degree r < degree q"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
322  | 
shows "long_divides p q (b, r)"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
323  | 
using assms unfolding long_divides_def by auto  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
324  | 
|
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
325  | 
lemma (in domain) exists_long_division:  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
326  | 
assumes "subfield K R" and "p \<in> carrier (K[X])" and "q \<in> carrier (K[X])" "q \<noteq> []"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
327  | 
obtains b r where "b \<in> carrier (K[X])" and "r \<in> carrier (K[X])" and "long_divides p q (b, r)"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
328  | 
using subfield_long_division_theorem_shell[OF assms(1-3)] assms(4)  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
329  | 
carrier_polynomial_shell[OF subfieldE(1)[OF assms(1)]]  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
330  | 
unfolding long_divides_def univ_poly_zero univ_poly_add univ_poly_mult by auto  | 
| 
 
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  | 
lemma (in domain) exists_unique_long_division:  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
333  | 
assumes "subfield K R" and "p \<in> carrier (K[X])" and "q \<in> carrier (K[X])" "q \<noteq> []"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
334  | 
shows "\<exists>!t. long_divides p q t"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
335  | 
proof -  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
336  | 
let ?padd = "\<lambda>a b. a \<oplus>\<^bsub>poly_ring R\<^esub> b"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
337  | 
let ?pmult = "\<lambda>a b. a \<otimes>\<^bsub>poly_ring R\<^esub> b"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
338  | 
let ?pminus = "\<lambda>a b. a \<ominus>\<^bsub>poly_ring R\<^esub> b"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
339  | 
|
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
340  | 
interpret UP: domain "poly_ring R"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
341  | 
using univ_poly_is_domain[OF carrier_is_subring] .  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
342  | 
|
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
343  | 
obtain b r where ldiv: "long_divides p q (b, r)"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
344  | 
using exists_long_division[OF assms] by metis  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
345  | 
|
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
346  | 
moreover have "(b, r) = (b', r')" if "long_divides p q (b', r')" for b' r'  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
347  | 
proof -  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
348  | 
have q: "q \<in> carrier (poly_ring R)" "q \<noteq> []"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
349  | 
using assms(3-4) carrier_polynomial[OF subfieldE(1)[OF assms(1)]]  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
350  | 
unfolding univ_poly_carrier by auto  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
351  | 
hence in_carrier: "q \<in> carrier (poly_ring R)"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
352  | 
"b \<in> carrier (poly_ring R)" "r \<in> carrier (poly_ring R)"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
353  | 
"b' \<in> carrier (poly_ring R)" "r' \<in> carrier (poly_ring R)"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
354  | 
using assms(3) that ldiv unfolding long_divides_def by auto  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
355  | 
have "?pminus (?padd (?pmult q b) r) r' = ?pminus (?padd (?pmult q b') r') r'"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
356  | 
using ldiv and that unfolding long_divides_def by auto  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
357  | 
hence eq: "?padd (?pmult q (?pminus b b')) (?pminus r r') = \<zero>\<^bsub>poly_ring R\<^esub>"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
358  | 
using in_carrier by algebra  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
359  | 
have "b = b'"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
360  | 
proof (rule ccontr)  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
361  | 
assume "b \<noteq> b'"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
362  | 
hence pminus: "?pminus b b' \<noteq> \<zero>\<^bsub>poly_ring R\<^esub>" "?pminus b b' \<in> carrier (poly_ring R)"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
363  | 
using in_carrier(2,4) by (metis UP.add.inv_closed UP.l_neg UP.minus_eq UP.minus_unique, algebra)  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
364  | 
hence degree_ge: "degree (?pmult q (?pminus b b')) \<ge> degree q"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
365  | 
using poly_mult_degree_eq[OF carrier_is_subring, of q "?pminus b b'"] q  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
366  | 
unfolding univ_poly_zero univ_poly_carrier univ_poly_mult by simp  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
367  | 
|
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
368  | 
have "?pminus b b' = \<zero>\<^bsub>poly_ring R\<^esub>" if "?pminus r r' = \<zero>\<^bsub>poly_ring R\<^esub>"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
369  | 
using eq pminus(2) q UP.integral univ_poly_zero unfolding that by auto  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
370  | 
hence "?pminus r r' \<noteq> []"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
371  | 
using pminus(1) unfolding univ_poly_zero by blast  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
372  | 
moreover have "?pminus r r' = []" if "r = []" and "r' = []"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
373  | 
using univ_poly_a_inv_def'[OF carrier_is_subring UP.zero_closed] that  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
374  | 
unfolding a_minus_def univ_poly_add univ_poly_zero by auto  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
375  | 
ultimately have "r \<noteq> [] \<or> r' \<noteq> []"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
376  | 
by blast  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
377  | 
hence "max (degree r) (degree r') < degree q"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
378  | 
using ldiv and that unfolding long_divides_def by auto  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
379  | 
moreover have "degree (?pminus r r') \<le> max (degree r) (degree r')"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
380  | 
using poly_add_degree[of r "map (a_inv R) r'"]  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
381  | 
unfolding a_minus_def univ_poly_add univ_poly_a_inv_def'[OF carrier_is_subring in_carrier(5)]  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
382  | 
by auto  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
383  | 
ultimately have degree_lt: "degree (?pminus r r') < degree q"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
384  | 
by linarith  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
385  | 
have is_poly: "polynomial (carrier R) (?pmult q (?pminus b b'))" "polynomial (carrier R) (?pminus r r')"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
386  | 
using in_carrier pminus(2) unfolding univ_poly_carrier by algebra+  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
387  | 
|
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
388  | 
have "degree (?padd (?pmult q (?pminus b b')) (?pminus r r')) = degree (?pmult q (?pminus b b'))"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
389  | 
using poly_add_degree_eq[OF carrier_is_subring is_poly] degree_ge degree_lt  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
390  | 
unfolding univ_poly_carrier sym[OF univ_poly_add[of R "carrier R"]] max_def by simp  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
391  | 
hence "degree (?padd (?pmult q (?pminus b b')) (?pminus r r')) > 0"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
392  | 
using degree_ge degree_lt by simp  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
393  | 
moreover have "degree (?padd (?pmult q (?pminus b b')) (?pminus r r')) = 0"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
394  | 
using eq unfolding univ_poly_zero by simp  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
395  | 
ultimately show False by simp  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
396  | 
qed  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
397  | 
hence "?pminus r r' = \<zero>\<^bsub>poly_ring R\<^esub>"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
398  | 
using in_carrier eq by algebra  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
399  | 
hence "r = r'"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
400  | 
using in_carrier by (metis UP.add.inv_closed UP.add.right_cancel UP.minus_eq UP.r_neg)  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
401  | 
with \<open>b = b'\<close> show ?thesis  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
402  | 
by simp  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
403  | 
qed  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
404  | 
|
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
405  | 
ultimately show ?thesis  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
406  | 
by auto  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
407  | 
qed  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
408  | 
|
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
409  | 
lemma (in domain) long_divisionE:  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
410  | 
assumes "subfield K R" and "p \<in> carrier (K[X])" and "q \<in> carrier (K[X])" "q \<noteq> []"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
411  | 
shows "long_divides p q (p pdiv q, p pmod q)"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
412  | 
using theI'[OF exists_unique_long_division[OF assms]] assms(4)  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
413  | 
unfolding pmod_def pdiv_def long_division_def by auto  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
414  | 
|
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
415  | 
lemma (in domain) long_divisionI:  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
416  | 
assumes "subfield K R" and "p \<in> carrier (K[X])" and "q \<in> carrier (K[X])" "q \<noteq> []"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
417  | 
shows "long_divides p q (b, r) \<Longrightarrow> (b, r) = (p pdiv q, p pmod q)"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
418  | 
using exists_unique_long_division[OF assms] long_divisionE[OF assms] by metis  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
419  | 
|
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
420  | 
lemma (in domain) long_division_closed:  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
421  | 
assumes "subfield K R" and "p \<in> carrier (K[X])" "q \<in> carrier (K[X])"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
422  | 
shows "p pdiv q \<in> carrier (K[X])" and "p pmod q \<in> carrier (K[X])"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
423  | 
proof -  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
424  | 
have "p pdiv q \<in> carrier (K[X]) \<and> p pmod q \<in> carrier (K[X])"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
425  | 
using assms univ_poly_zero_closed[of R] long_divisionI[of K] exists_long_division[OF assms]  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
426  | 
by (cases "q = []") (simp add: pdiv_def pmod_def, metis Pair_inject)+  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
427  | 
thus "p pdiv q \<in> carrier (K[X])" and "p pmod q \<in> carrier (K[X])"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
428  | 
by auto  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
429  | 
qed  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
430  | 
|
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
431  | 
lemma (in domain) pdiv_pmod:  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
432  | 
assumes "subfield K R" and "p \<in> carrier (K[X])" "q \<in> carrier (K[X])"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
433  | 
shows "p = (q \<otimes>\<^bsub>K[X]\<^esub> (p pdiv q)) \<oplus>\<^bsub>K[X]\<^esub> (p pmod q)"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
434  | 
proof (cases)  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
435  | 
interpret UP: ring "K[X]"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
436  | 
using univ_poly_is_ring[OF subfieldE(1)[OF assms(1)]] .  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
437  | 
assume "q = []" thus ?thesis  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
438  | 
using assms(2) unfolding pdiv_def pmod_def sym[OF univ_poly_zero[of R K]] by simp  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
439  | 
next  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
440  | 
assume "q \<noteq> []" thus ?thesis  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
441  | 
using long_divisionE[OF assms] unfolding long_divides_def univ_poly_mult univ_poly_add by simp  | 
| 
 
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  | 
lemma (in domain) pmod_degree:  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
445  | 
assumes "subfield K R" and "p \<in> carrier (K[X])" and "q \<in> carrier (K[X])" "q \<noteq> []"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
446  | 
shows "p pmod q = [] \<or> degree (p pmod q) < degree q"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
447  | 
using long_divisionE[OF assms] unfolding long_divides_def by auto  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
448  | 
|
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
449  | 
lemma (in domain) pmod_const:  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
450  | 
assumes "subfield K R" and "p \<in> carrier (K[X])" "q \<in> carrier (K[X])" and "degree q > degree p"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
451  | 
shows "p pdiv q = []" and "p pmod q = p"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
452  | 
proof -  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
453  | 
have "p pdiv q = [] \<and> p pmod q = p"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
454  | 
proof (cases)  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
455  | 
interpret UP: ring "K[X]"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
456  | 
using univ_poly_is_ring[OF subfieldE(1)[OF assms(1)]] .  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
457  | 
|
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
458  | 
assume "q \<noteq> []"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
459  | 
have "p = (q \<otimes>\<^bsub>K[X]\<^esub> []) \<oplus>\<^bsub>K[X]\<^esub> p"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
460  | 
using assms(2-3) unfolding sym[OF univ_poly_zero[of R K]] by simp  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
461  | 
moreover have "([], p) \<in> carrier (poly_ring R) \<times> carrier (poly_ring R)"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
462  | 
using carrier_polynomial_shell[OF subfieldE(1)[OF assms(1)] assms(2)] by auto  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
463  | 
ultimately have "long_divides p q ([], p)"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
464  | 
using assms(4) unfolding long_divides_def univ_poly_mult univ_poly_add by auto  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
465  | 
with \<open>q \<noteq> []\<close> show ?thesis  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
466  | 
using long_divisionI[OF assms(1-3)] by auto  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
467  | 
qed (simp add: pmod_def pdiv_def)  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
468  | 
thus "p pdiv q = []" and "p pmod q = p"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
469  | 
by auto  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
470  | 
qed  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
471  | 
|
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
472  | 
lemma (in domain) long_division_zero:  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
473  | 
assumes "subfield K R" and "q \<in> carrier (K[X])" shows "[] pdiv q = []" and "[] pmod q = []"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
474  | 
proof -  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
475  | 
interpret UP: ring "poly_ring R"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
476  | 
using univ_poly_is_ring[OF carrier_is_subring] .  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
477  | 
|
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
478  | 
have "[] pdiv q = [] \<and> [] pmod q = []"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
479  | 
proof (cases)  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
480  | 
assume "q \<noteq> []"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
481  | 
have "q \<in> carrier (poly_ring R)"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
482  | 
using carrier_polynomial_shell[OF subfieldE(1)[OF assms(1)] assms(2)] .  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
483  | 
hence "long_divides [] q ([], [])"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
484  | 
unfolding long_divides_def sym[OF univ_poly_zero[of R "carrier R"]] by auto  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
485  | 
with \<open>q \<noteq> []\<close> show ?thesis  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
486  | 
using long_divisionI[OF assms(1) univ_poly_zero_closed assms(2)] by simp  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
487  | 
qed (simp add: pmod_def pdiv_def)  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
488  | 
thus "[] pdiv q = []" and "[] pmod q = []"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
489  | 
by auto  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
490  | 
qed  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
491  | 
|
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
492  | 
lemma (in domain) long_division_a_inv:  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
493  | 
assumes "subfield K R" and "p \<in> carrier (K[X])" "q \<in> carrier (K[X])"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
494  | 
shows "((\<ominus>\<^bsub>K[X]\<^esub> p) pdiv q) = \<ominus>\<^bsub>K[X]\<^esub> (p pdiv q)" (is "?pdiv")  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
495  | 
and "((\<ominus>\<^bsub>K[X]\<^esub> p) pmod q) = \<ominus>\<^bsub>K[X]\<^esub> (p pmod q)" (is "?pmod")  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
496  | 
proof -  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
497  | 
interpret UP: ring "K[X]"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
498  | 
using univ_poly_is_ring[OF subfieldE(1)[OF assms(1)]] .  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
499  | 
|
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
500  | 
have "?pdiv \<and> ?pmod"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
501  | 
proof (cases)  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
502  | 
assume "q = []" thus ?thesis  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
503  | 
unfolding pmod_def pdiv_def sym[OF univ_poly_zero[of R K]] by simp  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
504  | 
next  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
505  | 
assume not_nil: "q \<noteq> []"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
506  | 
have "\<ominus>\<^bsub>K[X]\<^esub> p = \<ominus>\<^bsub>K[X]\<^esub> ((q \<otimes>\<^bsub>K[X]\<^esub> (p pdiv q)) \<oplus>\<^bsub>K[X]\<^esub> (p pmod q))"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
507  | 
using pdiv_pmod[OF assms] by simp  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
508  | 
hence "\<ominus>\<^bsub>K[X]\<^esub> p = (q \<otimes>\<^bsub>K[X]\<^esub> (\<ominus>\<^bsub>K[X]\<^esub> (p pdiv q))) \<oplus>\<^bsub>K[X]\<^esub> (\<ominus>\<^bsub>K[X]\<^esub> (p pmod q))"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
509  | 
using assms(2-3) long_division_closed[OF assms] by algebra  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
510  | 
moreover have "\<ominus>\<^bsub>K[X]\<^esub> (p pdiv q) \<in> carrier (K[X])" "\<ominus>\<^bsub>K[X]\<^esub> (p pmod q) \<in> carrier (K[X])"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
511  | 
using long_division_closed[OF assms] by algebra+  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
512  | 
hence "(\<ominus>\<^bsub>K[X]\<^esub> (p pdiv q), \<ominus>\<^bsub>K[X]\<^esub> (p pmod q)) \<in> carrier (poly_ring R) \<times> carrier (poly_ring R)"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
513  | 
using carrier_polynomial_shell[OF subfieldE(1)[OF assms(1)]] by auto  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
514  | 
moreover have "\<ominus>\<^bsub>K[X]\<^esub> (p pmod q) = [] \<or> degree (\<ominus>\<^bsub>K[X]\<^esub> (p pmod q)) < degree q"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
515  | 
using univ_poly_a_inv_length[OF subfieldE(1)[OF assms(1)]  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
516  | 
long_division_closed(2)[OF assms]] pmod_degree[OF assms not_nil]  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
517  | 
by auto  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
518  | 
ultimately have "long_divides (\<ominus>\<^bsub>K[X]\<^esub> p) q (\<ominus>\<^bsub>K[X]\<^esub> (p pdiv q), \<ominus>\<^bsub>K[X]\<^esub> (p pmod q))"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
519  | 
unfolding long_divides_def univ_poly_mult univ_poly_add by simp  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
520  | 
thus ?thesis  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
521  | 
using long_divisionI[OF assms(1) UP.a_inv_closed[OF assms(2)] assms(3) not_nil] by simp  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
522  | 
qed  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
523  | 
thus ?pdiv and ?pmod  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
524  | 
by auto  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
525  | 
qed  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
526  | 
|
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
527  | 
lemma (in domain) long_division_add:  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
528  | 
assumes "subfield K R" and "a \<in> carrier (K[X])" "b \<in> carrier (K[X])" "q \<in> carrier (K[X])"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
529  | 
shows "(a \<oplus>\<^bsub>K[X]\<^esub> b) pdiv q = (a pdiv q) \<oplus>\<^bsub>K[X]\<^esub> (b pdiv q)" (is "?pdiv")  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
530  | 
and "(a \<oplus>\<^bsub>K[X]\<^esub> b) pmod q = (a pmod q) \<oplus>\<^bsub>K[X]\<^esub> (b pmod q)" (is "?pmod")  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
531  | 
proof -  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
532  | 
let ?pdiv_add = "(a pdiv q) \<oplus>\<^bsub>K[X]\<^esub> (b pdiv q)"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
533  | 
let ?pmod_add = "(a pmod q) \<oplus>\<^bsub>K[X]\<^esub> (b pmod q)"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
534  | 
|
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
535  | 
interpret UP: ring "K[X]"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
536  | 
using univ_poly_is_ring[OF subfieldE(1)[OF assms(1)]] .  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
537  | 
|
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
538  | 
have "?pdiv \<and> ?pmod"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
539  | 
proof (cases)  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
540  | 
assume "q = []" thus ?thesis  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
541  | 
using assms(2-3) unfolding pmod_def pdiv_def sym[OF univ_poly_zero[of R K]] by simp  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
542  | 
next  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
543  | 
note in_carrier = long_division_closed[OF assms(1,2,4)]  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
544  | 
long_division_closed[OF assms(1,3,4)]  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
545  | 
|
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
546  | 
assume "q \<noteq> []"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
547  | 
have "a \<oplus>\<^bsub>K[X]\<^esub> b = ((q \<otimes>\<^bsub>K[X]\<^esub> (a pdiv q)) \<oplus>\<^bsub>K[X]\<^esub> (a pmod q)) \<oplus>\<^bsub>K[X]\<^esub>  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
548  | 
((q \<otimes>\<^bsub>K[X]\<^esub> (b pdiv q)) \<oplus>\<^bsub>K[X]\<^esub> (b pmod q))"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
549  | 
using assms(2-3)[THEN pdiv_pmod[OF assms(1) _ assms(4)]] by simp  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
550  | 
hence "a \<oplus>\<^bsub>K[X]\<^esub> b = (q \<otimes>\<^bsub>K[X]\<^esub> ?pdiv_add) \<oplus>\<^bsub>K[X]\<^esub> ?pmod_add"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
551  | 
using assms(4) in_carrier by algebra  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
552  | 
moreover have "(?pdiv_add, ?pmod_add) \<in> carrier (poly_ring R) \<times> carrier (poly_ring R)"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
553  | 
using in_carrier carrier_polynomial_shell[OF subfieldE(1)[OF assms(1)]] by auto  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
554  | 
moreover have "?pmod_add = [] \<or> degree ?pmod_add < degree q"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
555  | 
proof (cases)  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
556  | 
assume "?pmod_add \<noteq> []"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
557  | 
hence "a pmod q \<noteq> [] \<or> b pmod q \<noteq> []"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
558  | 
using in_carrier(2,4) unfolding sym[OF univ_poly_zero[of R K]] by auto  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
559  | 
moreover from \<open>q \<noteq> []\<close>  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
560  | 
have "a pmod q = [] \<or> degree (a pmod q) < degree q" and "b pmod q = [] \<or> degree (b pmod q) < degree q"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
561  | 
using assms(2-3)[THEN pmod_degree[OF assms(1) _ assms(4)]] by auto  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
562  | 
ultimately have "max (degree (a pmod q)) (degree (b pmod q)) < degree q"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
563  | 
by auto  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
564  | 
thus ?thesis  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
565  | 
using poly_add_degree le_less_trans unfolding univ_poly_add by blast  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
566  | 
qed simp  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
567  | 
ultimately have "long_divides (a \<oplus>\<^bsub>K[X]\<^esub> b) q (?pdiv_add, ?pmod_add)"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
568  | 
unfolding long_divides_def univ_poly_mult univ_poly_add by simp  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
569  | 
with \<open>q \<noteq> []\<close> show ?thesis  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
570  | 
using long_divisionI[OF assms(1) UP.a_closed[OF assms(2-3)] assms(4)] by simp  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
571  | 
qed  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
572  | 
thus ?pdiv and ?pmod  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
573  | 
by auto  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
574  | 
qed  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
575  | 
|
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
576  | 
lemma (in domain) long_division_add_iff:  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
577  | 
assumes "subfield K R"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
578  | 
and "a \<in> carrier (K[X])" "b \<in> carrier (K[X])" "c \<in> carrier (K[X])" "q \<in> carrier (K[X])"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
579  | 
shows "a pmod q = b pmod q \<longleftrightarrow> (a \<oplus>\<^bsub>K[X]\<^esub> c) pmod q = (b \<oplus>\<^bsub>K[X]\<^esub> c) pmod q"  | 
| 
 
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  | 
interpret UP: ring "K[X]"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
582  | 
using univ_poly_is_ring[OF subfieldE(1)[OF assms(1)]] .  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
583  | 
show ?thesis  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
584  | 
using assms(2-4)[THEN long_division_closed(2)[OF assms(1) _ assms(5)]]  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
585  | 
unfolding assms(2-3)[THEN long_division_add(2)[OF assms(1) _ assms(4-5)]] by auto  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
586  | 
qed  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
587  | 
|
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
588  | 
lemma (in domain) pdivides_iff:  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
589  | 
assumes "subfield K R" and "polynomial K p" "polynomial K q"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
590  | 
shows "p pdivides q \<longleftrightarrow> p divides\<^bsub>K[X]\<^esub> q"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
591  | 
proof  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
592  | 
show "p divides\<^bsub>K [X]\<^esub> q \<Longrightarrow> p pdivides q"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
593  | 
using carrier_polynomial[OF subfieldE(1)[OF assms(1)]]  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
594  | 
unfolding pdivides_def factor_def univ_poly_mult univ_poly_carrier by auto  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
595  | 
next  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
596  | 
interpret UP: ring "poly_ring R"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
597  | 
using univ_poly_is_ring[OF carrier_is_subring] .  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
598  | 
|
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
599  | 
have in_carrier: "p \<in> carrier (poly_ring R)" "q \<in> carrier (poly_ring R)"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
600  | 
using carrier_polynomial[OF subfieldE(1)[OF assms(1)]] assms  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
601  | 
unfolding univ_poly_carrier by auto  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
602  | 
|
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
603  | 
assume "p pdivides q"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
604  | 
then obtain b where "b \<in> carrier (poly_ring R)" and "q = p \<otimes>\<^bsub>poly_ring R\<^esub> b"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
605  | 
unfolding pdivides_def factor_def by blast  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
606  | 
show "p divides\<^bsub>K[X]\<^esub> q"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
607  | 
proof (cases)  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
608  | 
assume "p = []"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
609  | 
with \<open>b \<in> carrier (poly_ring R)\<close> and \<open>q = p \<otimes>\<^bsub>poly_ring R\<^esub> b\<close> have "q = []"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
610  | 
unfolding univ_poly_mult sym[OF univ_poly_carrier]  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
611  | 
using poly_mult_zero(1)[OF polynomial_incl] by simp  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
612  | 
with \<open>p = []\<close> show ?thesis  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
613  | 
using poly_mult_zero(2)[of "[]"]  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
614  | 
unfolding 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
 | 
615  | 
next  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
616  | 
interpret UP: ring "poly_ring R"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
617  | 
using univ_poly_is_ring[OF carrier_is_subring] .  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
618  | 
|
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
619  | 
assume "p \<noteq> []"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
620  | 
from \<open>p pdivides q\<close> obtain b where "b \<in> carrier (poly_ring R)" and "q = p \<otimes>\<^bsub>poly_ring R\<^esub> b"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
621  | 
unfolding pdivides_def factor_def by blast  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
622  | 
moreover have "p \<in> carrier (poly_ring R)" and "q \<in> carrier (poly_ring R)"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
623  | 
using assms carrier_polynomial[OF subfieldE(1)[OF assms(1)]] unfolding univ_poly_carrier by auto  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
624  | 
ultimately have "q = (p \<otimes>\<^bsub>poly_ring R\<^esub> b) \<oplus>\<^bsub>poly_ring R\<^esub> \<zero>\<^bsub>poly_ring R\<^esub>"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
625  | 
by algebra  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
626  | 
with \<open>b \<in> carrier (poly_ring R)\<close> have "long_divides q p (b, [])"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
627  | 
unfolding long_divides_def univ_poly_zero by auto  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
628  | 
with \<open>p \<noteq> []\<close> have "b \<in> carrier (K[X])"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
629  | 
using long_divisionI[of K q p b] long_division_closed[of K q p] assms  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
630  | 
unfolding univ_poly_carrier by auto  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
631  | 
with \<open>q = p \<otimes>\<^bsub>poly_ring R\<^esub> b\<close> show ?thesis  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
632  | 
unfolding factor_def univ_poly_mult by blast  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
633  | 
qed  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
634  | 
qed  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
635  | 
|
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
636  | 
lemma (in domain) pdivides_iff_shell:  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
637  | 
assumes "subfield K R" and "p \<in> carrier (K[X])" "q \<in> carrier (K[X])"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
638  | 
shows "p pdivides q \<longleftrightarrow> p divides\<^bsub>K[X]\<^esub> q"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
639  | 
using pdivides_iff assms by (simp add: univ_poly_carrier)  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
640  | 
|
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
641  | 
lemma (in domain) pmod_zero_iff_pdivides:  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
642  | 
assumes "subfield K R" and "p \<in> carrier (K[X])" "q \<in> carrier (K[X])"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
643  | 
shows "p pmod q = [] \<longleftrightarrow> q pdivides p"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
644  | 
proof -  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
645  | 
interpret UP: domain "K[X]"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
646  | 
using univ_poly_is_domain[OF subfieldE(1)[OF assms(1)]] .  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
647  | 
|
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
648  | 
show ?thesis  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
649  | 
proof  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
650  | 
assume pmod: "p pmod q = []"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
651  | 
have "p pdiv q \<in> carrier (K[X])" and "p pmod q \<in> carrier (K[X])"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
652  | 
using long_division_closed[OF assms] by auto  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
653  | 
hence "p = q \<otimes>\<^bsub>K[X]\<^esub> (p pdiv q)"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
654  | 
using pdiv_pmod[OF assms] assms(3) unfolding pmod sym[OF univ_poly_zero[of R K]] by algebra  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
655  | 
with \<open>p pdiv q \<in> carrier (K[X])\<close> show "q pdivides p"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
656  | 
unfolding pdivides_iff_shell[OF assms(1,3,2)] factor_def by blast  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
657  | 
next  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
658  | 
assume "q pdivides p" show "p pmod q = []"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
659  | 
proof (cases)  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
660  | 
assume "q = []" with \<open>q pdivides p\<close> show ?thesis  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
661  | 
using zero_pdivides unfolding pmod_def by simp  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
662  | 
next  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
663  | 
assume "q \<noteq> []"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
664  | 
from \<open>q pdivides p\<close> obtain r where "r \<in> carrier (K[X])" and "p = q \<otimes>\<^bsub>K[X]\<^esub> r"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
665  | 
unfolding pdivides_iff_shell[OF assms(1,3,2)] factor_def by blast  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
666  | 
hence "p = (q \<otimes>\<^bsub>K[X]\<^esub> r) \<oplus>\<^bsub>K[X]\<^esub> []"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
667  | 
using assms(2) unfolding sym[OF univ_poly_zero[of R K]] by simp  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
668  | 
moreover from \<open>r \<in> carrier (K[X])\<close> have "r \<in> carrier (poly_ring R)"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
669  | 
using carrier_polynomial_shell[OF subfieldE(1)[OF assms(1)]] by auto  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
670  | 
ultimately have "long_divides p q (r, [])"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
671  | 
unfolding long_divides_def univ_poly_mult univ_poly_add by auto  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
672  | 
with \<open>q \<noteq> []\<close> show ?thesis  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
673  | 
using long_divisionI[OF assms] by simp  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
674  | 
qed  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
675  | 
qed  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
676  | 
qed  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
677  | 
|
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
678  | 
lemma (in domain) same_pmod_iff_pdivides:  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
679  | 
assumes "subfield K R" and "a \<in> carrier (K[X])" "b \<in> carrier (K[X])" "q \<in> carrier (K[X])"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
680  | 
shows "a pmod q = b pmod q \<longleftrightarrow> q pdivides (a \<ominus>\<^bsub>K[X]\<^esub> b)"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
681  | 
proof -  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
682  | 
interpret UP: domain "K[X]"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
683  | 
using univ_poly_is_domain[OF subfieldE(1)[OF assms(1)]] .  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
684  | 
|
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
685  | 
have "a pmod q = b pmod q \<longleftrightarrow> (a \<oplus>\<^bsub>K[X]\<^esub> (\<ominus>\<^bsub>K[X]\<^esub> b)) pmod q = (b \<oplus>\<^bsub>K[X]\<^esub> (\<ominus>\<^bsub>K[X]\<^esub> b)) pmod q"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
686  | 
using long_division_add_iff[OF assms(1-3) UP.a_inv_closed[OF assms(3)] assms(4)] .  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
687  | 
also have " ... \<longleftrightarrow> (a \<ominus>\<^bsub>K[X]\<^esub> b) pmod q = \<zero>\<^bsub>K[X]\<^esub> pmod q"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
688  | 
using assms(2-3) by algebra  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
689  | 
also have " ... \<longleftrightarrow> q pdivides (a \<ominus>\<^bsub>K[X]\<^esub> b)"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
690  | 
using pmod_zero_iff_pdivides[OF assms(1) UP.minus_closed[OF assms(2-3)] assms(4)]  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
691  | 
unfolding univ_poly_zero long_division_zero(2)[OF assms(1,4)] .  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
692  | 
finally show ?thesis .  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
693  | 
qed  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
694  | 
|
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
695  | 
lemma (in domain) pdivides_imp_degree_le:  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
696  | 
assumes "subring K R" and "p \<in> carrier (K[X])" "q \<in> carrier (K[X])" "q \<noteq> []"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
697  | 
shows "p pdivides q \<Longrightarrow> degree p \<le> degree q"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
698  | 
proof -  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
699  | 
assume "p pdivides q"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
700  | 
then obtain r where r: "polynomial (carrier R) r" "q = poly_mult p r"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
701  | 
unfolding pdivides_def factor_def univ_poly_mult univ_poly_carrier by blast  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
702  | 
moreover have p: "polynomial (carrier R) p"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
703  | 
using assms(2) carrier_polynomial[OF assms(1)] unfolding univ_poly_carrier by auto  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
704  | 
moreover have "p \<noteq> []" and "r \<noteq> []"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
705  | 
using poly_mult_zero(2)[OF polynomial_incl[OF p]] r(2) assms(4) by auto  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
706  | 
ultimately show "degree p \<le> degree q"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
707  | 
using poly_mult_degree_eq[OF carrier_is_subring, of p r] by auto  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
708  | 
qed  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
709  | 
|
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
710  | 
lemma (in domain) pprimeE:  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
711  | 
assumes "subfield K R" "p \<in> carrier (K[X])" "pprime K p"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
712  | 
shows "p \<noteq> []" "p \<notin> Units (K[X])"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
713  | 
and "\<And>q r. \<lbrakk> q \<in> carrier (K[X]); r \<in> carrier (K[X])\<rbrakk> \<Longrightarrow>  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
714  | 
p pdivides (q \<otimes>\<^bsub>K[X]\<^esub> r) \<Longrightarrow> p pdivides q \<or> p pdivides r"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
715  | 
using assms(2-3) poly_mult_closed[OF subfieldE(1)[OF assms(1)]] pdivides_iff[OF assms(1)]  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
716  | 
unfolding ring_prime_def prime_def  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
717  | 
by (auto simp add: univ_poly_mult univ_poly_carrier univ_poly_zero)  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
718  | 
|
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
719  | 
lemma (in domain) pprimeI:  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
720  | 
assumes "subfield K R" "p \<in> carrier (K[X])" "p \<noteq> []" "p \<notin> Units (K[X])"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
721  | 
and "\<And>q r. \<lbrakk> q \<in> carrier (K[X]); r \<in> carrier (K[X])\<rbrakk> \<Longrightarrow>  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
722  | 
p pdivides (q \<otimes>\<^bsub>K[X]\<^esub> r) \<Longrightarrow> p pdivides q \<or> p pdivides r"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
723  | 
shows "pprime K p"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
724  | 
using assms(2-5) poly_mult_closed[OF subfieldE(1)[OF assms(1)]] pdivides_iff[OF assms(1)]  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
725  | 
unfolding ring_prime_def prime_def  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
726  | 
by (auto simp add: univ_poly_mult univ_poly_carrier univ_poly_zero)  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
727  | 
|
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
728  | 
lemma (in domain) associated_polynomials_iff:  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
729  | 
assumes "subfield K R" and "p \<in> carrier (K[X])" "q \<in> carrier (K[X])"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
730  | 
  shows "p \<sim>\<^bsub>K[X]\<^esub> q \<longleftrightarrow> (\<exists>k \<in> K - { \<zero> }. p = [ k ] \<otimes>\<^bsub>K[X]\<^esub> q)"
 | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
731  | 
using domain.ring_associated_iff[OF univ_poly_is_domain[OF subfieldE(1)[OF assms(1)]] assms(2-3)]  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
732  | 
unfolding univ_poly_units[OF assms(1)] by auto  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
733  | 
|
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
734  | 
corollary (in domain) associated_polynomials_imp_same_length: (* stronger than "imp_same_degree" *)  | 
| 
70214
 
58191e01f0b1
moving around some material from Algebraic_Closure
 
paulson <lp15@cam.ac.uk> 
parents: 
70162 
diff
changeset
 | 
735  | 
assumes "subring K R" and "p \<in> carrier (K[X])" and "q \<in> carrier (K[X])"  | 
| 
70160
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
736  | 
shows "p \<sim>\<^bsub>K[X]\<^esub> q \<Longrightarrow> length p = length q"  | 
| 
70214
 
58191e01f0b1
moving around some material from Algebraic_Closure
 
paulson <lp15@cam.ac.uk> 
parents: 
70162 
diff
changeset
 | 
737  | 
proof -  | 
| 
 
58191e01f0b1
moving around some material from Algebraic_Closure
 
paulson <lp15@cam.ac.uk> 
parents: 
70162 
diff
changeset
 | 
738  | 
  { fix p q
 | 
| 
 
58191e01f0b1
moving around some material from Algebraic_Closure
 
paulson <lp15@cam.ac.uk> 
parents: 
70162 
diff
changeset
 | 
739  | 
assume p: "p \<in> carrier (K[X])" and q: "q \<in> carrier (K[X])" and "p \<sim>\<^bsub>K[X]\<^esub> q"  | 
| 
 
58191e01f0b1
moving around some material from Algebraic_Closure
 
paulson <lp15@cam.ac.uk> 
parents: 
70162 
diff
changeset
 | 
740  | 
have "length p \<le> length q"  | 
| 
 
58191e01f0b1
moving around some material from Algebraic_Closure
 
paulson <lp15@cam.ac.uk> 
parents: 
70162 
diff
changeset
 | 
741  | 
proof (cases "q = []")  | 
| 
 
58191e01f0b1
moving around some material from Algebraic_Closure
 
paulson <lp15@cam.ac.uk> 
parents: 
70162 
diff
changeset
 | 
742  | 
case True with \<open>p \<sim>\<^bsub>K[X]\<^esub> q\<close> have "p = []"  | 
| 
 
58191e01f0b1
moving around some material from Algebraic_Closure
 
paulson <lp15@cam.ac.uk> 
parents: 
70162 
diff
changeset
 | 
743  | 
unfolding associated_def True factor_def univ_poly_def by auto  | 
| 
 
58191e01f0b1
moving around some material from Algebraic_Closure
 
paulson <lp15@cam.ac.uk> 
parents: 
70162 
diff
changeset
 | 
744  | 
thus ?thesis  | 
| 
 
58191e01f0b1
moving around some material from Algebraic_Closure
 
paulson <lp15@cam.ac.uk> 
parents: 
70162 
diff
changeset
 | 
745  | 
using True by simp  | 
| 
 
58191e01f0b1
moving around some material from Algebraic_Closure
 
paulson <lp15@cam.ac.uk> 
parents: 
70162 
diff
changeset
 | 
746  | 
next  | 
| 
 
58191e01f0b1
moving around some material from Algebraic_Closure
 
paulson <lp15@cam.ac.uk> 
parents: 
70162 
diff
changeset
 | 
747  | 
case False  | 
| 
 
58191e01f0b1
moving around some material from Algebraic_Closure
 
paulson <lp15@cam.ac.uk> 
parents: 
70162 
diff
changeset
 | 
748  | 
from \<open>p \<sim>\<^bsub>K[X]\<^esub> q\<close> have "p divides\<^bsub>K [X]\<^esub> q"  | 
| 
 
58191e01f0b1
moving around some material from Algebraic_Closure
 
paulson <lp15@cam.ac.uk> 
parents: 
70162 
diff
changeset
 | 
749  | 
unfolding associated_def by simp  | 
| 
 
58191e01f0b1
moving around some material from Algebraic_Closure
 
paulson <lp15@cam.ac.uk> 
parents: 
70162 
diff
changeset
 | 
750  | 
hence "p divides\<^bsub>poly_ring R\<^esub> q"  | 
| 
 
58191e01f0b1
moving around some material from Algebraic_Closure
 
paulson <lp15@cam.ac.uk> 
parents: 
70162 
diff
changeset
 | 
751  | 
using carrier_polynomial[OF assms(1)]  | 
| 
 
58191e01f0b1
moving around some material from Algebraic_Closure
 
paulson <lp15@cam.ac.uk> 
parents: 
70162 
diff
changeset
 | 
752  | 
unfolding factor_def univ_poly_carrier univ_poly_mult by auto  | 
| 
 
58191e01f0b1
moving around some material from Algebraic_Closure
 
paulson <lp15@cam.ac.uk> 
parents: 
70162 
diff
changeset
 | 
753  | 
with \<open>q \<noteq> []\<close> have "degree p \<le> degree q"  | 
| 
 
58191e01f0b1
moving around some material from Algebraic_Closure
 
paulson <lp15@cam.ac.uk> 
parents: 
70162 
diff
changeset
 | 
754  | 
using pdivides_imp_degree_le[OF assms(1) p q] unfolding pdivides_def by simp  | 
| 
 
58191e01f0b1
moving around some material from Algebraic_Closure
 
paulson <lp15@cam.ac.uk> 
parents: 
70162 
diff
changeset
 | 
755  | 
with \<open>q \<noteq> []\<close> show ?thesis  | 
| 
 
58191e01f0b1
moving around some material from Algebraic_Closure
 
paulson <lp15@cam.ac.uk> 
parents: 
70162 
diff
changeset
 | 
756  | 
by (cases "p = []", auto simp add: Suc_leI le_diff_iff)  | 
| 
 
58191e01f0b1
moving around some material from Algebraic_Closure
 
paulson <lp15@cam.ac.uk> 
parents: 
70162 
diff
changeset
 | 
757  | 
qed  | 
| 
 
58191e01f0b1
moving around some material from Algebraic_Closure
 
paulson <lp15@cam.ac.uk> 
parents: 
70162 
diff
changeset
 | 
758  | 
} note aux_lemma = this  | 
| 
70160
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
759  | 
|
| 
70214
 
58191e01f0b1
moving around some material from Algebraic_Closure
 
paulson <lp15@cam.ac.uk> 
parents: 
70162 
diff
changeset
 | 
760  | 
interpret UP: domain "K[X]"  | 
| 
 
58191e01f0b1
moving around some material from Algebraic_Closure
 
paulson <lp15@cam.ac.uk> 
parents: 
70162 
diff
changeset
 | 
761  | 
using univ_poly_is_domain[OF assms(1)] .  | 
| 
 
58191e01f0b1
moving around some material from Algebraic_Closure
 
paulson <lp15@cam.ac.uk> 
parents: 
70162 
diff
changeset
 | 
762  | 
|
| 
 
58191e01f0b1
moving around some material from Algebraic_Closure
 
paulson <lp15@cam.ac.uk> 
parents: 
70162 
diff
changeset
 | 
763  | 
assume "p \<sim>\<^bsub>K[X]\<^esub> q" thus ?thesis  | 
| 
 
58191e01f0b1
moving around some material from Algebraic_Closure
 
paulson <lp15@cam.ac.uk> 
parents: 
70162 
diff
changeset
 | 
764  | 
using aux_lemma[OF assms(2-3)] aux_lemma[OF assms(3,2) UP.associated_sym] by simp  | 
| 
 
58191e01f0b1
moving around some material from Algebraic_Closure
 
paulson <lp15@cam.ac.uk> 
parents: 
70162 
diff
changeset
 | 
765  | 
qed  | 
| 
70160
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
766  | 
|
| 
70215
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
767  | 
lemma (in ring) divides_pirreducible_condition:  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
768  | 
assumes "pirreducible K q" and "p \<in> carrier (K[X])"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
769  | 
shows "p divides\<^bsub>K[X]\<^esub> q \<Longrightarrow> p \<in> Units (K[X]) \<or> p \<sim>\<^bsub>K[X]\<^esub> q"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
770  | 
using divides_irreducible_condition[of "K[X]" q p] assms  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
771  | 
unfolding ring_irreducible_def by auto  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
772  | 
|
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
773  | 
subsection \<open>Polynomial Power\<close>  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
774  | 
|
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
775  | 
lemma (in domain) polynomial_pow_not_zero:  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
776  | 
assumes "p \<in> carrier (poly_ring R)" and "p \<noteq> []"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
777  | 
shows "p [^]\<^bsub>poly_ring R\<^esub> (n::nat) \<noteq> []"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
778  | 
proof -  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
779  | 
interpret UP: domain "poly_ring R"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
780  | 
using univ_poly_is_domain[OF carrier_is_subring] .  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
781  | 
|
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
782  | 
from assms UP.integral show ?thesis  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
783  | 
unfolding sym[OF univ_poly_zero[of R "carrier R"]]  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
784  | 
by (induction n, auto)  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
785  | 
qed  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
786  | 
|
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
787  | 
lemma (in domain) subring_polynomial_pow_not_zero:  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
788  | 
assumes "subring K R" and "p \<in> carrier (K[X])" and "p \<noteq> []"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
789  | 
shows "p [^]\<^bsub>K[X]\<^esub> (n::nat) \<noteq> []"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
790  | 
using domain.polynomial_pow_not_zero[OF subring_is_domain, of K p n] assms  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
791  | 
unfolding univ_poly_consistent[OF assms(1)] by simp  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
792  | 
|
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
793  | 
lemma (in domain) polynomial_pow_degree:  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
794  | 
assumes "p \<in> carrier (poly_ring R)"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
795  | 
shows "degree (p [^]\<^bsub>poly_ring R\<^esub> n) = n * degree p"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
796  | 
proof -  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
797  | 
interpret UP: domain "poly_ring R"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
798  | 
using univ_poly_is_domain[OF carrier_is_subring] .  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
799  | 
|
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
800  | 
show ?thesis  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
801  | 
proof (induction n)  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
802  | 
case 0 thus ?case  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
803  | 
using UP.nat_pow_0 unfolding univ_poly_one by auto  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
804  | 
next  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
805  | 
let ?ppow = "\<lambda>n. p [^]\<^bsub>poly_ring R\<^esub> n"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
806  | 
case (Suc n) thus ?case  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
807  | 
proof (cases "p = []")  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
808  | 
case True thus ?thesis  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
809  | 
using univ_poly_zero[of R "carrier R"] UP.r_null assms by auto  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
810  | 
next  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
811  | 
case False  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
812  | 
hence "?ppow n \<in> carrier (poly_ring R)" and "?ppow n \<noteq> []" and "p \<noteq> []"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
813  | 
using polynomial_pow_not_zero[of p n] assms by (auto simp add: univ_poly_one)  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
814  | 
thus ?thesis  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
815  | 
using poly_mult_degree_eq[OF carrier_is_subring, of "?ppow n" p] Suc assms  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
816  | 
unfolding univ_poly_carrier univ_poly_zero  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
817  | 
by (auto simp add: add.commute univ_poly_mult)  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
818  | 
qed  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
819  | 
qed  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
820  | 
qed  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
821  | 
|
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
822  | 
lemma (in domain) subring_polynomial_pow_degree:  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
823  | 
assumes "subring K R" and "p \<in> carrier (K[X])"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
824  | 
shows "degree (p [^]\<^bsub>K[X]\<^esub> n) = n * degree p"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
825  | 
using domain.polynomial_pow_degree[OF subring_is_domain, of K p n] assms  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
826  | 
unfolding univ_poly_consistent[OF assms(1)] by simp  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
827  | 
|
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
828  | 
lemma (in domain) polynomial_pow_division:  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
829  | 
assumes "p \<in> carrier (poly_ring R)" and "(n::nat) \<le> m"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
830  | 
shows "(p [^]\<^bsub>poly_ring R\<^esub> n) pdivides (p [^]\<^bsub>poly_ring R\<^esub> m)"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
831  | 
proof -  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
832  | 
interpret UP: domain "poly_ring R"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
833  | 
using univ_poly_is_domain[OF carrier_is_subring] .  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
834  | 
|
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
835  | 
let ?ppow = "\<lambda>n. p [^]\<^bsub>poly_ring R\<^esub> n"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
836  | 
|
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
837  | 
have "?ppow n \<otimes>\<^bsub>poly_ring R\<^esub> ?ppow k = ?ppow (n + k)" for k  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
838  | 
using assms(1) by (simp add: UP.nat_pow_mult)  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
839  | 
thus ?thesis  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
840  | 
using dividesI[of "?ppow (m - n)" "poly_ring R" "?ppow m" "?ppow n"] assms  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
841  | 
unfolding pdivides_def by auto  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
842  | 
qed  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
843  | 
|
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
844  | 
lemma (in domain) subring_polynomial_pow_division:  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
845  | 
assumes "subring K R" and "p \<in> carrier (K[X])" and "(n::nat) \<le> m"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
846  | 
shows "(p [^]\<^bsub>K[X]\<^esub> n) divides\<^bsub>K[X]\<^esub> (p [^]\<^bsub>K[X]\<^esub> m)"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
847  | 
using domain.polynomial_pow_division[OF subring_is_domain, of K p n m] assms  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
848  | 
unfolding univ_poly_consistent[OF assms(1)] pdivides_def by simp  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
849  | 
|
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
850  | 
lemma (in domain) pirreducible_pow_pdivides_iff:  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
851  | 
assumes "subfield K R" "p \<in> carrier (K[X])" "q \<in> carrier (K[X])" "r \<in> carrier (K[X])"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
852  | 
and "pirreducible K p" and "\<not> (p pdivides q)"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
853  | 
shows "(p [^]\<^bsub>K[X]\<^esub> (n :: nat)) pdivides (q \<otimes>\<^bsub>K[X]\<^esub> r) \<longleftrightarrow> (p [^]\<^bsub>K[X]\<^esub> n) pdivides r"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
854  | 
proof -  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
855  | 
interpret UP: principal_domain "K[X]"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
856  | 
using univ_poly_is_principal[OF assms(1)] .  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
857  | 
show ?thesis  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
858  | 
proof (cases "r = []")  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
859  | 
case True with \<open>q \<in> carrier (K[X])\<close> have "q \<otimes>\<^bsub>K[X]\<^esub> r = []" and "r = []"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
860  | 
unfolding sym[OF univ_poly_zero[of R K]] by auto  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
861  | 
thus ?thesis  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
862  | 
using pdivides_zero[OF subfieldE(1),of K] assms by auto  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
863  | 
next  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
864  | 
case False then have not_zero: "p \<noteq> []" "q \<noteq> []" "r \<noteq> []" "q \<otimes>\<^bsub>K[X]\<^esub> r \<noteq> []"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
865  | 
using subfieldE(1) pdivides_zero[OF _ assms(2)] assms(1-2,5-6) pirreducibleE(1)  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
866  | 
UP.integral_iff[OF assms(3-4)] univ_poly_zero[of R K] by auto  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
867  | 
from \<open>p \<noteq> []\<close>  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
868  | 
have ppow: "p [^]\<^bsub>K[X]\<^esub> (n :: nat) \<noteq> []" "p [^]\<^bsub>K[X]\<^esub> (n :: nat) \<in> carrier (K[X])"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
869  | 
using subring_polynomial_pow_not_zero[OF subfieldE(1)] assms(1-2) by auto  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
870  | 
have not_pdiv: "\<not> (p divides\<^bsub>mult_of (K[X])\<^esub> q)"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
871  | 
using assms(6) pdivides_iff_shell[OF assms(1-3)] unfolding pdivides_def by auto  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
872  | 
have prime: "prime (mult_of (K[X])) p"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
873  | 
using assms(5) pprime_iff_pirreducible[OF assms(1-2)]  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
874  | 
unfolding sym[OF UP.prime_eq_prime_mult[OF assms(2)]] ring_prime_def by simp  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
875  | 
have "a pdivides b \<longleftrightarrow> a divides\<^bsub>mult_of (K[X])\<^esub> b"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
876  | 
if "a \<in> carrier (K[X])" "a \<noteq> \<zero>\<^bsub>K[X]\<^esub>" "b \<in> carrier (K[X])" "b \<noteq> \<zero>\<^bsub>K[X]\<^esub>" for a b  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
877  | 
using that UP.divides_imp_divides_mult[of a b] divides_mult_imp_divides[of "K[X]" a b]  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
878  | 
unfolding pdivides_iff_shell[OF assms(1) that(1,3)] by blast  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
879  | 
thus ?thesis  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
880  | 
using UP.mult_of.prime_pow_divides_iff[OF _ _ _ prime not_pdiv, of r] ppow not_zero assms(2-4)  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
881  | 
unfolding nat_pow_mult_of carrier_mult_of mult_mult_of sym[OF univ_poly_zero[of R K]]  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
882  | 
by (metis DiffI UP.m_closed singletonD)  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
883  | 
qed  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
884  | 
qed  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
885  | 
|
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
886  | 
lemma (in domain) subring_degree_one_imp_pirreducible:  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
887  | 
assumes "subring K R" and "a \<in> Units (R \<lparr> carrier := K \<rparr>)" and "b \<in> K"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
888  | 
shows "pirreducible K [ a, b ]"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
889  | 
proof (rule pirreducibleI[OF assms(1)])  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
890  | 
have "a \<in> K" and "a \<noteq> \<zero>"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
891  | 
using assms(2) subringE(1)[OF assms(1)] unfolding Units_def by auto  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
892  | 
thus "[ a, b ] \<in> carrier (K[X])" and "[ a, b ] \<noteq> []" and "[ a, b ] \<notin> Units (K [X])"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
893  | 
using univ_poly_units_incl[OF assms(1)] assms(2-3)  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
894  | 
unfolding sym[OF univ_poly_carrier] polynomial_def by auto  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
895  | 
next  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
896  | 
interpret UP: domain "K[X]"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
897  | 
using univ_poly_is_domain[OF assms(1)] .  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
898  | 
|
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
899  | 
  { fix q r
 | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
900  | 
assume q: "q \<in> carrier (K[X])" and r: "r \<in> carrier (K[X])" and "[ a, b ] = q \<otimes>\<^bsub>K[X]\<^esub> r"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
901  | 
hence not_zero: "q \<noteq> []" "r \<noteq> []"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
902  | 
by (metis UP.integral_iff list.distinct(1) univ_poly_zero)+  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
903  | 
have "degree (q \<otimes>\<^bsub>K[X]\<^esub> r) = degree q + degree r"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
904  | 
using not_zero poly_mult_degree_eq[OF assms(1)] q r  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
905  | 
by (simp add: univ_poly_carrier univ_poly_mult)  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
906  | 
with sym[OF \<open>[ a, b ] = q \<otimes>\<^bsub>K[X]\<^esub> r\<close>] have "degree q + degree r = 1" and "q \<noteq> []" "r \<noteq> []"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
907  | 
using not_zero by auto  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
908  | 
} note aux_lemma1 = this  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
909  | 
|
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
910  | 
  { fix q r
 | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
911  | 
assume q: "q \<in> carrier (K[X])" "q \<noteq> []" and r: "r \<in> carrier (K[X])" "r \<noteq> []"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
912  | 
and "[ a, b ] = q \<otimes>\<^bsub>K[X]\<^esub> r" and "degree q = 1" and "degree r = 0"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
913  | 
hence "length q = Suc (Suc 0)" and "length r = Suc 0"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
914  | 
by (linarith, metis add.right_neutral add_eq_if length_0_conv)  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
915  | 
from \<open>length q = Suc (Suc 0)\<close> obtain c d where q_def: "q = [ c, d ]"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
916  | 
by (metis length_0_conv length_Cons list.exhaust nat.inject)  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
917  | 
from \<open>length r = Suc 0\<close> obtain e where r_def: "r = [ e ]"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
918  | 
by (metis length_0_conv length_Suc_conv)  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
919  | 
from \<open>r = [ e ]\<close> and \<open>q = [ c, d ]\<close>  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
920  | 
have c: "c \<in> K" "c \<noteq> \<zero>" and d: "d \<in> K" and e: "e \<in> K" "e \<noteq> \<zero>"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
921  | 
using r q subringE(1)[OF assms(1)] unfolding sym[OF univ_poly_carrier] polynomial_def by auto  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
922  | 
with sym[OF \<open>[ a, b ] = q \<otimes>\<^bsub>K[X]\<^esub> r\<close>] have "a = c \<otimes> e"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
923  | 
using poly_mult_lead_coeff[OF assms(1), of q r]  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
924  | 
unfolding polynomial_def sym[OF univ_poly_mult[of R K]] r_def q_def by auto  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
925  | 
obtain inv_a where a: "a \<in> K" and inv_a: "inv_a \<in> K" "a \<otimes> inv_a = \<one>" "inv_a \<otimes> a = \<one>"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
926  | 
using assms(2) unfolding Units_def by auto  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
927  | 
hence "a \<noteq> \<zero>" and "inv_a \<noteq> \<zero>"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
928  | 
using subringE(1)[OF assms(1)] integral_iff by auto  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
929  | 
with \<open>c \<in> K\<close> and \<open>c \<noteq> \<zero>\<close> have in_carrier: "[ c \<otimes> inv_a ] \<in> carrier (K[X])"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
930  | 
using subringE(1,6)[OF assms(1)] inv_a integral  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
931  | 
unfolding sym[OF univ_poly_carrier] polynomial_def  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
932  | 
by (auto, meson subsetD)  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
933  | 
moreover have "[ c \<otimes> inv_a ] \<otimes>\<^bsub>K[X]\<^esub> r = [ \<one> ]"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
934  | 
using \<open>a = c \<otimes> e\<close> a inv_a c e subsetD[OF subringE(1)[OF assms(1)]]  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
935  | 
unfolding r_def univ_poly_mult by (auto) (simp add: m_assoc m_lcomm integral_iff)+  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
936  | 
ultimately have "r \<in> Units (K[X])"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
937  | 
using r(1) UP.m_comm[OF in_carrier r(1)] unfolding sym[OF univ_poly_one[of R K]] Units_def by auto  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
938  | 
} note aux_lemma2 = this  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
939  | 
|
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
940  | 
fix q r  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
941  | 
assume q: "q \<in> carrier (K[X])" and r: "r \<in> carrier (K[X])" and qr: "[ a, b ] = q \<otimes>\<^bsub>K[X]\<^esub> r"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
942  | 
thus "q \<in> Units (K[X]) \<or> r \<in> Units (K[X])"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
943  | 
using aux_lemma1[OF q r qr] aux_lemma2[of q r] aux_lemma2[of r q] UP.m_comm add_is_1 by auto  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
944  | 
qed  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
945  | 
|
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
946  | 
lemma (in domain) degree_one_imp_pirreducible:  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
947  | 
assumes "subfield K R" and "p \<in> carrier (K[X])" and "degree p = 1"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
948  | 
shows "pirreducible K p"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
949  | 
proof -  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
950  | 
from \<open>degree p = 1\<close> have "length p = Suc (Suc 0)"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
951  | 
by simp  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
952  | 
then obtain a b where p: "p = [ a, b ]"  | 
| 72004 | 953  | 
by (metis length_0_conv length_Cons nat.inject neq_Nil_conv)  | 
| 
70215
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
954  | 
with \<open>p \<in> carrier (K[X])\<close> show ?thesis  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
955  | 
using subring_degree_one_imp_pirreducible[OF subfieldE(1)[OF assms(1)], of a b]  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
956  | 
subfield.subfield_Units[OF assms(1)]  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
957  | 
unfolding sym[OF univ_poly_carrier] polynomial_def by auto  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
958  | 
qed  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
959  | 
|
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
960  | 
lemma (in ring) degree_oneE[elim]:  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
961  | 
assumes "p \<in> carrier (K[X])" and "degree p = 1"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
962  | 
and "\<And>a b. \<lbrakk> a \<in> K; a \<noteq> \<zero>; b \<in> K; p = [ a, b ] \<rbrakk> \<Longrightarrow> P"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
963  | 
shows P  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
964  | 
proof -  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
965  | 
from \<open>degree p = 1\<close> have "length p = Suc (Suc 0)"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
966  | 
by simp  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
967  | 
then obtain a b where "p = [ a, b ]"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
968  | 
by (metis length_0_conv length_Cons nat.inject neq_Nil_conv)  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
969  | 
with \<open>p \<in> carrier (K[X])\<close> have "a \<in> K" and "a \<noteq> \<zero>" and "b \<in> K"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
970  | 
unfolding sym[OF univ_poly_carrier] polynomial_def by auto  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
971  | 
with \<open>p = [ a, b ]\<close> show ?thesis  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
972  | 
using assms(3) by simp  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
973  | 
qed  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
974  | 
|
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
975  | 
lemma (in domain) subring_degree_one_associatedI:  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
976  | 
assumes "subring K R" and "a \<in> K" "a' \<in> K" and "b \<in> K" and "a \<otimes> a' = \<one>"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
977  | 
shows "[ a , b ] \<sim>\<^bsub>K[X]\<^esub> [ \<one>, a' \<otimes> b ]"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
978  | 
proof -  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
979  | 
from \<open>a \<otimes> a' = \<one>\<close> have not_zero: "a \<noteq> \<zero>" "a' \<noteq> \<zero>"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
980  | 
using subringE(1)[OF assms(1)] assms(2-3) by auto  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
981  | 
hence "[ a, b ] = [ a ] \<otimes>\<^bsub>K[X]\<^esub> [ \<one>, a' \<otimes> b ]"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
982  | 
using assms(2-4)[THEN subsetD[OF subringE(1)[OF assms(1)]]] assms(5) m_assoc  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
983  | 
unfolding univ_poly_mult by fastforce  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
984  | 
moreover have "[ a, b ] \<in> carrier (K[X])" and "[ \<one>, a' \<otimes> b ] \<in> carrier (K[X])"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
985  | 
using subringE(1,3,6)[OF assms(1)] not_zero one_not_zero assms  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
986  | 
unfolding sym[OF univ_poly_carrier] polynomial_def by auto  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
987  | 
moreover have "[ a ] \<in> Units (K[X])"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
988  | 
proof -  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
989  | 
from \<open>a \<noteq> \<zero>\<close> and \<open>a' \<noteq> \<zero>\<close> have "[ a ] \<in> carrier (K[X])" and "[ a' ] \<in> carrier (K[X])"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
990  | 
using assms(2-3) unfolding sym[OF univ_poly_carrier] polynomial_def by auto  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
991  | 
moreover have "a' \<otimes> a = \<one>"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
992  | 
using subsetD[OF subringE(1)[OF assms(1)]] assms m_comm by simp  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
993  | 
hence "[ a ] \<otimes>\<^bsub>K[X]\<^esub> [ a' ] = [ \<one> ]" and "[ a' ] \<otimes>\<^bsub>K[X]\<^esub> [ a ] = [ \<one> ]"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
994  | 
using assms unfolding univ_poly_mult by auto  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
995  | 
ultimately show ?thesis  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
996  | 
unfolding sym[OF univ_poly_one[of R K]] Units_def by blast  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
997  | 
qed  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
998  | 
ultimately show ?thesis  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
999  | 
using domain.ring_associated_iff[OF univ_poly_is_domain[OF assms(1)]] by blast  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1000  | 
qed  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1001  | 
|
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1002  | 
lemma (in domain) degree_one_associatedI:  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1003  | 
assumes "subfield K R" and "p \<in> carrier (K[X])" and "degree p = 1"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1004  | 
shows "p \<sim>\<^bsub>K[X]\<^esub> [ \<one>, inv (lead_coeff p) \<otimes> (const_term p) ]"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1005  | 
proof -  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1006  | 
from \<open>p \<in> carrier (K[X])\<close> and \<open>degree p = 1\<close>  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1007  | 
obtain a b where "p = [ a, b ]" and "a \<in> K" "a \<noteq> \<zero>" and "b \<in> K"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1008  | 
by auto  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1009  | 
thus ?thesis  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1010  | 
using subring_degree_one_associatedI[OF subfieldE(1)[OF assms(1)]]  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1011  | 
subfield_m_inv[OF assms(1)] subsetD[OF subfieldE(3)[OF assms(1)]]  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1012  | 
unfolding const_term_def  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1013  | 
by auto  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1014  | 
qed  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1015  | 
|
| 
70160
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1016  | 
subsection \<open>Ideals\<close>  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1017  | 
|
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1018  | 
lemma (in domain) exists_unique_gen:  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1019  | 
  assumes "subfield K R" "ideal I (K[X])" "I \<noteq> { [] }"
 | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1020  | 
shows "\<exists>!p \<in> carrier (K[X]). lead_coeff p = \<one> \<and> I = PIdl\<^bsub>K[X]\<^esub> p"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1021  | 
(is "\<exists>!p. ?generator p")  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1022  | 
proof -  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1023  | 
interpret UP: principal_domain "K[X]"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1024  | 
using univ_poly_is_principal[OF assms(1)] .  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1025  | 
obtain q where q: "q \<in> carrier (K[X])" "I = PIdl\<^bsub>K[X]\<^esub> q"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1026  | 
using UP.exists_gen[OF assms(2)] by blast  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1027  | 
hence not_nil: "q \<noteq> []"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1028  | 
using UP.genideal_zero UP.cgenideal_eq_genideal[OF UP.zero_closed] assms(3)  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1029  | 
by (auto simp add: univ_poly_zero)  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1030  | 
  hence "lead_coeff q \<in> K - { \<zero> }"
 | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1031  | 
using q(1) unfolding univ_poly_def polynomial_def by auto  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1032  | 
  hence inv_lc_q: "inv (lead_coeff q) \<in> K - { \<zero> }" "inv (lead_coeff q) \<otimes> lead_coeff q = \<one>"
 | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1033  | 
using subfield_m_inv[OF assms(1)] by auto  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1034  | 
|
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1035  | 
define p where "p = [ inv (lead_coeff q) ] \<otimes>\<^bsub>K[X]\<^esub> q"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1036  | 
have is_poly: "polynomial K [ inv (lead_coeff q) ]" "polynomial K q"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1037  | 
using inv_lc_q(1) q(1) unfolding univ_poly_def polynomial_def by auto  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1038  | 
hence in_carrier: "p \<in> carrier (K[X])"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1039  | 
using UP.m_closed unfolding univ_poly_carrier p_def by simp  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1040  | 
have lc_p: "lead_coeff p = \<one>"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1041  | 
using poly_mult_lead_coeff[OF subfieldE(1)[OF assms(1)] is_poly _ not_nil] inv_lc_q(2)  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1042  | 
unfolding p_def univ_poly_mult[of R K] by simp  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1043  | 
moreover have PIdl_p: "I = PIdl\<^bsub>K[X]\<^esub> p"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1044  | 
using UP.associated_iff_same_ideal[OF in_carrier q(1)] q(2) inv_lc_q(1) p_def  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1045  | 
associated_polynomials_iff[OF assms(1) in_carrier q(1)]  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1046  | 
by auto  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1047  | 
ultimately have "?generator p"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1048  | 
using in_carrier by simp  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1049  | 
|
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1050  | 
moreover  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1051  | 
have "\<And>r. \<lbrakk> r \<in> carrier (K[X]); lead_coeff r = \<one>; I = PIdl\<^bsub>K[X]\<^esub> r \<rbrakk> \<Longrightarrow> r = p"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1052  | 
proof -  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1053  | 
fix r assume r: "r \<in> carrier (K[X])" "lead_coeff r = \<one>" "I = PIdl\<^bsub>K[X]\<^esub> r"  | 
| 
70214
 
58191e01f0b1
moving around some material from Algebraic_Closure
 
paulson <lp15@cam.ac.uk> 
parents: 
70162 
diff
changeset
 | 
1054  | 
have "subring K R"  | 
| 
 
58191e01f0b1
moving around some material from Algebraic_Closure
 
paulson <lp15@cam.ac.uk> 
parents: 
70162 
diff
changeset
 | 
1055  | 
by (simp add: \<open>subfield K R\<close> subfieldE(1))  | 
| 
70160
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1056  | 
    obtain k where k: "k \<in> K - { \<zero> }" "r = [ k ] \<otimes>\<^bsub>K[X]\<^esub> p"
 | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1057  | 
using UP.associated_iff_same_ideal[OF r(1) in_carrier] PIdl_p r(3)  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1058  | 
associated_polynomials_iff[OF assms(1) r(1) in_carrier]  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1059  | 
by auto  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1060  | 
hence "polynomial K [ k ]"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1061  | 
unfolding polynomial_def by simp  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1062  | 
moreover have "p \<noteq> []"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1063  | 
using not_nil UP.associated_iff_same_ideal[OF in_carrier q(1)] q(2) PIdl_p  | 
| 
70214
 
58191e01f0b1
moving around some material from Algebraic_Closure
 
paulson <lp15@cam.ac.uk> 
parents: 
70162 
diff
changeset
 | 
1064  | 
associated_polynomials_imp_same_length[OF \<open>subring K R\<close> in_carrier q(1)] by auto  | 
| 
70160
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1065  | 
ultimately have "lead_coeff r = k \<otimes> (lead_coeff p)"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1066  | 
using poly_mult_lead_coeff[OF subfieldE(1)[OF assms(1)]] in_carrier k(2)  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1067  | 
unfolding univ_poly_def by (auto simp del: poly_mult.simps)  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1068  | 
hence "k = \<one>"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1069  | 
using lc_p r(2) k(1) subfieldE(3)[OF assms(1)] by auto  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1070  | 
hence "r = map ((\<otimes>) \<one>) p"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1071  | 
using poly_mult_const(1)[OF subfieldE(1)[OF assms(1)] _ k(1), of p] in_carrier  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1072  | 
unfolding k(2) univ_poly_carrier[of R K] univ_poly_mult[of R K] by auto  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1073  | 
moreover have "set p \<subseteq> carrier R"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1074  | 
using polynomial_in_carrier[OF subfieldE(1)[OF assms(1)]]  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1075  | 
in_carrier univ_poly_carrier[of R K] by auto  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1076  | 
hence "map ((\<otimes>) \<one>) p = p"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1077  | 
by (induct p) (auto)  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1078  | 
ultimately show "r = p" by simp  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1079  | 
qed  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1080  | 
|
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1081  | 
ultimately show ?thesis by blast  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1082  | 
qed  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1083  | 
|
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1084  | 
proposition (in domain) exists_unique_pirreducible_gen:  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1085  | 
assumes "subfield K R" "ring_hom_ring (K[X]) R h"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1086  | 
    and "a_kernel (K[X]) R h \<noteq> { [] }" "a_kernel (K[X]) R h \<noteq> carrier (K[X])"
 | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1087  | 
shows "\<exists>!p \<in> carrier (K[X]). pirreducible K p \<and> lead_coeff p = \<one> \<and> a_kernel (K[X]) R h = PIdl\<^bsub>K[X]\<^esub> p"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1088  | 
(is "\<exists>!p. ?generator p")  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1089  | 
proof -  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1090  | 
interpret UP: principal_domain "K[X]"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1091  | 
using univ_poly_is_principal[OF assms(1)] .  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1092  | 
|
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1093  | 
have "ideal (a_kernel (K[X]) R h) (K[X])"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1094  | 
using ring_hom_ring.kernel_is_ideal[OF assms(2)] .  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1095  | 
then obtain p  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1096  | 
where p: "p \<in> carrier (K[X])" "lead_coeff p = \<one>" "a_kernel (K[X]) R h = PIdl\<^bsub>K[X]\<^esub> p"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1097  | 
and unique:  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1098  | 
"\<And>q. \<lbrakk> q \<in> carrier (K[X]); lead_coeff q = \<one>; a_kernel (K[X]) R h = PIdl\<^bsub>K[X]\<^esub> q \<rbrakk> \<Longrightarrow> q = p"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1099  | 
using exists_unique_gen[OF assms(1) _ assms(3)] by metis  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1100  | 
|
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1101  | 
  have "p \<in> carrier (K[X]) - { [] }"
 | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1102  | 
using UP.genideal_zero UP.cgenideal_eq_genideal[OF UP.zero_closed] assms(3) p(1,3)  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1103  | 
by (auto simp add: univ_poly_zero)  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1104  | 
hence "pprime K p"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1105  | 
using ring_hom_ring.primeideal_vimage[OF assms(2) UP.is_cring zeroprimeideal]  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1106  | 
UP.primeideal_iff_prime[of p]  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1107  | 
unfolding univ_poly_zero sym[OF p(3)] a_kernel_def' by simp  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1108  | 
hence "pirreducible K p"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1109  | 
using pprime_iff_pirreducible[OF assms(1) p(1)] by simp  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1110  | 
thus ?thesis  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1111  | 
using p unique by metis  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1112  | 
qed  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1113  | 
|
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1114  | 
lemma (in domain) cgenideal_pirreducible:  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1115  | 
assumes "subfield K R" and "p \<in> carrier (K[X])" "pirreducible K p"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1116  | 
shows "\<lbrakk> pirreducible K q; q \<in> PIdl\<^bsub>K[X]\<^esub> p \<rbrakk> \<Longrightarrow> p \<sim>\<^bsub>K[X]\<^esub> q"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1117  | 
proof -  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1118  | 
interpret UP: principal_domain "K[X]"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1119  | 
using univ_poly_is_principal[OF assms(1)] .  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1120  | 
|
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1121  | 
assume q: "pirreducible K q" "q \<in> PIdl\<^bsub>K[X]\<^esub> p"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1122  | 
hence in_carrier: "q \<in> carrier (K[X])"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1123  | 
using additive_subgroup.a_subset[OF ideal.axioms(1)[OF UP.cgenideal_ideal[OF assms(2)]]] by auto  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1124  | 
hence "p divides\<^bsub>K[X]\<^esub> q"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1125  | 
by (meson q assms(2) UP.cgenideal_ideal UP.cgenideal_minimal UP.to_contain_is_to_divide)  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1126  | 
then obtain r where r: "r \<in> carrier (K[X])" "q = p \<otimes>\<^bsub>K[X]\<^esub> r"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1127  | 
by auto  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1128  | 
hence "r \<in> Units (K[X])"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1129  | 
using pirreducibleE(3)[OF _ in_carrier q(1) assms(2) r(1)] subfieldE(1)[OF assms(1)]  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1130  | 
pirreducibleE(2)[OF _ assms(2-3)] by auto  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1131  | 
thus "p \<sim>\<^bsub>K[X]\<^esub> q"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1132  | 
using UP.ring_associated_iff[OF in_carrier assms(2)] r(2) UP.associated_sym  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1133  | 
unfolding UP.m_comm[OF assms(2) r(1)] by auto  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1134  | 
qed  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1135  | 
|
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1136  | 
|
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1137  | 
subsection \<open>Roots and Multiplicity\<close>  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1138  | 
|
| 
70215
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1139  | 
definition (in ring) is_root :: "'a list \<Rightarrow> 'a \<Rightarrow> bool"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1140  | 
where "is_root p x \<longleftrightarrow> (x \<in> carrier R \<and> eval p x = \<zero> \<and> p \<noteq> [])"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1141  | 
|
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1142  | 
definition (in ring) alg_mult :: "'a list \<Rightarrow> 'a \<Rightarrow> nat"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1143  | 
where "alg_mult p x =  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1144  | 
(if p = [] then 0 else  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1145  | 
(if x \<in> carrier R then Greatest (\<lambda> n. ([ \<one>, \<ominus> x ] [^]\<^bsub>poly_ring R\<^esub> n) pdivides p) else 0))"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1146  | 
|
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1147  | 
definition (in ring) roots :: "'a list \<Rightarrow> 'a multiset"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1148  | 
where "roots p = Abs_multiset (alg_mult p)"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1149  | 
|
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1150  | 
definition (in ring) roots_on :: "'a set \<Rightarrow> 'a list \<Rightarrow> 'a multiset"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1151  | 
where "roots_on K p = roots p \<inter># mset_set K"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1152  | 
|
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1153  | 
definition (in ring) splitted :: "'a list \<Rightarrow> bool"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1154  | 
where "splitted p \<longleftrightarrow> size (roots p) = degree p"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1155  | 
|
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1156  | 
definition (in ring) splitted_on :: "'a set \<Rightarrow> 'a list \<Rightarrow> bool"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1157  | 
where "splitted_on K p \<longleftrightarrow> size (roots_on K p) = degree p"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1158  | 
|
| 
70160
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1159  | 
lemma (in domain) pdivides_imp_root_sharing:  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1160  | 
assumes "p \<in> carrier (poly_ring R)" "p pdivides q" and "a \<in> carrier R"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1161  | 
shows "eval p a = \<zero> \<Longrightarrow> eval q a = \<zero>"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1162  | 
proof -  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1163  | 
from \<open>p pdivides q\<close> obtain r where r: "q = p \<otimes>\<^bsub>poly_ring R\<^esub> r" "r \<in> carrier (poly_ring R)"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1164  | 
unfolding pdivides_def factor_def by auto  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1165  | 
hence "eval q a = (eval p a) \<otimes> (eval r a)"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1166  | 
using ring_hom_memE(2)[OF eval_is_hom[OF carrier_is_subring assms(3)] assms(1) r(2)] by simp  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1167  | 
thus "eval p a = \<zero> \<Longrightarrow> eval q a = \<zero>"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1168  | 
using ring_hom_memE(1)[OF eval_is_hom[OF carrier_is_subring assms(3)] r(2)] by auto  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1169  | 
qed  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1170  | 
|
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1171  | 
lemma (in domain) degree_one_root:  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1172  | 
assumes "subfield K R" and "p \<in> carrier (K[X])" and "degree p = 1"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1173  | 
shows "eval p (\<ominus> (inv (lead_coeff p) \<otimes> (const_term p))) = \<zero>"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1174  | 
and "inv (lead_coeff p) \<otimes> (const_term p) \<in> K"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1175  | 
proof -  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1176  | 
from \<open>degree p = 1\<close> have "length p = Suc (Suc 0)"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1177  | 
by simp  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1178  | 
then obtain a b where p: "p = [ a, b ]"  | 
| 
73932
 
fd21b4a93043
added opaque_combs and renamed hide_lams to opaque_lifting
 
desharna 
parents: 
73270 
diff
changeset
 | 
1179  | 
by (metis (no_types, opaque_lifting) Suc_length_conv length_0_conv)  | 
| 
70160
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1180  | 
  hence "a \<in> K - { \<zero> }" "b \<in> K"  and in_carrier: "a \<in> carrier R" "b \<in> carrier R"
 | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1181  | 
using assms(2) subfieldE(3)[OF assms(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
 | 
1182  | 
hence inv_a: "inv a \<in> carrier R" "a \<otimes> inv a = \<one>" and "inv a \<in> K"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1183  | 
using subfield_m_inv(1-2)[OF assms(1), of a] subfieldE(3)[OF assms(1)] by auto  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1184  | 
hence "eval p (\<ominus> (inv a \<otimes> b)) = a \<otimes> (\<ominus> (inv a \<otimes> b)) \<oplus> b"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1185  | 
using in_carrier unfolding p by simp  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1186  | 
also have " ... = \<ominus> (a \<otimes> (inv a \<otimes> b)) \<oplus> b"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1187  | 
using inv_a in_carrier by (simp add: r_minus)  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1188  | 
also have " ... = \<zero>"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1189  | 
using in_carrier(2) unfolding sym[OF m_assoc[OF in_carrier(1) inv_a(1) in_carrier(2)]] inv_a(2) by algebra  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1190  | 
finally have "eval p (\<ominus> (inv a \<otimes> b)) = \<zero>" .  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1191  | 
moreover have ct: "const_term p = b"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1192  | 
using in_carrier unfolding p const_term_def by auto  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1193  | 
ultimately show "eval p (\<ominus> (inv (lead_coeff p) \<otimes> (const_term p))) = \<zero>"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1194  | 
unfolding p by simp  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1195  | 
from \<open>inv a \<in> K\<close> and \<open>b \<in> K\<close>  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1196  | 
show "inv (lead_coeff p) \<otimes> (const_term p) \<in> K"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1197  | 
using p subringE(6)[OF subfieldE(1)[OF assms(1)]] unfolding ct by auto  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1198  | 
qed  | 
| 
70215
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1199  | 
lemma (in domain) is_root_imp_pdivides:  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1200  | 
assumes "p \<in> carrier (poly_ring R)"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1201  | 
shows "is_root p x \<Longrightarrow> [ \<one>, \<ominus> x ] pdivides p"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1202  | 
proof -  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1203  | 
let ?b = "[ \<one> , \<ominus> x ]"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1204  | 
|
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1205  | 
interpret UP: domain "poly_ring R"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1206  | 
using univ_poly_is_domain[OF carrier_is_subring] .  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1207  | 
|
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1208  | 
assume "is_root p x" hence x: "x \<in> carrier R" and is_root: "eval p x = \<zero>"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1209  | 
unfolding is_root_def by auto  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1210  | 
hence b: "?b \<in> carrier (poly_ring R)"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1211  | 
unfolding sym[OF univ_poly_carrier] polynomial_def by auto  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1212  | 
then obtain q r where q: "q \<in> carrier (poly_ring R)" and r: "r \<in> carrier (poly_ring R)"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1213  | 
and long_divides: "p = (?b \<otimes>\<^bsub>poly_ring R\<^esub> q) \<oplus>\<^bsub>poly_ring R\<^esub> r" "r = [] \<or> degree r < degree ?b"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1214  | 
using long_division_theorem[OF carrier_is_subring, of p ?b] assms by (auto simp add: univ_poly_carrier)  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1215  | 
|
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1216  | 
show ?thesis  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1217  | 
proof (cases "r = []")  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1218  | 
case True then have "r = \<zero>\<^bsub>poly_ring R\<^esub>"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1219  | 
unfolding univ_poly_zero[of R "carrier R"] .  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1220  | 
thus ?thesis  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1221  | 
using long_divides(1) q r b dividesI[OF q, of p ?b] by (simp add: pdivides_def)  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1222  | 
next  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1223  | 
case False then have "length r = Suc 0"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1224  | 
using long_divides(2) le_SucE by fastforce  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1225  | 
then obtain a where "r = [ a ]" and a: "a \<in> carrier R" and "a \<noteq> \<zero>"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1226  | 
using r unfolding sym[OF univ_poly_carrier] polynomial_def  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1227  | 
by (metis False length_0_conv length_Suc_conv list.sel(1) list.set_sel(1) subset_code(1))  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1228  | 
|
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1229  | 
have "eval p x = ((eval ?b x) \<otimes> (eval q x)) \<oplus> (eval r x)"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1230  | 
using long_divides(1) ring_hom_memE[OF eval_is_hom[OF carrier_is_subring x]] by (simp add: b q r)  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1231  | 
also have " ... = eval r x"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1232  | 
using ring_hom_memE[OF eval_is_hom[OF carrier_is_subring x]] x b q r by (auto, algebra)  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1233  | 
finally have "a = \<zero>"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1234  | 
using a unfolding \<open>r = [ a ]\<close> is_root by simp  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1235  | 
with \<open>a \<noteq> \<zero>\<close> have False .. thus ?thesis ..  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1236  | 
qed  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1237  | 
qed  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1238  | 
|
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1239  | 
lemma (in domain) pdivides_imp_is_root:  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1240  | 
assumes "p \<noteq> []" and "x \<in> carrier R"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1241  | 
shows "[ \<one>, \<ominus> x ] pdivides p \<Longrightarrow> is_root p x"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1242  | 
proof -  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1243  | 
assume "[ \<one>, \<ominus> x ] pdivides p"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1244  | 
then obtain q where q: "q \<in> carrier (poly_ring R)" and pdiv: "p = [ \<one>, \<ominus> x ] \<otimes>\<^bsub>poly_ring R\<^esub> q"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1245  | 
unfolding pdivides_def by auto  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1246  | 
moreover have "[ \<one>, \<ominus> x ] \<in> carrier (poly_ring R)"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1247  | 
using assms(2) unfolding sym[OF univ_poly_carrier] polynomial_def by simp  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1248  | 
ultimately have "eval p x = \<zero>"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1249  | 
using ring_hom_memE[OF eval_is_hom[OF carrier_is_subring, of x]] assms(2) by (auto, algebra)  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1250  | 
with \<open>p \<noteq> []\<close> and \<open>x \<in> carrier R\<close> show "is_root p x"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1251  | 
unfolding is_root_def by simp  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1252  | 
qed  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1253  | 
|
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1254  | 
lemma (in domain) associated_polynomials_imp_same_is_root:  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1255  | 
assumes "p \<in> carrier (poly_ring R)" and "q \<in> carrier (poly_ring R)" and "p \<sim>\<^bsub>poly_ring R\<^esub> q"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1256  | 
shows "is_root p x \<longleftrightarrow> is_root q x"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1257  | 
proof (cases "p = []")  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1258  | 
case True with \<open>p \<sim>\<^bsub>poly_ring R\<^esub> q\<close> have "q = []"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1259  | 
unfolding associated_def True factor_def univ_poly_def by auto  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1260  | 
thus ?thesis  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1261  | 
using True unfolding is_root_def by simp  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1262  | 
next  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1263  | 
case False  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1264  | 
interpret UP: domain "poly_ring R"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1265  | 
using univ_poly_is_domain[OF carrier_is_subring] .  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1266  | 
|
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1267  | 
  { fix p q
 | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1268  | 
assume p: "p \<in> carrier (poly_ring R)" and q: "q \<in> carrier (poly_ring R)" and pq: "p \<sim>\<^bsub>poly_ring R\<^esub> q"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1269  | 
have "is_root p x \<Longrightarrow> is_root q x"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1270  | 
proof -  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1271  | 
assume is_root: "is_root p x"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1272  | 
then have "[ \<one>, \<ominus> x ] pdivides p" and "p \<noteq> []" and "x \<in> carrier R"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1273  | 
using is_root_imp_pdivides[OF p] unfolding is_root_def by auto  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1274  | 
moreover have "[ \<one>, \<ominus> x ] \<in> carrier (poly_ring R)"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1275  | 
using is_root unfolding is_root_def sym[OF univ_poly_carrier] polynomial_def by simp  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1276  | 
ultimately have "[ \<one>, \<ominus> x ] pdivides q"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1277  | 
using UP.divides_cong_r[OF _ pq ] unfolding pdivides_def by simp  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1278  | 
with \<open>p \<noteq> []\<close> and \<open>x \<in> carrier R\<close> show ?thesis  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1279  | 
using associated_polynomials_imp_same_length[OF carrier_is_subring p q pq]  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1280  | 
pdivides_imp_is_root[of q x]  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1281  | 
by fastforce  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1282  | 
qed  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1283  | 
}  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1284  | 
|
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1285  | 
then show ?thesis  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1286  | 
using assms UP.associated_sym[OF assms(3)] by blast  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1287  | 
qed  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1288  | 
|
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1289  | 
lemma (in ring) monic_degree_one_root_condition:  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1290  | 
assumes "a \<in> carrier R" shows "is_root [ \<one>, \<ominus> a ] b \<longleftrightarrow> a = b"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1291  | 
using assms minus_equality r_neg[OF assms] unfolding is_root_def by (auto, fastforce)  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1292  | 
|
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1293  | 
lemma (in field) degree_one_root_condition:  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1294  | 
assumes "p \<in> carrier (poly_ring R)" and "degree p = 1"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1295  | 
shows "is_root p x \<longleftrightarrow> x = \<ominus> (inv (lead_coeff p) \<otimes> (const_term p))"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1296  | 
proof -  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1297  | 
interpret UP: domain "poly_ring R"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1298  | 
using univ_poly_is_domain[OF carrier_is_subring] .  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1299  | 
|
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1300  | 
from \<open>degree p = 1\<close> have "length p = Suc (Suc 0)"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1301  | 
by simp  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1302  | 
then obtain a b where p: "p = [ a, b ]"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1303  | 
by (metis length_0_conv length_Cons list.exhaust nat.inject)  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1304  | 
hence a: "a \<in> carrier R" "a \<noteq> \<zero>" and b: "b \<in> carrier R"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1305  | 
using assms(1) unfolding sym[OF univ_poly_carrier] polynomial_def by auto  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1306  | 
hence inv_a: "inv a \<in> carrier R" "(inv a) \<otimes> a = \<one>"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1307  | 
using subfield_m_inv[OF carrier_is_subfield, of a] by auto  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1308  | 
hence in_carrier: "[ \<one>, (inv a) \<otimes> b ] \<in> carrier (poly_ring R)"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1309  | 
using b unfolding sym[OF univ_poly_carrier] polynomial_def by auto  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1310  | 
|
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1311  | 
have "p \<sim>\<^bsub>poly_ring R\<^esub> [ \<one>, (inv a) \<otimes> b ]"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1312  | 
proof (rule UP.associatedI2'[OF _ _ in_carrier, of _ "[ a ]"])  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1313  | 
have "p = [ a ] \<otimes>\<^bsub>poly_ring R\<^esub> [ \<one>, inv a \<otimes> b ]"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1314  | 
using a inv_a b m_assoc[of a "inv a" b] unfolding p univ_poly_mult by (auto, algebra)  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1315  | 
also have " ... = [ \<one>, inv a \<otimes> b ] \<otimes>\<^bsub>poly_ring R\<^esub> [ a ]"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1316  | 
using UP.m_comm[OF in_carrier, of "[ a ]"] a  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1317  | 
by (auto simp add: sym[OF univ_poly_carrier] polynomial_def)  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1318  | 
finally show "p = [ \<one>, inv a \<otimes> b ] \<otimes>\<^bsub>poly_ring R\<^esub> [ a ]" .  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1319  | 
next  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1320  | 
from \<open>a \<in> carrier R\<close> and \<open>a \<noteq> \<zero>\<close> show "[ a ] \<in> Units (poly_ring R)"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1321  | 
unfolding univ_poly_units[OF carrier_is_subfield] by simp  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1322  | 
qed  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1323  | 
|
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1324  | 
moreover have "(inv a) \<otimes> b = \<ominus> (\<ominus> (inv (lead_coeff p) \<otimes> (const_term p)))"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1325  | 
and "inv (lead_coeff p) \<otimes> (const_term p) \<in> carrier R"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1326  | 
using inv_a a b unfolding p const_term_def by auto  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1327  | 
|
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1328  | 
ultimately show ?thesis  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1329  | 
using associated_polynomials_imp_same_is_root[OF assms(1) in_carrier]  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1330  | 
monic_degree_one_root_condition  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1331  | 
by (metis add.inv_closed)  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1332  | 
qed  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1333  | 
|
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1334  | 
lemma (in domain) is_root_poly_mult_imp_is_root:  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1335  | 
assumes "p \<in> carrier (poly_ring R)" and "q \<in> carrier (poly_ring R)"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1336  | 
shows "is_root (p \<otimes>\<^bsub>poly_ring R\<^esub> q) x \<Longrightarrow> (is_root p x) \<or> (is_root q x)"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1337  | 
proof -  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1338  | 
interpret UP: domain "poly_ring R"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1339  | 
using univ_poly_is_domain[OF carrier_is_subring] .  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1340  | 
|
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1341  | 
assume is_root: "is_root (p \<otimes>\<^bsub>poly_ring R\<^esub> q) x"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1342  | 
hence "p \<noteq> []" and "q \<noteq> []"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1343  | 
unfolding is_root_def sym[OF univ_poly_zero[of R "carrier R"]]  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1344  | 
using UP.l_null[OF assms(2)] UP.r_null[OF assms(1)] by blast+  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1345  | 
moreover have x: "x \<in> carrier R" and "eval (p \<otimes>\<^bsub>poly_ring R\<^esub> q) x = \<zero>"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1346  | 
using is_root unfolding is_root_def by simp+  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1347  | 
hence "eval p x = \<zero> \<or> eval q x = \<zero>"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1348  | 
using ring_hom_memE[OF eval_is_hom[OF carrier_is_subring], of x] assms integral by auto  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1349  | 
ultimately show "(is_root p x) \<or> (is_root q x)"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1350  | 
using x unfolding is_root_def by auto  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1351  | 
qed  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1352  | 
|
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1353  | 
lemma (in domain) degree_zero_imp_not_is_root:  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1354  | 
assumes "p \<in> carrier (poly_ring R)" and "degree p = 0" shows "\<not> is_root p x"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1355  | 
proof (cases "p = []", simp add: is_root_def)  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1356  | 
case False with \<open>degree p = 0\<close> have "length p = Suc 0"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1357  | 
using le_SucE by fastforce  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1358  | 
then obtain a where "p = [ a ]" and "a \<in> carrier R" and "a \<noteq> \<zero>"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1359  | 
using assms unfolding sym[OF univ_poly_carrier] polynomial_def  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1360  | 
by (metis False length_0_conv length_Suc_conv list.sel(1) list.set_sel(1) subset_code(1))  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1361  | 
thus ?thesis  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1362  | 
unfolding is_root_def by auto  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1363  | 
qed  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1364  | 
|
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1365  | 
lemma (in domain) finite_number_of_roots:  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1366  | 
  assumes "p \<in> carrier (poly_ring R)" shows "finite { x. is_root p x }"
 | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1367  | 
using assms  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1368  | 
proof (induction "degree p" arbitrary: p)  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1369  | 
case 0 thus ?case  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1370  | 
by (simp add: degree_zero_imp_not_is_root)  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1371  | 
next  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1372  | 
case (Suc n) show ?case  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1373  | 
  proof (cases "{ x. is_root p x } = {}")
 | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1374  | 
case True thus ?thesis  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1375  | 
by (simp add: True)  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1376  | 
next  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1377  | 
interpret UP: domain "poly_ring R"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1378  | 
using univ_poly_is_domain[OF carrier_is_subring] .  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1379  | 
|
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1380  | 
case False  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1381  | 
then obtain a where is_root: "is_root p a"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1382  | 
by blast  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1383  | 
hence a: "a \<in> carrier R" and eval: "eval p a = \<zero>" and p_not_zero: "p \<noteq> []"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1384  | 
unfolding is_root_def by auto  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1385  | 
hence in_carrier: "[ \<one>, \<ominus> a ] \<in> carrier (poly_ring R)"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1386  | 
unfolding sym[OF univ_poly_carrier] polynomial_def by auto  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1387  | 
|
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1388  | 
obtain q where q: "q \<in> carrier (poly_ring R)" and p: "p = [ \<one>, \<ominus> a ] \<otimes>\<^bsub>poly_ring R\<^esub> q"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1389  | 
using is_root_imp_pdivides[OF Suc(3) is_root] unfolding pdivides_def by auto  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1390  | 
with \<open>p \<noteq> []\<close> have q_not_zero: "q \<noteq> []"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1391  | 
using UP.r_null UP.integral in_carrier unfolding sym[OF univ_poly_zero[of R "carrier R"]]  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1392  | 
by metis  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1393  | 
hence "degree q = n"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1394  | 
using poly_mult_degree_eq[OF carrier_is_subring, of "[ \<one>, \<ominus> a ]" q]  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1395  | 
in_carrier q p_not_zero p Suc(2)  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1396  | 
unfolding univ_poly_carrier  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1397  | 
by (metis One_nat_def Suc_eq_plus1 diff_Suc_1 list.distinct(1)  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1398  | 
list.size(3-4) plus_1_eq_Suc univ_poly_mult)  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1399  | 
    hence "finite { x. is_root q x }"
 | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1400  | 
using Suc(1)[OF _ q] by simp  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1401  | 
|
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1402  | 
    moreover have "{ x. is_root p x } \<subseteq> insert a { x. is_root q x }"
 | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1403  | 
using is_root_poly_mult_imp_is_root[OF in_carrier q]  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1404  | 
monic_degree_one_root_condition[OF a]  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1405  | 
unfolding p by auto  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1406  | 
|
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1407  | 
ultimately show ?thesis  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1408  | 
using finite_subset by auto  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1409  | 
qed  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1410  | 
qed  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1411  | 
|
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1412  | 
lemma (in domain) alg_multE:  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1413  | 
assumes "x \<in> carrier R" and "p \<in> carrier (poly_ring R)" and "p \<noteq> []"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1414  | 
shows "([ \<one>, \<ominus> x ] [^]\<^bsub>poly_ring R\<^esub> (alg_mult p x)) pdivides p"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1415  | 
and "\<And>n. ([ \<one>, \<ominus> x ] [^]\<^bsub>poly_ring R\<^esub> n) pdivides p \<Longrightarrow> n \<le> alg_mult p x"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1416  | 
proof -  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1417  | 
interpret UP: domain "poly_ring R"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1418  | 
using univ_poly_is_domain[OF carrier_is_subring] .  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1419  | 
|
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1420  | 
let ?ppow = "\<lambda>n :: nat. ([ \<one>, \<ominus> x ] [^]\<^bsub>poly_ring R\<^esub> n)"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1421  | 
|
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1422  | 
  define S :: "nat set" where "S = { n. ?ppow n pdivides p }"
 | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1423  | 
have "?ppow 0 = \<one>\<^bsub>poly_ring R\<^esub>"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1424  | 
using UP.nat_pow_0 by simp  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1425  | 
hence "0 \<in> S"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1426  | 
using UP.one_divides[OF assms(2)] unfolding S_def pdivides_def by simp  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1427  | 
  hence "S \<noteq> {}"
 | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1428  | 
by auto  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1429  | 
|
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1430  | 
moreover have "n \<le> degree p" if "n \<in> S" for n :: nat  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1431  | 
proof -  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1432  | 
have "[ \<one>, \<ominus> x ] \<in> carrier (poly_ring R)"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1433  | 
using assms unfolding sym[OF univ_poly_carrier] polynomial_def by auto  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1434  | 
hence "?ppow n \<in> carrier (poly_ring R)"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1435  | 
using assms unfolding univ_poly_zero by auto  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1436  | 
with \<open>n \<in> S\<close> have "degree (?ppow n) \<le> degree p"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1437  | 
using pdivides_imp_degree_le[OF carrier_is_subring _ assms(2-3), of "?ppow n"] by (simp add: S_def)  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1438  | 
with \<open>[ \<one>, \<ominus> x ] \<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
 | 
1439  | 
using polynomial_pow_degree by simp  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1440  | 
qed  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1441  | 
hence "finite S"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1442  | 
using finite_nat_set_iff_bounded_le by blast  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1443  | 
|
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1444  | 
ultimately have MaxS: "\<And>n. n \<in> S \<Longrightarrow> n \<le> Max S" "Max S \<in> S"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1445  | 
using Max_ge[of S] Max_in[of S] by auto  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1446  | 
with \<open>x \<in> carrier R\<close> have "alg_mult p x = Max S"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1447  | 
using Greatest_equality[of "\<lambda>n. ?ppow n pdivides p" "Max S"] assms(3)  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1448  | 
unfolding S_def alg_mult_def by auto  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1449  | 
thus "([ \<one>, \<ominus> x ] [^]\<^bsub>poly_ring R\<^esub> (alg_mult p x)) pdivides p"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1450  | 
and "\<And>n. ([ \<one>, \<ominus> x ] [^]\<^bsub>poly_ring R\<^esub> n) pdivides p \<Longrightarrow> n \<le> alg_mult p x"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1451  | 
using MaxS unfolding S_def by auto  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1452  | 
qed  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1453  | 
|
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1454  | 
lemma (in domain) le_alg_mult_imp_pdivides:  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1455  | 
assumes "x \<in> carrier R" and "p \<in> carrier (poly_ring R)"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1456  | 
shows "n \<le> alg_mult p x \<Longrightarrow> ([ \<one>, \<ominus> x ] [^]\<^bsub>poly_ring R\<^esub> n) pdivides p"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1457  | 
proof -  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1458  | 
interpret UP: domain "poly_ring R"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1459  | 
using univ_poly_is_domain[OF carrier_is_subring] .  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1460  | 
|
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1461  | 
assume le_alg_mult: "n \<le> alg_mult p x"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1462  | 
have in_carrier: "[ \<one>, \<ominus> x ] \<in> carrier (poly_ring R)"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1463  | 
using assms(1) unfolding sym[OF univ_poly_carrier] polynomial_def by auto  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1464  | 
hence ppow_pdivides:  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1465  | 
"([ \<one>, \<ominus> x ] [^]\<^bsub>poly_ring R\<^esub> n) pdivides  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1466  | 
([ \<one>, \<ominus> x ] [^]\<^bsub>poly_ring R\<^esub> (alg_mult p x))"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1467  | 
using polynomial_pow_division[OF _ le_alg_mult] by simp  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1468  | 
|
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1469  | 
show ?thesis  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1470  | 
proof (cases "p = []")  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1471  | 
case True thus ?thesis  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1472  | 
using in_carrier pdivides_zero[OF carrier_is_subring] by auto  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1473  | 
next  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1474  | 
case False thus ?thesis  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1475  | 
using ppow_pdivides UP.divides_trans UP.nat_pow_closed alg_multE(1)[OF assms] in_carrier  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1476  | 
unfolding pdivides_def by meson  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1477  | 
qed  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1478  | 
qed  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1479  | 
|
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1480  | 
lemma (in domain) alg_mult_gt_zero_iff_is_root:  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1481  | 
assumes "p \<in> carrier (poly_ring R)" shows "alg_mult p x > 0 \<longleftrightarrow> is_root p x"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1482  | 
proof -  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1483  | 
interpret UP: domain "poly_ring R"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1484  | 
using univ_poly_is_domain[OF carrier_is_subring] .  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1485  | 
show ?thesis  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1486  | 
proof  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1487  | 
assume is_root: "is_root p x" hence x: "x \<in> carrier R" and not_zero: "p \<noteq> []"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1488  | 
unfolding is_root_def by auto  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1489  | 
have "[\<one>, \<ominus> x] [^]\<^bsub>poly_ring R\<^esub> (Suc 0) = [\<one>, \<ominus> x]"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1490  | 
using x unfolding univ_poly_def by auto  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1491  | 
thus "alg_mult p x > 0"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1492  | 
using is_root_imp_pdivides[OF _ is_root] alg_multE(2)[OF x, of p "Suc 0"] not_zero assms by auto  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1493  | 
next  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1494  | 
assume gt_zero: "alg_mult p x > 0"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1495  | 
hence x: "x \<in> carrier R" and not_zero: "p \<noteq> []"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1496  | 
unfolding alg_mult_def by (cases "p = []", auto, cases "x \<in> carrier R", auto)  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1497  | 
hence in_carrier: "[ \<one>, \<ominus> x ] \<in> carrier (poly_ring R)"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1498  | 
unfolding sym[OF univ_poly_carrier] polynomial_def by auto  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1499  | 
with \<open>x \<in> carrier R\<close> have "[ \<one>, \<ominus> x ] pdivides p" and "eval [ \<one>, \<ominus> x ] x = \<zero>"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1500  | 
using le_alg_mult_imp_pdivides[of x p "1::nat"] gt_zero assms by (auto, algebra)  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1501  | 
thus "is_root p x"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1502  | 
using pdivides_imp_root_sharing[OF in_carrier] not_zero x by (simp add: is_root_def)  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1503  | 
qed  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1504  | 
qed  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1505  | 
|
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1506  | 
lemma (in domain) alg_mult_eq_count_roots:  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1507  | 
assumes "p \<in> carrier (poly_ring R)" shows "alg_mult p = count (roots p)"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1508  | 
using finite_number_of_roots[OF assms]  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1509  | 
unfolding sym[OF alg_mult_gt_zero_iff_is_root[OF assms]]  | 
| 73270 | 1510  | 
by (simp add: roots_def)  | 
| 
70215
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1511  | 
|
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1512  | 
lemma (in domain) roots_mem_iff_is_root:  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1513  | 
assumes "p \<in> carrier (poly_ring R)" shows "x \<in># roots p \<longleftrightarrow> is_root p x"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1514  | 
using alg_mult_eq_count_roots[OF assms] count_greater_zero_iff  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1515  | 
unfolding roots_def sym[OF alg_mult_gt_zero_iff_is_root[OF assms]] by metis  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1516  | 
|
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1517  | 
lemma (in domain) degree_zero_imp_empty_roots:  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1518  | 
  assumes "p \<in> carrier (poly_ring R)" and "degree p = 0" shows "roots p = {#}"
 | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1519  | 
using degree_zero_imp_not_is_root[of p] roots_mem_iff_is_root[of p] assms by auto  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1520  | 
|
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1521  | 
lemma (in domain) degree_zero_imp_splitted:  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1522  | 
assumes "p \<in> carrier (poly_ring R)" and "degree p = 0" shows "splitted p"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1523  | 
unfolding splitted_def degree_zero_imp_empty_roots[OF assms] assms(2) by simp  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1524  | 
|
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1525  | 
lemma (in domain) roots_inclI':  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1526  | 
assumes "p \<in> carrier (poly_ring R)" and "\<And>a. \<lbrakk> a \<in> carrier R; p \<noteq> [] \<rbrakk> \<Longrightarrow> alg_mult p a \<le> count m a"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1527  | 
shows "roots p \<subseteq># m"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1528  | 
proof (intro mset_subset_eqI)  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1529  | 
fix a show "count (roots p) a \<le> count m a"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1530  | 
using assms unfolding sym[OF alg_mult_eq_count_roots[OF assms(1)]] alg_mult_def  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1531  | 
by (cases "p = []", simp, cases "a \<in> carrier R", auto)  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1532  | 
qed  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1533  | 
|
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1534  | 
lemma (in domain) roots_inclI:  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1535  | 
assumes "p \<in> carrier (poly_ring R)" and "q \<in> carrier (poly_ring R)" "q \<noteq> []"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1536  | 
and "\<And>a. \<lbrakk> a \<in> carrier R; p \<noteq> [] \<rbrakk> \<Longrightarrow> ([ \<one>, \<ominus> a ] [^]\<^bsub>poly_ring R\<^esub> (alg_mult p a)) pdivides q"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1537  | 
shows "roots p \<subseteq># roots q"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1538  | 
using roots_inclI'[OF assms(1), of "roots q"] assms alg_multE(2)[OF _ assms(2-3)]  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1539  | 
unfolding sym[OF alg_mult_eq_count_roots[OF assms(2)]] by auto  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1540  | 
|
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1541  | 
lemma (in domain) pdivides_imp_roots_incl:  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1542  | 
assumes "p \<in> carrier (poly_ring R)" and "q \<in> carrier (poly_ring R)" "q \<noteq> []"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1543  | 
shows "p pdivides q \<Longrightarrow> roots p \<subseteq># roots q"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1544  | 
proof (rule roots_inclI[OF assms])  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1545  | 
interpret UP: domain "poly_ring R"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1546  | 
using univ_poly_is_domain[OF carrier_is_subring] .  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1547  | 
|
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1548  | 
fix a assume "p pdivides q" and a: "a \<in> carrier R"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1549  | 
hence "[ \<one> , \<ominus> a ] \<in> carrier (poly_ring R)"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1550  | 
unfolding sym[OF univ_poly_carrier] polynomial_def by simp  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1551  | 
with \<open>p pdivides q\<close> show "([\<one>, \<ominus> a] [^]\<^bsub>poly_ring R\<^esub> (alg_mult p a)) pdivides q"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1552  | 
using UP.divides_trans[of _p q] le_alg_mult_imp_pdivides[OF a assms(1)]  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1553  | 
by (auto simp add: pdivides_def)  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1554  | 
qed  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1555  | 
|
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1556  | 
lemma (in domain) associated_polynomials_imp_same_roots:  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1557  | 
assumes "p \<in> carrier (poly_ring R)" and "q \<in> carrier (poly_ring R)" and "p \<sim>\<^bsub>poly_ring R\<^esub> q"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1558  | 
shows "roots p = roots q"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1559  | 
using assms pdivides_imp_roots_incl zero_pdivides  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1560  | 
unfolding pdivides_def associated_def  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1561  | 
by (metis subset_mset.eq_iff)  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1562  | 
|
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1563  | 
lemma (in domain) monic_degree_one_roots:  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1564  | 
  assumes "a \<in> carrier R" shows "roots [ \<one> , \<ominus> a ] = {# a #}"
 | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1565  | 
proof -  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1566  | 
let ?p = "[ \<one> , \<ominus> a ]"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1567  | 
|
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1568  | 
interpret UP: domain "poly_ring R"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1569  | 
using univ_poly_is_domain[OF carrier_is_subring] .  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1570  | 
|
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1571  | 
from \<open>a \<in> carrier R\<close> have in_carrier: "?p \<in> carrier (poly_ring R)"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1572  | 
unfolding sym[OF univ_poly_carrier] polynomial_def by simp  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1573  | 
show ?thesis  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1574  | 
proof (rule subset_mset.antisym)  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1575  | 
    show "{# a #} \<subseteq># roots ?p"
 | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1576  | 
using roots_mem_iff_is_root[OF in_carrier]  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1577  | 
monic_degree_one_root_condition[OF assms]  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1578  | 
by simp  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1579  | 
next  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1580  | 
    show "roots ?p \<subseteq># {# a #}"
 | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1581  | 
proof (rule mset_subset_eqI, auto)  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1582  | 
fix b assume "a \<noteq> b" thus "count (roots ?p) b = 0"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1583  | 
using alg_mult_gt_zero_iff_is_root[OF in_carrier]  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1584  | 
monic_degree_one_root_condition[OF assms]  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1585  | 
unfolding sym[OF alg_mult_eq_count_roots[OF in_carrier]]  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1586  | 
by fastforce  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1587  | 
next  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1588  | 
have "(?p [^]\<^bsub>poly_ring R\<^esub> (alg_mult ?p a)) pdivides ?p"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1589  | 
using le_alg_mult_imp_pdivides[OF assms in_carrier] by simp  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1590  | 
hence "degree (?p [^]\<^bsub>poly_ring R\<^esub> (alg_mult ?p a)) \<le> degree ?p"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1591  | 
using pdivides_imp_degree_le[OF carrier_is_subring, of _ ?p] in_carrier by auto  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1592  | 
thus "count (roots ?p) a \<le> Suc 0"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1593  | 
using polynomial_pow_degree[OF in_carrier]  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1594  | 
unfolding sym[OF alg_mult_eq_count_roots[OF in_carrier]]  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1595  | 
by auto  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1596  | 
qed  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1597  | 
qed  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1598  | 
qed  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1599  | 
|
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1600  | 
lemma (in domain) degree_one_roots:  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1601  | 
assumes "a \<in> carrier R" "a' \<in> carrier R" and "b \<in> carrier R" and "a \<otimes> a' = \<one>"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1602  | 
  shows "roots [ a , b ] = {# \<ominus> (a' \<otimes> b) #}"
 | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1603  | 
proof -  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1604  | 
have "[ a, b ] \<in> carrier (poly_ring R)" and "[ \<one>, a' \<otimes> b ] \<in> carrier (poly_ring R)"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1605  | 
using assms unfolding sym[OF univ_poly_carrier] polynomial_def by auto  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1606  | 
thus ?thesis  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1607  | 
using subring_degree_one_associatedI[OF carrier_is_subring assms] assms  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1608  | 
monic_degree_one_roots associated_polynomials_imp_same_roots  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1609  | 
by (metis add.inv_closed local.minus_minus m_closed)  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1610  | 
qed  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1611  | 
|
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1612  | 
lemma (in field) degree_one_imp_singleton_roots:  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1613  | 
assumes "p \<in> carrier (poly_ring R)" and "degree p = 1"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1614  | 
  shows "roots p = {# \<ominus> (inv (lead_coeff p) \<otimes> (const_term p)) #}"
 | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1615  | 
proof -  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1616  | 
from \<open>p \<in> carrier (poly_ring R)\<close> and \<open>degree p = 1\<close>  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1617  | 
obtain a b where "p = [ a, b ]" and "a \<in> carrier R" "a \<noteq> \<zero>" and "b \<in> carrier R"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1618  | 
by auto  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1619  | 
thus ?thesis  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1620  | 
using degree_one_roots[of a "inv a" b]  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1621  | 
by (auto simp add: const_term_def field_Units)  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1622  | 
qed  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1623  | 
|
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1624  | 
lemma (in field) degree_one_imp_splitted:  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1625  | 
assumes "p \<in> carrier (poly_ring R)" and "degree p = 1" shows "splitted p"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1626  | 
using degree_one_imp_singleton_roots[OF assms] assms(2) unfolding splitted_def by simp  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1627  | 
|
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1628  | 
lemma (in field) no_roots_imp_same_roots:  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1629  | 
assumes "p \<in> carrier (poly_ring R)" "p \<noteq> []" and "q \<in> carrier (poly_ring R)"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1630  | 
  shows "roots p = {#} \<Longrightarrow> roots (p \<otimes>\<^bsub>poly_ring R\<^esub> q) = roots q"
 | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1631  | 
proof -  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1632  | 
interpret UP: domain "poly_ring R"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1633  | 
using univ_poly_is_domain[OF carrier_is_subring] .  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1634  | 
|
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1635  | 
  assume no_roots: "roots p = {#}" show "roots (p \<otimes>\<^bsub>poly_ring R\<^esub> q) = roots q"
 | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1636  | 
proof (intro subset_mset.antisym)  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1637  | 
have pdiv: "q pdivides (p \<otimes>\<^bsub>poly_ring R\<^esub> q)"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1638  | 
using UP.divides_prod_l assms unfolding pdivides_def by blast  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1639  | 
show "roots q \<subseteq># roots (p \<otimes>\<^bsub>poly_ring R\<^esub> q)"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1640  | 
using pdivides_imp_roots_incl[OF _ _ _ pdiv] assms  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1641  | 
degree_zero_imp_empty_roots[OF assms(3)]  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1642  | 
by (cases "q = []", auto, metis UP.l_null UP.m_rcancel UP.zero_closed univ_poly_zero)  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1643  | 
next  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1644  | 
show "roots (p \<otimes>\<^bsub>poly_ring R\<^esub> q) \<subseteq># roots q"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1645  | 
proof (cases "p \<otimes>\<^bsub>poly_ring R\<^esub> q = []")  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1646  | 
case True thus ?thesis  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1647  | 
using degree_zero_imp_empty_roots[OF UP.m_closed[OF assms(1,3)]] by simp  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1648  | 
next  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1649  | 
case False with \<open>p \<noteq> []\<close> have q_not_zero: "q \<noteq> []"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1650  | 
by (metis UP.r_null assms(1) univ_poly_zero)  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1651  | 
show ?thesis  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1652  | 
proof (rule roots_inclI[OF UP.m_closed[OF assms(1,3)] assms(3) q_not_zero])  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1653  | 
fix a assume a: "a \<in> carrier R"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1654  | 
hence "\<not> ([ \<one>, \<ominus> a ] pdivides p)"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1655  | 
using assms(1-2) no_roots pdivides_imp_is_root roots_mem_iff_is_root[of p] by auto  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1656  | 
moreover have in_carrier: "[ \<one>, \<ominus> a ] \<in> carrier (poly_ring R)"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1657  | 
using a unfolding sym[OF univ_poly_carrier] polynomial_def by auto  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1658  | 
hence "pirreducible (carrier R) [ \<one>, \<ominus> a ]"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1659  | 
using degree_one_imp_pirreducible[OF carrier_is_subfield] by simp  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1660  | 
moreover  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1661  | 
have "([ \<one>, \<ominus> a ] [^]\<^bsub>poly_ring R\<^esub> (alg_mult (p \<otimes>\<^bsub>poly_ring R\<^esub> q) a)) pdivides (p \<otimes>\<^bsub>poly_ring R\<^esub> q)"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1662  | 
using le_alg_mult_imp_pdivides[OF a UP.m_closed, of p q] assms by simp  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1663  | 
ultimately show "([ \<one>, \<ominus> a ] [^]\<^bsub>poly_ring R\<^esub> (alg_mult (p \<otimes>\<^bsub>poly_ring R\<^esub> q) a)) pdivides q"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1664  | 
using pirreducible_pow_pdivides_iff[OF carrier_is_subfield in_carrier] assms by auto  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1665  | 
qed  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1666  | 
qed  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1667  | 
qed  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1668  | 
qed  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1669  | 
|
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1670  | 
lemma (in field) poly_mult_degree_one_monic_imp_same_roots:  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1671  | 
assumes "a \<in> carrier R" and "p \<in> carrier (poly_ring R)" "p \<noteq> []"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1672  | 
shows "roots ([ \<one>, \<ominus> a ] \<otimes>\<^bsub>poly_ring R\<^esub> p) = add_mset a (roots p)"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1673  | 
proof -  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1674  | 
interpret UP: domain "poly_ring R"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1675  | 
using univ_poly_is_domain[OF carrier_is_subring] .  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1676  | 
|
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1677  | 
from \<open>a \<in> carrier R\<close> have in_carrier: "[ \<one>, \<ominus> a ] \<in> carrier (poly_ring R)"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1678  | 
unfolding sym[OF univ_poly_carrier] polynomial_def by simp  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1679  | 
|
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1680  | 
show ?thesis  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1681  | 
proof (intro subset_mset.antisym[OF roots_inclI' mset_subset_eqI])  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1682  | 
show "([ \<one>, \<ominus> a ] \<otimes>\<^bsub>poly_ring R\<^esub> p) \<in> carrier (poly_ring R)"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1683  | 
using in_carrier assms(2) by simp  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1684  | 
next  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1685  | 
fix b assume b: "b \<in> carrier R" and "[ \<one>, \<ominus> a ] \<otimes>\<^bsub>poly_ring R\<^esub> p \<noteq> []"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1686  | 
hence not_zero: "p \<noteq> []"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1687  | 
unfolding univ_poly_def by auto  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1688  | 
from \<open>b \<in> carrier R\<close> have in_carrier': "[ \<one>, \<ominus> b ] \<in> carrier (poly_ring R)"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1689  | 
unfolding sym[OF univ_poly_carrier] polynomial_def by simp  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1690  | 
show "alg_mult ([ \<one>, \<ominus> a ] \<otimes>\<^bsub>poly_ring R\<^esub> p) b \<le> count (add_mset a (roots p)) b"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1691  | 
proof (cases "a = b")  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1692  | 
case False  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1693  | 
hence "\<not> [ \<one>, \<ominus> b ] pdivides [ \<one>, \<ominus> a ]"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1694  | 
using assms(1) b monic_degree_one_root_condition pdivides_imp_is_root by blast  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1695  | 
moreover have "pirreducible (carrier R) [ \<one>, \<ominus> b ]"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1696  | 
using degree_one_imp_pirreducible[OF carrier_is_subfield in_carrier'] by simp  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1697  | 
ultimately  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1698  | 
have "[ \<one>, \<ominus> b ] [^]\<^bsub>poly_ring R\<^esub> (alg_mult ([ \<one>, \<ominus> a ] \<otimes>\<^bsub>poly_ring R\<^esub> p) b) pdivides p"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1699  | 
using le_alg_mult_imp_pdivides[OF b UP.m_closed, of _ p] assms(2) in_carrier  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1700  | 
pirreducible_pow_pdivides_iff[OF carrier_is_subfield in_carrier' in_carrier, of p]  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1701  | 
by auto  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1702  | 
with \<open>a \<noteq> b\<close> show ?thesis  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1703  | 
using alg_mult_eq_count_roots[OF assms(2)] alg_multE(2)[OF b assms(2) not_zero] by auto  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1704  | 
next  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1705  | 
case True  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1706  | 
have "[ \<one>, \<ominus> a ] pdivides ([ \<one>, \<ominus> a ] \<otimes>\<^bsub>poly_ring R\<^esub> p)"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1707  | 
using dividesI[OF assms(2)] unfolding pdivides_def by auto  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1708  | 
with \<open>[ \<one>, \<ominus> a ] \<otimes>\<^bsub>poly_ring R\<^esub> p \<noteq> []\<close>  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1709  | 
have "alg_mult ([ \<one>, \<ominus> a ] \<otimes>\<^bsub>poly_ring R\<^esub> p) a \<ge> Suc 0"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1710  | 
using alg_multE(2)[of a _ "Suc 0"] in_carrier assms by auto  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1711  | 
then obtain m where m: "alg_mult ([ \<one>, \<ominus> a ] \<otimes>\<^bsub>poly_ring R\<^esub> p) a = Suc m"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1712  | 
using Suc_le_D by blast  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1713  | 
hence "([ \<one>, \<ominus> a ] \<otimes>\<^bsub>poly_ring R\<^esub> ([ \<one>, \<ominus> a ] [^]\<^bsub>poly_ring R\<^esub> m)) pdivides  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1714  | 
([ \<one>, \<ominus> a ] \<otimes>\<^bsub>poly_ring R\<^esub> p)"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1715  | 
using le_alg_mult_imp_pdivides[OF _ UP.m_closed, of a _ p]  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1716  | 
in_carrier assms UP.nat_pow_Suc2 by force  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1717  | 
hence "([ \<one>, \<ominus> a ] [^]\<^bsub>poly_ring R\<^esub> m) pdivides p"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1718  | 
using UP.mult_divides in_carrier assms(2)  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1719  | 
unfolding univ_poly_zero pdivides_def factor_def  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1720  | 
by (simp add: UP.m_assoc UP.m_lcancel univ_poly_zero)  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1721  | 
with \<open>a = b\<close> show ?thesis  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1722  | 
using alg_mult_eq_count_roots assms in_carrier UP.nat_pow_Suc2  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1723  | 
alg_multE(2)[OF assms(1) _ not_zero] m  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1724  | 
by auto  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1725  | 
qed  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1726  | 
next  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1727  | 
fix b  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1728  | 
have not_zero: "[ \<one>, \<ominus> a ] \<otimes>\<^bsub>poly_ring R\<^esub> p \<noteq> []"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1729  | 
using assms in_carrier univ_poly_zero[of R] UP.integral by auto  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1730  | 
|
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1731  | 
show "count (add_mset a (roots p)) b \<le> count (roots ([\<one>, \<ominus> a] \<otimes>\<^bsub>poly_ring R\<^esub> p)) b"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1732  | 
proof (cases "a = b")  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1733  | 
case True  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1734  | 
have "([ \<one>, \<ominus> a ] \<otimes>\<^bsub>poly_ring R\<^esub> ([ \<one>, \<ominus> a ] [^]\<^bsub>poly_ring R\<^esub> (alg_mult p a))) pdivides  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1735  | 
([ \<one>, \<ominus> a ] \<otimes>\<^bsub>poly_ring R\<^esub> p)"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1736  | 
using UP.divides_mult[OF _ in_carrier] le_alg_mult_imp_pdivides[OF assms(1,2)] in_carrier assms  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1737  | 
by (auto simp add: pdivides_def)  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1738  | 
with \<open>a = b\<close> show ?thesis  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1739  | 
using alg_mult_eq_count_roots assms in_carrier UP.nat_pow_Suc2  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1740  | 
alg_multE(2)[OF assms(1) _ not_zero]  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1741  | 
by auto  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1742  | 
next  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1743  | 
case False  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1744  | 
have "p pdivides ([ \<one>, \<ominus> a ] \<otimes>\<^bsub>poly_ring R\<^esub> p)"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1745  | 
using dividesI[OF in_carrier] UP.m_comm in_carrier assms unfolding pdivides_def by auto  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1746  | 
thus ?thesis  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1747  | 
using False pdivides_imp_roots_incl assms in_carrier not_zero  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1748  | 
by (simp add: subseteq_mset_def)  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1749  | 
qed  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1750  | 
qed  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1751  | 
qed  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1752  | 
|
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1753  | 
lemma (in domain) not_empty_rootsE[elim]:  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1754  | 
  assumes "p \<in> carrier (poly_ring R)" and "roots p \<noteq> {#}"
 | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1755  | 
and "\<And>a. \<lbrakk> a \<in> carrier R; a \<in># roots p;  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1756  | 
[ \<one>, \<ominus> a ] \<in> carrier (poly_ring R); [ \<one>, \<ominus> a ] pdivides p \<rbrakk> \<Longrightarrow> P"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1757  | 
shows P  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1758  | 
proof -  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1759  | 
  from \<open>roots p \<noteq> {#}\<close> obtain a where "a \<in># roots p"
 | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1760  | 
by blast  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1761  | 
with \<open>p \<in> carrier (poly_ring R)\<close> have "[ \<one>, \<ominus> a ] pdivides p"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1762  | 
and "[ \<one>, \<ominus> a ] \<in> carrier (poly_ring R)" and "a \<in> carrier R"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1763  | 
using is_root_imp_pdivides[of p] roots_mem_iff_is_root[of p] is_root_def[of p a]  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1764  | 
unfolding sym[OF univ_poly_carrier] polynomial_def by auto  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1765  | 
with \<open>a \<in># roots p\<close> show ?thesis  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1766  | 
using assms(3)[of a] by auto  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1767  | 
qed  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1768  | 
|
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1769  | 
lemma (in field) associated_polynomials_imp_same_roots:  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1770  | 
assumes "p \<in> carrier (poly_ring R)" "p \<noteq> []" and "q \<in> carrier (poly_ring R)" "q \<noteq> []"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1771  | 
shows "roots (p \<otimes>\<^bsub>poly_ring R\<^esub> q) = roots p + roots q"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1772  | 
proof -  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1773  | 
interpret UP: domain "poly_ring R"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1774  | 
using univ_poly_is_domain[OF carrier_is_subring] .  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1775  | 
from assms show ?thesis  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1776  | 
proof (induction "degree p" arbitrary: p rule: less_induct)  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1777  | 
case less show ?case  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1778  | 
    proof (cases "roots p = {#}")
 | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1779  | 
case True thus ?thesis  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1780  | 
using no_roots_imp_same_roots[of p q] less by simp  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1781  | 
next  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1782  | 
case False with \<open>p \<in> carrier (poly_ring R)\<close>  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1783  | 
obtain a where a: "a \<in> carrier R" and "a \<in># roots p" and pdiv: "[ \<one>, \<ominus> a ] pdivides p"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1784  | 
and in_carrier: "[ \<one>, \<ominus> a ] \<in> carrier (poly_ring R)"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1785  | 
by blast  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1786  | 
show ?thesis  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1787  | 
proof (cases "degree p = 1")  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1788  | 
case True with \<open>p \<in> carrier (poly_ring R)\<close>  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1789  | 
obtain b c where p: "p = [ b, c ]" and b: "b \<in> carrier R" "b \<noteq> \<zero>" and c: "c \<in> carrier R"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1790  | 
by auto  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1791  | 
        with \<open>a \<in># roots p\<close> have roots: "roots p = {# a #}" and a: "\<ominus> a = inv b \<otimes> c" "a \<in> carrier R"
 | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1792  | 
and lead: "lead_coeff p = b" and const: "const_term p = c"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1793  | 
using degree_one_imp_singleton_roots[of p] less(2) field_Units  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1794  | 
unfolding const_term_def by auto  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1795  | 
hence "(p \<otimes>\<^bsub>poly_ring R\<^esub> q) \<sim>\<^bsub>poly_ring R\<^esub> ([ \<one>, \<ominus> a ] \<otimes>\<^bsub>poly_ring R\<^esub> q)"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1796  | 
using UP.mult_cong_l[OF degree_one_associatedI[OF carrier_is_subfield _ True]] less(2,4)  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1797  | 
by (auto simp add: a lead const)  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1798  | 
hence "roots (p \<otimes>\<^bsub>poly_ring R\<^esub> q) = roots ([ \<one>, \<ominus> a ] \<otimes>\<^bsub>poly_ring R\<^esub> q)"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1799  | 
using associated_polynomials_imp_same_roots in_carrier less(2,4) unfolding a by simp  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1800  | 
thus ?thesis  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1801  | 
unfolding poly_mult_degree_one_monic_imp_same_roots[OF a(2) less(4,5)] roots by simp  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1802  | 
next  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1803  | 
case False  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1804  | 
from \<open>[ \<one>, \<ominus> a ] pdivides p\<close>  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1805  | 
obtain r where p: "p = [ \<one>, \<ominus> a ] \<otimes>\<^bsub>poly_ring R\<^esub> r" and r: "r \<in> carrier (poly_ring R)"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1806  | 
unfolding pdivides_def by auto  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1807  | 
with \<open>p \<noteq> []\<close> have not_zero: "r \<noteq> []"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1808  | 
using in_carrier univ_poly_zero[of R "carrier R"] UP.integral_iff by auto  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1809  | 
with \<open>p = [ \<one>, \<ominus> a ] \<otimes>\<^bsub>poly_ring R\<^esub> r\<close> have deg: "degree p = Suc (degree r)"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1810  | 
using poly_mult_degree_eq[OF carrier_is_subring, of _ r] in_carrier r  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1811  | 
unfolding univ_poly_carrier sym[OF univ_poly_mult[of R "carrier R"]] by auto  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1812  | 
with \<open>r \<noteq> []\<close> and \<open>q \<noteq> []\<close> have "r \<otimes>\<^bsub>poly_ring R\<^esub> q \<noteq> []"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1813  | 
using in_carrier univ_poly_zero[of R "carrier R"] UP.integral less(4) r by auto  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1814  | 
hence "roots (p \<otimes>\<^bsub>poly_ring R\<^esub> q) = add_mset a (roots (r \<otimes>\<^bsub>poly_ring R\<^esub> q))"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1815  | 
using poly_mult_degree_one_monic_imp_same_roots[OF a UP.m_closed[OF r less(4)]]  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1816  | 
UP.m_assoc[OF in_carrier r less(4)] p by auto  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1817  | 
also have " ... = add_mset a (roots r + roots q)"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1818  | 
using less(1)[OF _ r not_zero less(4-5)] deg by simp  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1819  | 
also have " ... = (add_mset a (roots r)) + roots q"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1820  | 
by simp  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1821  | 
also have " ... = roots p + roots q"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1822  | 
using poly_mult_degree_one_monic_imp_same_roots[OF a r not_zero] p by simp  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1823  | 
finally show ?thesis .  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1824  | 
qed  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1825  | 
qed  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1826  | 
qed  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1827  | 
qed  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1828  | 
|
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1829  | 
lemma (in field) size_roots_le_degree:  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1830  | 
assumes "p \<in> carrier (poly_ring R)" shows "size (roots p) \<le> degree p"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1831  | 
using assms  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1832  | 
proof (induction "degree p" arbitrary: p rule: less_induct)  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1833  | 
case less show ?case  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1834  | 
  proof (cases "roots p = {#}", simp)
 | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1835  | 
interpret UP: domain "poly_ring R"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1836  | 
using univ_poly_is_domain[OF carrier_is_subring] .  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1837  | 
|
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1838  | 
case False with \<open>p \<in> carrier (poly_ring R)\<close>  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1839  | 
obtain a where a: "a \<in> carrier R" and "a \<in># roots p" and "[ \<one>, \<ominus> a ] pdivides p"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1840  | 
and in_carrier: "[ \<one>, \<ominus> a ] \<in> carrier (poly_ring R)"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1841  | 
by blast  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1842  | 
then obtain q where p: "p = [ \<one>, \<ominus> a ] \<otimes>\<^bsub>poly_ring R\<^esub> q" and q: "q \<in> carrier (poly_ring R)"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1843  | 
unfolding pdivides_def by auto  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1844  | 
with \<open>a \<in># roots p\<close> have "p \<noteq> []"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1845  | 
using degree_zero_imp_empty_roots[OF less(2)] by auto  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1846  | 
hence not_zero: "q \<noteq> []"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1847  | 
using in_carrier univ_poly_zero[of R "carrier R"] UP.integral_iff p by auto  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1848  | 
hence "degree p = Suc (degree q)"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1849  | 
using poly_mult_degree_eq[OF carrier_is_subring, of _ q] in_carrier p q  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1850  | 
unfolding univ_poly_carrier sym[OF univ_poly_mult[of R "carrier R"]] by auto  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1851  | 
with \<open>q \<noteq> []\<close> show ?thesis  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1852  | 
using poly_mult_degree_one_monic_imp_same_roots[OF a q] p less(1)[OF _ q]  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1853  | 
by (metis Suc_le_mono lessI size_add_mset)  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1854  | 
qed  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1855  | 
qed  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1856  | 
|
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1857  | 
lemma (in domain) pirreducible_roots:  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1858  | 
assumes "p \<in> carrier (poly_ring R)" and "pirreducible (carrier R) p" and "degree p \<noteq> 1"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1859  | 
  shows "roots p = {#}"
 | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1860  | 
proof (rule ccontr)  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1861  | 
  assume "roots p \<noteq> {#}" with \<open>p \<in> carrier (poly_ring R)\<close>
 | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1862  | 
obtain a where a: "a \<in> carrier R" and "a \<in># roots p" and "[ \<one>, \<ominus> a ] pdivides p"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1863  | 
and in_carrier: "[ \<one>, \<ominus> a ] \<in> carrier (poly_ring R)"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1864  | 
by blast  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1865  | 
hence "[ \<one>, \<ominus> a ] \<sim>\<^bsub>poly_ring R\<^esub> p"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1866  | 
using divides_pirreducible_condition[OF assms(2) in_carrier]  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1867  | 
univ_poly_units_incl[OF carrier_is_subring]  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1868  | 
unfolding pdivides_def by auto  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1869  | 
hence "degree p = 1"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1870  | 
using associated_polynomials_imp_same_length[OF carrier_is_subring in_carrier assms(1)] by auto  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1871  | 
with \<open>degree p \<noteq> 1\<close> show False ..  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1872  | 
qed  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1873  | 
|
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1874  | 
lemma (in field) pirreducible_imp_not_splitted:  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1875  | 
assumes "p \<in> carrier (poly_ring R)" and "pirreducible (carrier R) p" and "degree p \<noteq> 1"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1876  | 
shows "\<not> splitted p"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1877  | 
using pirreducible_roots[of p] pirreducible_degree[OF carrier_is_subfield, of p] assms  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1878  | 
by (simp add: splitted_def)  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1879  | 
|
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1880  | 
lemma (in field)  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1881  | 
assumes "p \<in> carrier (poly_ring R)" and "q \<in> carrier (poly_ring R)"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1882  | 
and "pirreducible (carrier R) p" and "degree p \<noteq> 1"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1883  | 
shows "roots (p \<otimes>\<^bsub>poly_ring R\<^esub> q) = roots q"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1884  | 
using no_roots_imp_same_roots[of p q] pirreducible_roots[of p] assms  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1885  | 
unfolding ring_irreducible_def univ_poly_zero by auto  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1886  | 
|
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1887  | 
lemma (in field) trivial_factors_imp_splitted:  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1888  | 
assumes "p \<in> carrier (poly_ring R)"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1889  | 
and "\<And>q. \<lbrakk> q \<in> carrier (poly_ring R); pirreducible (carrier R) q; q pdivides p \<rbrakk> \<Longrightarrow> degree q \<le> 1"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1890  | 
shows "splitted p"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1891  | 
using assms  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1892  | 
proof (induction "degree p" arbitrary: p rule: less_induct)  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1893  | 
interpret UP: principal_domain "poly_ring R"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1894  | 
using univ_poly_is_principal[OF carrier_is_subfield] .  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1895  | 
case less show ?case  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1896  | 
proof (cases "degree p = 0", simp add: degree_zero_imp_splitted[OF less(2)])  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1897  | 
case False show ?thesis  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1898  | 
    proof (cases "roots p = {#}")
 | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1899  | 
case True  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1900  | 
      from \<open>degree p \<noteq> 0\<close> have "p \<notin> Units (poly_ring R)" and "p \<in> carrier (poly_ring R) - { [] }"
 | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1901  | 
using univ_poly_units'[OF carrier_is_subfield, of p] less(2) by auto  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1902  | 
then obtain q where "q \<in> carrier (poly_ring R)" "pirreducible (carrier R) q" and "q pdivides p"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1903  | 
using UP.exists_irreducible_divisor[of p] unfolding univ_poly_zero pdivides_def by auto  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1904  | 
      with \<open>degree p \<noteq> 0\<close> have "roots p \<noteq> {#}"
 | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1905  | 
using degree_one_imp_singleton_roots[OF _ , of q] less(3)[of q]  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1906  | 
pdivides_imp_roots_incl[OF _ less(2), of q]  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1907  | 
pirreducible_degree[OF carrier_is_subfield, of q]  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1908  | 
by force  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1909  | 
      from \<open>roots p = {#}\<close> and \<open>roots p \<noteq> {#}\<close> have False
 | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1910  | 
by simp  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1911  | 
thus ?thesis ..  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1912  | 
next  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1913  | 
case False with \<open>p \<in> carrier (poly_ring R)\<close>  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1914  | 
obtain a where a: "a \<in> carrier R" and "a \<in># roots p" and "[ \<one>, \<ominus> a ] pdivides p"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1915  | 
and in_carrier: "[ \<one>, \<ominus> a ] \<in> carrier (poly_ring R)"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1916  | 
by blast  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1917  | 
then obtain q where p: "p = [ \<one>, \<ominus> a ] \<otimes>\<^bsub>poly_ring R\<^esub> q" and q: "q \<in> carrier (poly_ring R)"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1918  | 
unfolding pdivides_def by blast  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1919  | 
with \<open>degree p \<noteq> 0\<close> have "p \<noteq> []"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1920  | 
by auto  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1921  | 
with \<open>p = [ \<one>, \<ominus> a ] \<otimes>\<^bsub>poly_ring R\<^esub> q\<close> have "q \<noteq> []"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1922  | 
using in_carrier q unfolding sym[OF univ_poly_zero[of R "carrier R"]] by auto  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1923  | 
with \<open>p = [ \<one>, \<ominus> a ] \<otimes>\<^bsub>poly_ring R\<^esub> q\<close> and \<open>p \<noteq> []\<close> have "degree p = Suc (degree q)"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1924  | 
using poly_mult_degree_eq[OF carrier_is_subring] in_carrier q  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1925  | 
unfolding univ_poly_carrier sym[OF univ_poly_mult[of R "carrier R"]] by auto  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1926  | 
moreover have "q pdivides p"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1927  | 
using p dividesI[OF in_carrier] UP.m_comm[OF in_carrier q] by (auto simp add: pdivides_def)  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1928  | 
hence "degree r = 1" if "r \<in> carrier (poly_ring R)" and "pirreducible (carrier R) r"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1929  | 
and "r pdivides q" for r  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1930  | 
using less(3)[OF that(1-2)] UP.divides_trans[OF _ _ that(1), of q p] that(3)  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1931  | 
pirreducible_degree[OF carrier_is_subfield that(1-2)]  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1932  | 
by (auto simp add: pdivides_def)  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1933  | 
ultimately have "splitted q"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1934  | 
using less(1)[OF _ q] by auto  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1935  | 
with \<open>degree p = Suc (degree q)\<close> and \<open>q \<noteq> []\<close> show ?thesis  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1936  | 
using poly_mult_degree_one_monic_imp_same_roots[OF a q]  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1937  | 
unfolding sym[OF p] splitted_def  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1938  | 
by simp  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1939  | 
qed  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1940  | 
qed  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1941  | 
qed  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1942  | 
|
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1943  | 
lemma (in field) pdivides_imp_splitted:  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1944  | 
assumes "p \<in> carrier (poly_ring R)" and "q \<in> carrier (poly_ring R)" "q \<noteq> []" and "splitted q"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1945  | 
shows "p pdivides q \<Longrightarrow> splitted p"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1946  | 
proof (cases "p = []")  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1947  | 
case True thus ?thesis  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1948  | 
using degree_zero_imp_splitted[OF assms(1)] by simp  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1949  | 
next  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1950  | 
interpret UP: principal_domain "poly_ring R"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1951  | 
using univ_poly_is_principal[OF carrier_is_subfield] .  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1952  | 
|
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1953  | 
case False  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1954  | 
assume "p pdivides q"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1955  | 
then obtain b where b: "b \<in> carrier (poly_ring R)" and q: "q = p \<otimes>\<^bsub>poly_ring R\<^esub> b"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1956  | 
unfolding pdivides_def by auto  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1957  | 
with \<open>q \<noteq> []\<close> have "p \<noteq> []" and "b \<noteq> []"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1958  | 
using assms UP.integral_iff[of p b] unfolding sym[OF univ_poly_zero[of R "carrier R"]] by auto  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1959  | 
hence "degree p + degree b = size (roots p) + size (roots b)"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1960  | 
using associated_polynomials_imp_same_roots[of p b] assms b q splitted_def  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1961  | 
poly_mult_degree_eq[OF carrier_is_subring,of p b]  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1962  | 
unfolding univ_poly_carrier sym[OF univ_poly_mult[of R "carrier R"]]  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1963  | 
by auto  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1964  | 
moreover have "size (roots p) \<le> degree p" and "size (roots b) \<le> degree b"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1965  | 
using size_roots_le_degree assms(1) b by auto  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1966  | 
ultimately show ?thesis  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1967  | 
unfolding splitted_def by linarith  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1968  | 
qed  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1969  | 
|
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1970  | 
lemma (in field) splitted_imp_trivial_factors:  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1971  | 
assumes "p \<in> carrier (poly_ring R)" "p \<noteq> []" and "splitted p"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1972  | 
shows "\<And>q. \<lbrakk> q \<in> carrier (poly_ring R); pirreducible (carrier R) q; q pdivides p \<rbrakk> \<Longrightarrow> degree q = 1"  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1973  | 
using pdivides_imp_splitted[OF _ assms] pirreducible_imp_not_splitted  | 
| 
 
8371a25ca177
Algebraic closure: moving more theorems into their rightful places
 
paulson <lp15@cam.ac.uk> 
parents: 
70214 
diff
changeset
 | 
1974  | 
by auto  | 
| 
70160
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1975  | 
|
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1976  | 
|
| 70162 | 1977  | 
subsection \<open>Link between @{term \<open>(pmod)\<close>} and @{term rupture_surj}\<close>
 | 
| 
70160
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1978  | 
|
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1979  | 
lemma (in domain) rupture_surj_composed_with_pmod:  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1980  | 
assumes "subfield K R" and "p \<in> carrier (K[X])" and "q \<in> carrier (K[X])"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1981  | 
shows "rupture_surj K p q = rupture_surj K p (q pmod p)"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1982  | 
proof -  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1983  | 
interpret UP: principal_domain "K[X]"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1984  | 
using univ_poly_is_principal[OF assms(1)] .  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1985  | 
interpret Rupt: ring "Rupt K p"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1986  | 
using assms by (simp add: UP.cgenideal_ideal ideal.quotient_is_ring rupture_def)  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1987  | 
|
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1988  | 
let ?h = "rupture_surj K p"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1989  | 
|
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1990  | 
have "?h q = (?h p \<otimes>\<^bsub>Rupt K p\<^esub> ?h (q pdiv p)) \<oplus>\<^bsub>Rupt K p\<^esub> ?h (q pmod p)"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1991  | 
and "?h (q pdiv p) \<in> carrier (Rupt K p)" "?h (q pmod p) \<in> carrier (Rupt K p)"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1992  | 
using pdiv_pmod[OF assms(1,3,2)] long_division_closed[OF assms(1,3,2)] assms UP.m_closed  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1993  | 
ring_hom_memE[OF rupture_surj_hom(1)[OF subfieldE(1)[OF assms(1)] assms(2)]]  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1994  | 
by metis+  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1995  | 
moreover have "?h p = PIdl\<^bsub>K[X]\<^esub> p"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1996  | 
using assms by (simp add: UP.a_rcos_zero UP.cgenideal_ideal UP.cgenideal_self)  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1997  | 
hence "?h p = \<zero>\<^bsub>Rupt K p\<^esub>"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1998  | 
unfolding rupture_def FactRing_def by simp  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
1999  | 
ultimately show ?thesis  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2000  | 
by simp  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2001  | 
qed  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2002  | 
|
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2003  | 
corollary (in domain) rupture_carrier_as_pmod_image:  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2004  | 
assumes "subfield K R" and "p \<in> carrier (K[X])"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2005  | 
shows "(rupture_surj K p) ` ((\<lambda>q. q pmod p) ` (carrier (K[X]))) = carrier (Rupt K p)"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2006  | 
(is "?lhs = ?rhs")  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2007  | 
proof  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2008  | 
have "(\<lambda>q. q pmod p) ` carrier (K[X]) \<subseteq> carrier (K[X])"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2009  | 
using long_division_closed(2)[OF assms(1) _ assms(2)] by auto  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2010  | 
thus "?lhs \<subseteq> ?rhs"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2011  | 
using ring_hom_memE(1)[OF rupture_surj_hom(1)[OF subfieldE(1)[OF assms(1)] assms(2)]] by auto  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2012  | 
next  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2013  | 
show "?rhs \<subseteq> ?lhs"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2014  | 
proof  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2015  | 
fix a assume "a \<in> carrier (Rupt K p)"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2016  | 
then obtain q where "q \<in> carrier (K[X])" and "a = rupture_surj K p q"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2017  | 
unfolding rupture_def FactRing_def A_RCOSETS_def' by auto  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2018  | 
thus "a \<in> ?lhs"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2019  | 
using rupture_surj_composed_with_pmod[OF assms] by auto  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2020  | 
qed  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2021  | 
qed  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2022  | 
|
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2023  | 
lemma (in domain) rupture_surj_inj_on:  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2024  | 
assumes "subfield K R" and "p \<in> carrier (K[X])"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2025  | 
shows "inj_on (rupture_surj K p) ((\<lambda>q. q pmod p) ` (carrier (K[X])))"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2026  | 
proof (intro inj_onI)  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2027  | 
interpret UP: principal_domain "K[X]"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2028  | 
using univ_poly_is_principal[OF assms(1)] .  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2029  | 
|
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2030  | 
fix a b  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2031  | 
assume "a \<in> (\<lambda>q. q pmod p) ` carrier (K[X])"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2032  | 
and "b \<in> (\<lambda>q. q pmod p) ` carrier (K[X])"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2033  | 
then obtain q s  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2034  | 
where q: "q \<in> carrier (K[X])" "a = q pmod p"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2035  | 
and s: "s \<in> carrier (K[X])" "b = s pmod p"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2036  | 
by auto  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2037  | 
moreover assume "rupture_surj K p a = rupture_surj K p b"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2038  | 
ultimately have "q \<ominus>\<^bsub>K[X]\<^esub> s \<in> (PIdl\<^bsub>K[X]\<^esub> p)"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2039  | 
using UP.quotient_eq_iff_same_a_r_cos[OF UP.cgenideal_ideal[OF assms(2)], of q s]  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2040  | 
rupture_surj_composed_with_pmod[OF assms] by auto  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2041  | 
hence "p pdivides (q \<ominus>\<^bsub>K[X]\<^esub> s)"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2042  | 
using assms q(1) s(1) UP.to_contain_is_to_divide pdivides_iff_shell  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2043  | 
by (meson UP.cgenideal_ideal UP.cgenideal_minimal UP.minus_closed)  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2044  | 
thus "a = b"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2045  | 
unfolding q s same_pmod_iff_pdivides[OF assms(1) q(1) s(1) assms(2)] .  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2046  | 
qed  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2047  | 
|
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2048  | 
|
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2049  | 
subsection \<open>Dimension\<close>  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2050  | 
|
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2051  | 
definition (in ring) exp_base :: "'a \<Rightarrow> nat \<Rightarrow> 'a list"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2052  | 
where "exp_base x n = map (\<lambda>i. x [^] i) (rev [0..< n])"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2053  | 
|
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2054  | 
lemma (in ring) exp_base_closed:  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2055  | 
assumes "x \<in> carrier R" shows "set (exp_base x n) \<subseteq> carrier R"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2056  | 
using assms by (induct n) (auto simp add: exp_base_def)  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2057  | 
|
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2058  | 
lemma (in ring) exp_base_append:  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2059  | 
shows "exp_base x (n + m) = (map (\<lambda>i. x [^] i) (rev [n..< n + m])) @ exp_base x n"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2060  | 
unfolding exp_base_def by (metis map_append rev_append upt_add_eq_append zero_le)  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2061  | 
|
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2062  | 
lemma (in ring) drop_exp_base:  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2063  | 
shows "drop n (exp_base x m) = exp_base x (m - n)"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2064  | 
proof -  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2065  | 
have ?thesis if "n > m"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2066  | 
using that by (simp add: exp_base_def)  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2067  | 
moreover have ?thesis if "n \<le> m"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2068  | 
using exp_base_append[of x "m - n" n] that by auto  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2069  | 
ultimately show ?thesis  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2070  | 
by linarith  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2071  | 
qed  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2072  | 
|
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2073  | 
lemma (in ring) combine_eq_eval:  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2074  | 
shows "combine Ks (exp_base x (length Ks)) = eval Ks x"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2075  | 
unfolding exp_base_def by (induct Ks) (auto)  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2076  | 
|
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2077  | 
lemma (in domain) pmod_image_characterization:  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2078  | 
assumes "subfield K R" and "p \<in> carrier (K[X])" and "p \<noteq> []"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2079  | 
  shows "(\<lambda>q. q pmod p) ` carrier (K[X]) = { q \<in> carrier (K[X]). length q \<le> degree p }"
 | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2080  | 
proof -  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2081  | 
interpret UP: principal_domain "K[X]"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2082  | 
using univ_poly_is_principal[OF assms(1)] .  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2083  | 
|
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2084  | 
show ?thesis  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2085  | 
proof (intro no_atp(10)[OF subsetI subsetI])  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2086  | 
    fix q assume "q \<in> { q \<in> carrier (K[X]). length q \<le> degree p }"
 | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2087  | 
then have "q \<in> carrier (K[X])" and "length q \<le> degree p"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2088  | 
by simp+  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2089  | 
|
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2090  | 
show "q \<in> (\<lambda>q. q pmod p) ` carrier (K[X])"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2091  | 
proof (cases "q = []")  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2092  | 
case True  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2093  | 
have "p pmod p = q"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2094  | 
unfolding True pmod_zero_iff_pdivides[OF assms(1,2,2)]  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2095  | 
using assms(1-2) pdivides_iff_shell by auto  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2096  | 
thus ?thesis  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2097  | 
using assms(2) by blast  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2098  | 
next  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2099  | 
case False  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2100  | 
with \<open>length q \<le> degree p\<close> have "degree q < degree p"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2101  | 
using le_eq_less_or_eq by fastforce  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2102  | 
with \<open>q \<in> carrier (K[X])\<close> show ?thesis  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2103  | 
using pmod_const(2)[OF assms(1) _ assms(2), of q] by (metis imageI)  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2104  | 
qed  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2105  | 
next  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2106  | 
fix q assume "q \<in> (\<lambda>q. q pmod p) ` carrier (K[X])"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2107  | 
then obtain q' where "q' \<in> carrier (K[X])" and "q = q' pmod p"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2108  | 
by auto  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2109  | 
    thus "q \<in> { q \<in> carrier (K[X]). length q \<le> degree p }"
 | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2110  | 
using long_division_closed(2)[OF assms(1) _ assms(2), of q']  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2111  | 
pmod_degree[OF assms(1) _ assms(2-3), of q']  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2112  | 
by auto  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2113  | 
qed  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2114  | 
qed  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2115  | 
|
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2116  | 
lemma (in domain) Span_var_pow_base:  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2117  | 
assumes "subfield K R"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2118  | 
shows "ring.Span (K[X]) (poly_of_const ` K) (ring.exp_base (K[X]) X n) =  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2119  | 
         { q \<in> carrier (K[X]). length q \<le> n }" (is "?lhs = ?rhs")
 | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2120  | 
proof -  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2121  | 
note subring = subfieldE(1)[OF assms]  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2122  | 
note subfield = univ_poly_subfield_of_consts[OF assms]  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2123  | 
|
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2124  | 
interpret UP: domain "K[X]"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2125  | 
using univ_poly_is_domain[OF subring] .  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2126  | 
|
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2127  | 
show ?thesis  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2128  | 
proof (intro no_atp(10)[OF subsetI subsetI])  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2129  | 
    fix q assume "q \<in> { q \<in> carrier (K[X]). length q \<le> n }"
 | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2130  | 
then have q: "q \<in> carrier (K[X])" "length q \<le> n"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2131  | 
by simp+  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2132  | 
|
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2133  | 
let ?repl = "replicate (n - length q) \<zero>\<^bsub>K[X]\<^esub>"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2134  | 
let ?map = "map poly_of_const q"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2135  | 
let ?comb = "UP.combine"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2136  | 
define Ks where "Ks = ?repl @ ?map"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2137  | 
|
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2138  | 
have "q = ?comb ?map (UP.exp_base X (length q))"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2139  | 
using q eval_rewrite[OF subring q(1)] unfolding sym[OF UP.combine_eq_eval] by auto  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2140  | 
moreover from \<open>length q \<le> n\<close>  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2141  | 
have "?comb (?repl @ Ks) (UP.exp_base X n) = ?comb Ks (UP.exp_base X (length q))"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2142  | 
if "set Ks \<subseteq> carrier (K[X])" for Ks  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2143  | 
using UP.combine_prepend_replicate[OF that UP.exp_base_closed[OF var_closed(1)[OF subring]]]  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2144  | 
unfolding UP.drop_exp_base by auto  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2145  | 
|
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2146  | 
moreover have "set ?map \<subseteq> carrier (K[X])"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2147  | 
using map_norm_in_poly_ring_carrier[OF subring q(1)]  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2148  | 
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
 | 
2149  | 
|
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2150  | 
moreover have "?repl = map poly_of_const (replicate (n - length q) \<zero>)"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2151  | 
unfolding poly_of_const_def univ_poly_zero by (induct "n - length q") (auto)  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2152  | 
hence "set ?repl \<subseteq> poly_of_const ` K"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2153  | 
using subringE(2)[OF subring] by auto  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2154  | 
moreover from \<open>q \<in> carrier (K[X])\<close> have "set q \<subseteq> K"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2155  | 
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
 | 
2156  | 
hence "set ?map \<subseteq> poly_of_const ` K"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2157  | 
by auto  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2158  | 
|
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2159  | 
ultimately have "q = ?comb Ks (UP.exp_base X n)" and "set Ks \<subseteq> poly_of_const ` K"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2160  | 
by (simp add: Ks_def)+  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2161  | 
thus "q \<in> UP.Span (poly_of_const ` K) (UP.exp_base X n)"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2162  | 
using UP.Span_eq_combine_set[OF subfield UP.exp_base_closed[OF var_closed(1)[OF subring]]] by auto  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2163  | 
next  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2164  | 
fix q assume "q \<in> UP.Span (poly_of_const ` K) (UP.exp_base X n)"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2165  | 
    thus "q \<in> { q \<in> carrier (K[X]). length q \<le> n }"
 | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2166  | 
proof (induction n arbitrary: q)  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2167  | 
case 0 thus ?case  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2168  | 
unfolding UP.exp_base_def by (auto simp add: univ_poly_zero)  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2169  | 
next  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2170  | 
case (Suc n)  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2171  | 
then obtain k p where k: "k \<in> K" and p: "p \<in> UP.Span (poly_of_const ` K) (UP.exp_base X n)"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2172  | 
and q: "q = ((poly_of_const k) \<otimes>\<^bsub>K[X]\<^esub> (X [^]\<^bsub>K[X]\<^esub> n)) \<oplus>\<^bsub>K[X]\<^esub> p"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2173  | 
unfolding UP.exp_base_def using UP.line_extension_mem_iff by auto  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2174  | 
have p_in_carrier: "p \<in> carrier (K[X])" and "length p \<le> n"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2175  | 
using Suc(1)[OF p] by simp+  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2176  | 
moreover from \<open>k \<in> K\<close> have "poly_of_const k \<in> carrier (K[X])"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2177  | 
unfolding poly_of_const_def 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
 | 
2178  | 
ultimately have "q \<in> carrier (K[X])"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2179  | 
unfolding q using var_pow_closed[OF subring, of n] by algebra  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2180  | 
|
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2181  | 
moreover have "poly_of_const k = \<zero>\<^bsub>K[X]\<^esub>" if "k = \<zero>"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2182  | 
unfolding poly_of_const_def that univ_poly_zero by simp  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2183  | 
with \<open>p \<in> carrier (K[X])\<close> have "q = p" if "k = \<zero>"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2184  | 
unfolding q using var_pow_closed[OF subring, of n] that by algebra  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2185  | 
with \<open>length p \<le> n\<close> have "length q \<le> Suc n" if "k = \<zero>"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2186  | 
using that by simp  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2187  | 
|
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2188  | 
moreover have "poly_of_const k = [ k ]" if "k \<noteq> \<zero>"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2189  | 
unfolding poly_of_const_def using that by simp  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2190  | 
hence monom: "monom k n = (poly_of_const k) \<otimes>\<^bsub>K[X]\<^esub> (X [^]\<^bsub>K[X]\<^esub> n)" if "k \<noteq> \<zero>"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2191  | 
using that monom_eq_var_pow[OF subring] subfieldE(3)[OF assms] k by auto  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2192  | 
with \<open>p \<in> carrier (K[X])\<close> and \<open>k \<in> K\<close> and \<open>length p \<le> n\<close>  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2193  | 
have "length q = Suc n" if "k \<noteq> \<zero>"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2194  | 
using that poly_add_length_eq[OF subring monom_is_polynomial[OF subring, of k n], of p]  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2195  | 
unfolding univ_poly_carrier monom_def univ_poly_add sym[OF monom[OF that]] q by auto  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2196  | 
ultimately show ?case  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2197  | 
by (cases "k = \<zero>", auto)  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2198  | 
qed  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2199  | 
qed  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2200  | 
qed  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2201  | 
|
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2202  | 
lemma (in domain) var_pow_base_independent:  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2203  | 
assumes "subfield K R"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2204  | 
shows "ring.independent (K[X]) (poly_of_const ` K) (ring.exp_base (K[X]) X n)"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2205  | 
proof -  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2206  | 
note subring = subfieldE(1)[OF assms]  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2207  | 
interpret UP: domain "K[X]"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2208  | 
using univ_poly_is_domain[OF subring] .  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2209  | 
|
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2210  | 
show ?thesis  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2211  | 
proof (induction n, simp add: UP.exp_base_def)  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2212  | 
case (Suc n)  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2213  | 
have "X [^]\<^bsub>K[X]\<^esub> n \<notin> UP.Span (poly_of_const ` K) (ring.exp_base (K[X]) X n)"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2214  | 
unfolding sym[OF unitary_monom_eq_var_pow[OF subring]] monom_def  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2215  | 
Span_var_pow_base[OF assms] by auto  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2216  | 
moreover have "X [^]\<^bsub>K[X]\<^esub> n # UP.exp_base X n = UP.exp_base X (Suc n)"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2217  | 
unfolding UP.exp_base_def by simp  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2218  | 
ultimately show ?case  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2219  | 
using UP.li_Cons[OF var_pow_closed[OF subring, of n] _Suc] by simp  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2220  | 
qed  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2221  | 
qed  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2222  | 
|
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2223  | 
lemma (in domain) bounded_degree_dimension:  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2224  | 
assumes "subfield K R"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2225  | 
  shows "ring.dimension (K[X]) n (poly_of_const ` K) { q \<in> carrier (K[X]). length q \<le> n }"
 | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2226  | 
proof -  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2227  | 
interpret UP: domain "K[X]"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2228  | 
using univ_poly_is_domain[OF subfieldE(1)[OF assms]] .  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2229  | 
have "length (UP.exp_base X n) = n"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2230  | 
unfolding UP.exp_base_def by simp  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2231  | 
thus ?thesis  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2232  | 
using UP.dimension_independent[OF var_pow_base_independent[OF assms], of n]  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2233  | 
unfolding Span_var_pow_base[OF assms] by simp  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2234  | 
qed  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2235  | 
|
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2236  | 
corollary (in domain) univ_poly_infinite_dimension:  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2237  | 
assumes "subfield K R" shows "ring.infinite_dimension (K[X]) (poly_of_const ` K) (carrier (K[X]))"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2238  | 
proof (rule ccontr)  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2239  | 
interpret UP: domain "K[X]"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2240  | 
using univ_poly_is_domain[OF subfieldE(1)[OF assms]] .  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2241  | 
|
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2242  | 
assume "\<not> UP.infinite_dimension (poly_of_const ` K) (carrier (K[X]))"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2243  | 
then obtain n where n: "UP.dimension n (poly_of_const ` K) (carrier (K[X]))"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2244  | 
by blast  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2245  | 
show False  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2246  | 
using UP.independent_length_le_dimension[OF univ_poly_subfield_of_consts[OF assms] n  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2247  | 
var_pow_base_independent[OF assms, of "Suc n"]  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2248  | 
UP.exp_base_closed[OF var_closed(1)[OF subfieldE(1)[OF assms]]]]  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2249  | 
unfolding UP.exp_base_def by simp  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2250  | 
qed  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2251  | 
|
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2252  | 
corollary (in domain) rupture_dimension:  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2253  | 
assumes "subfield K R" and "p \<in> carrier (K[X])" and "degree p > 0"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2254  | 
shows "ring.dimension (Rupt K p) (degree p) ((rupture_surj K p) ` poly_of_const ` K) (carrier (Rupt K p))"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2255  | 
proof -  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2256  | 
interpret UP: domain "K[X]"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2257  | 
using univ_poly_is_domain[OF subfieldE(1)[OF assms(1)]] .  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2258  | 
interpret Hom: ring_hom_ring "K[X]" "Rupt K p" "rupture_surj K p"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2259  | 
using rupture_surj_hom(2)[OF subfieldE(1)[OF assms(1)] assms(2)] .  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2260  | 
|
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2261  | 
have not_nil: "p \<noteq> []"  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2262  | 
using assms(3) by auto  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2263  | 
|
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2264  | 
show ?thesis  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2265  | 
using Hom.inj_hom_dimension[OF univ_poly_subfield_of_consts rupture_one_not_zero  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2266  | 
rupture_surj_inj_on] bounded_degree_dimension assms  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2267  | 
unfolding sym[OF rupture_carrier_as_pmod_image[OF assms(1-2)]]  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2268  | 
pmod_image_characterization[OF assms(1-2) not_nil]  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2269  | 
by simp  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2270  | 
qed  | 
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2271  | 
|
| 
 
8e9100dcde52
Towards a proof of algebraic closure (NB not finished)
 
paulson <lp15@cam.ac.uk> 
parents:  
diff
changeset
 | 
2272  | 
end  |