src/HOL/Algebra/Indexed_Polynomials.thy
author wenzelm
Fri, 21 Apr 2023 15:30:59 +0200
changeset 77901 5728d5ebce34
parent 70162 13b10ca71150
child 80914 d97fdabd9e2b
permissions -rw-r--r--
more uniform operations wrt. Thm.full_prop_of;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
70160
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
     1
(*  Title:      HOL/Algebra/Indexed_Polynomials.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 Indexed_Polynomials
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
     6
  imports Weak_Morphisms "HOL-Library.Multiset" Polynomial_Divisibility
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>Indexed 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
text \<open>In this theory, we build a basic framework to the study of polynomials on letters
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    13
      indexed by a set. The main interest is to then apply these concepts to the construction
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    14
      of the algebraic closure of a field. \<close>
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    15
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    16
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    17
subsection \<open>Definitions\<close>
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    18
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    19
text \<open>We formalize indexed monomials as multisets with its support a subset of the index set.
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    20
      On top of those, we build indexed polynomials which are simply functions mapping a monomial
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    21
      to its coefficient. \<close>
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 (in ring) indexed_const :: "'a \<Rightarrow> ('c multiset \<Rightarrow> 'a)" 
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    24
  where "indexed_const k = (\<lambda>m. if m = {#} then k else \<zero>)"
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 (in ring) indexed_pmult :: "('c multiset \<Rightarrow> 'a) \<Rightarrow> 'c \<Rightarrow> ('c multiset \<Rightarrow> 'a)" (infixl "\<Otimes>" 65)
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    27
  where "indexed_pmult P i = (\<lambda>m. if i \<in># m then P (m - {# i #}) else \<zero>)"
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
definition (in ring) indexed_padd :: "_ \<Rightarrow> _ \<Rightarrow> ('c multiset \<Rightarrow> 'a)" (infixl "\<Oplus>" 65)
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    30
  where "indexed_padd P Q = (\<lambda>m. (P m) \<oplus> (Q m))"
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
definition (in ring) indexed_var :: "'c \<Rightarrow> ('c multiset \<Rightarrow> 'a)" ("\<X>\<index>")
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    33
  where "indexed_var i = (indexed_const \<one>) \<Otimes> i"
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
definition (in ring) index_free :: "('c multiset \<Rightarrow> 'a) \<Rightarrow> 'c \<Rightarrow> bool"
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    36
  where "index_free P i \<longleftrightarrow> (\<forall>m. i \<in># m \<longrightarrow> P m = \<zero>)"
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    37
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    38
definition (in ring) carrier_coeff :: "('c multiset \<Rightarrow> 'a) \<Rightarrow> bool"
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    39
  where "carrier_coeff P \<longleftrightarrow> (\<forall>m. P m \<in> carrier R)"
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    40
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    41
inductive_set (in ring) indexed_pset :: "'c set \<Rightarrow> 'a set \<Rightarrow> ('c multiset \<Rightarrow> 'a) set" ("_ [\<X>\<index>]" 80)
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    42
  for I and K where
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    43
    indexed_const:  "k \<in> K \<Longrightarrow> indexed_const k \<in> (K[\<X>\<^bsub>I\<^esub>])"
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    44
  | indexed_padd:  "\<lbrakk> P \<in> (K[\<X>\<^bsub>I\<^esub>]); Q \<in> (K[\<X>\<^bsub>I\<^esub>]) \<rbrakk> \<Longrightarrow> P \<Oplus> Q \<in> (K[\<X>\<^bsub>I\<^esub>])"
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    45
  | indexed_pmult: "\<lbrakk> P \<in> (K[\<X>\<^bsub>I\<^esub>]); i \<in> I \<rbrakk> \<Longrightarrow> P \<Otimes> i \<in> (K[\<X>\<^bsub>I\<^esub>])"
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    46
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    47
fun (in ring) indexed_eval_aux :: "('c multiset \<Rightarrow> 'a) list \<Rightarrow> 'c \<Rightarrow> ('c multiset \<Rightarrow> 'a)"
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    48
  where "indexed_eval_aux Ps i = foldr (\<lambda>P Q. (Q \<Otimes> i) \<Oplus> P) Ps (indexed_const \<zero>)"
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    49
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    50
fun (in ring) indexed_eval :: "('c multiset \<Rightarrow> 'a) list \<Rightarrow> 'c \<Rightarrow> ('c multiset \<Rightarrow> 'a)"
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    51
  where "indexed_eval Ps i = indexed_eval_aux (rev Ps) i"
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
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    54
subsection \<open>Basic Properties\<close>
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    55
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    56
lemma (in ring) carrier_coeffE:
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    57
  assumes "carrier_coeff P" shows "P m \<in> carrier R"
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    58
  using assms unfolding carrier_coeff_def by simp
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    59
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    60
lemma (in ring) indexed_zero_def: "indexed_const \<zero> = (\<lambda>_. \<zero>)"
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    61
  unfolding indexed_const_def by simp
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    62
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    63
lemma (in ring) indexed_const_index_free: "index_free (indexed_const k) i"
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    64
  unfolding index_free_def indexed_const_def by auto
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) indexed_var_not_index_free: "\<not> index_free \<X>\<^bsub>i\<^esub> i"
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    67
proof -
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    68
  have "\<X>\<^bsub>i\<^esub> {# i #} = \<one>"
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    69
    unfolding indexed_var_def indexed_pmult_def indexed_const_def by simp
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    70
  thus ?thesis
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    71
    using one_not_zero unfolding index_free_def by fastforce 
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    72
qed
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 ring) indexed_pmult_zero [simp]:
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    75
  shows "indexed_pmult (indexed_const \<zero>) i = indexed_const \<zero>"
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    76
  unfolding indexed_zero_def indexed_pmult_def by auto
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    77
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    78
lemma (in ring) indexed_padd_zero:
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    79
  assumes "carrier_coeff P" shows "P \<Oplus> (indexed_const \<zero>) = P" and "(indexed_const \<zero>) \<Oplus> P = P"
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    80
  using assms unfolding carrier_coeff_def indexed_zero_def indexed_padd_def by auto
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    81
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    82
lemma (in ring) indexed_padd_const:
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    83
  shows "(indexed_const k1) \<Oplus> (indexed_const k2) = indexed_const (k1 \<oplus> k2)"
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    84
  unfolding indexed_padd_def indexed_const_def by auto
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    85
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    86
lemma (in ring) indexed_const_in_carrier:
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    87
  assumes "K \<subseteq> carrier R" and "k \<in> K" shows "\<And>m. (indexed_const k) m \<in> carrier R"
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    88
  using assms unfolding indexed_const_def by auto
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    89
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    90
lemma (in ring) indexed_padd_in_carrier:
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    91
  assumes "carrier_coeff P" and "carrier_coeff Q" shows "carrier_coeff (indexed_padd P Q)"
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    92
  using assms unfolding carrier_coeff_def indexed_padd_def by simp
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    93
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    94
lemma (in ring) indexed_pmult_in_carrier:
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    95
  assumes "carrier_coeff P" shows "carrier_coeff (P \<Otimes> i)"
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    96
  using assms unfolding carrier_coeff_def indexed_pmult_def by simp
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    97
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    98
lemma (in ring) indexed_eval_aux_in_carrier:
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    99
  assumes "list_all carrier_coeff Ps" shows "carrier_coeff (indexed_eval_aux Ps i)"
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   100
  using assms unfolding carrier_coeff_def
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   101
  by (induct Ps) (auto simp add: indexed_zero_def indexed_padd_def indexed_pmult_def)
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   102
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   103
lemma (in ring) indexed_eval_in_carrier:
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   104
  assumes "list_all carrier_coeff Ps" shows "carrier_coeff (indexed_eval Ps i)"
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   105
  using assms indexed_eval_aux_in_carrier[of "rev Ps"] by auto
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   106
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   107
lemma (in ring) indexed_pset_in_carrier:
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   108
  assumes "K \<subseteq> carrier R" and "P \<in> (K[\<X>\<^bsub>I\<^esub>])" shows "carrier_coeff P"
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   109
  using assms(2,1) indexed_const_in_carrier unfolding carrier_coeff_def
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   110
  by (induction) (auto simp add: indexed_zero_def indexed_padd_def indexed_pmult_def)
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   111
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   112
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   113
subsection \<open>Indexed Eval\<close>
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 ring) exists_indexed_eval_aux_monomial:
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   116
  assumes "carrier_coeff P" and "list_all carrier_coeff Qs"
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   117
    and "count n i = k" and "P n \<noteq> \<zero>" and "list_all (\<lambda>Q. index_free Q i) Qs"
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   118
  obtains m where "count m i = length Qs + k" and "(indexed_eval_aux (Qs @ [ P ]) i) m \<noteq> \<zero>"
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   119
proof -
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   120
  from assms(2,5) have "\<exists>m. count m i = length Qs + k \<and> (indexed_eval_aux (Qs @ [ P ]) i) m \<noteq> \<zero>"
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   121
  proof (induct Qs)
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   122
    case Nil thus ?case
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   123
      using indexed_padd_zero(2)[OF assms(1)] assms(3-4) by auto
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   124
  next
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   125
    case (Cons Q Qs)
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   126
    then obtain m where m: "count m i = length Qs + k" "(indexed_eval_aux (Qs @ [ P ]) i) m \<noteq> \<zero>"
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   127
      by auto
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   128
    define m' where "m' = m + {# i #}"
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   129
    hence "Q m' = \<zero>"
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   130
      using Cons(3) unfolding index_free_def by simp
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   131
    moreover have "(indexed_eval_aux (Qs @ [ P ]) i) m \<in> carrier R"
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   132
      using indexed_eval_aux_in_carrier[of "Qs @ [ P ]" i] Cons(2) assms(1) carrier_coeffE by auto
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   133
    hence "((indexed_eval_aux (Qs @ [ P ]) i) \<Otimes> i) m' \<in> carrier R - { \<zero> }"
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   134
      using m unfolding indexed_pmult_def m'_def by simp
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   135
    ultimately have "(indexed_eval_aux (Q # (Qs @ [ P ])) i) m' \<noteq> \<zero>"
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   136
      by (auto simp add: indexed_padd_def)
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   137
    moreover from \<open>count m i = length Qs + k\<close> have "count m' i = length (Q # Qs) + k"
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   138
      unfolding m'_def by simp
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   139
    ultimately show ?case
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   140
      by auto
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   141
  qed
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   142
  thus thesis
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   143
    using that by blast
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   144
qed
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   145
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   146
lemma (in ring) indexed_eval_aux_monomial_degree_le:
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   147
  assumes "list_all carrier_coeff Ps" and "list_all (\<lambda>P. index_free P i) Ps"
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   148
    and "(indexed_eval_aux Ps i) m \<noteq> \<zero>" shows "count m i \<le> length Ps - 1"
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   149
  using assms(1-3)
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   150
proof (induct Ps arbitrary: m, simp add: indexed_zero_def)
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   151
  case (Cons P Ps) show ?case
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   152
  proof (cases "count m i = 0", simp)
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   153
    assume "count m i \<noteq> 0"
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   154
    hence "P m = \<zero>"
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   155
      using Cons(3) unfolding index_free_def by simp
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   156
    moreover have "(indexed_eval_aux Ps i) m \<in> carrier R"
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   157
      using carrier_coeffE[OF indexed_eval_aux_in_carrier[of Ps i]] Cons(2) by simp 
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   158
    ultimately have "((indexed_eval_aux Ps i) \<Otimes> i) m \<noteq> \<zero>"
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   159
      using Cons(4) by (auto simp add: indexed_padd_def)
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   160
    with \<open>count m i \<noteq> 0\<close> have "(indexed_eval_aux Ps i) (m - {# i #}) \<noteq> \<zero>"
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   161
      unfolding indexed_pmult_def by (auto simp del: indexed_eval_aux.simps)
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   162
    hence "count m i - 1 \<le> length Ps - 1"
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   163
      using Cons(1)[of "m - {# i #}"] Cons(2-3) by auto
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   164
    moreover from \<open>(indexed_eval_aux Ps i) (m - {# i #}) \<noteq> \<zero>\<close> have "length Ps > 0"
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   165
      by (auto simp add: indexed_zero_def)
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   166
    moreover from \<open>count m i \<noteq> 0\<close> have "count m i > 0"
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   167
      by simp
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   168
    ultimately show ?thesis
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   169
      by (simp add: Suc_leI le_diff_iff)
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   170
  qed
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   171
qed
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   172
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   173
lemma (in ring) indexed_eval_aux_is_inj:
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   174
  assumes "list_all carrier_coeff Ps" and "list_all (\<lambda>P. index_free P i) Ps"
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   175
      and "list_all carrier_coeff Qs" and "list_all (\<lambda>Q. index_free Q i) Qs"
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   176
    and "indexed_eval_aux Ps i = indexed_eval_aux Qs i" and "length Ps = length Qs"
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   177
  shows "Ps = Qs"
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   178
  using assms
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   179
proof (induct Ps arbitrary: Qs, simp)
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   180
  case (Cons P Ps)
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   181
  from \<open>length (P # Ps) = length Qs\<close> obtain Q' Qs' where Qs: "Qs = Q' # Qs'" and "length Ps = length Qs'"
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   182
    by (metis Suc_length_conv)
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   183
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   184
  have in_carrier:
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   185
    "((indexed_eval_aux Ps  i) \<Otimes> i) m \<in> carrier R" "P  m \<in> carrier R"
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   186
    "((indexed_eval_aux Qs' i) \<Otimes> i) m \<in> carrier R" "Q' m \<in> carrier R" for m
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   187
    using indexed_eval_aux_in_carrier[of Ps  i]
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   188
          indexed_eval_aux_in_carrier[of Qs' i] Cons(2,4) carrier_coeffE
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   189
    unfolding Qs indexed_pmult_def by auto
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   190
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   191
  have "(indexed_eval_aux (P # Ps) i) m = (indexed_eval_aux (Q' # Qs') i) m" for m
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   192
    using Cons(6) unfolding Qs by simp
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   193
  hence eq: "((indexed_eval_aux Ps i) \<Otimes> i) m \<oplus> P m = ((indexed_eval_aux Qs' i) \<Otimes> i) m \<oplus> Q' m" for m
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   194
    by (simp add: indexed_padd_def)
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   195
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   196
  have "P m = Q' m" if "i \<in># m" for m
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   197
    using that Cons(3,5) unfolding index_free_def Qs by auto
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   198
  moreover have "P m = Q' m" if "i \<notin># m" for m
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   199
    using in_carrier(2,4) eq[of m] that by (auto simp add: indexed_pmult_def)
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   200
  ultimately have "P = Q'"
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   201
    by auto
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   202
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   203
  hence "(indexed_eval_aux Ps i) m = (indexed_eval_aux Qs' i) m" for m
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   204
    using eq[of "m + {# i #}"] in_carrier[of "m + {# i #}"] unfolding indexed_pmult_def by auto
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   205
  with \<open>length Ps = length Qs'\<close> have "Ps = Qs'"
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   206
    using Cons(1)[of Qs'] Cons(2-5) unfolding Qs by auto
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   207
  with \<open>P = Q'\<close> show ?case
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   208
    unfolding Qs by simp
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   209
qed
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   210
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   211
lemma (in ring) indexed_eval_aux_is_inj':
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   212
  assumes "list_all carrier_coeff Ps" and "list_all (\<lambda>P. index_free P i) Ps"
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   213
      and "list_all carrier_coeff Qs" and "list_all (\<lambda>Q. index_free Q i) Qs"
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   214
      and "carrier_coeff P" and "index_free P i" "P \<noteq> indexed_const \<zero>"
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   215
      and "carrier_coeff Q" and "index_free Q i" "Q \<noteq> indexed_const \<zero>"
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   216
    and "indexed_eval_aux (Ps @ [ P ]) i = indexed_eval_aux (Qs @ [ Q ]) i"
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   217
  shows "Ps = Qs" and "P = Q"
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   218
proof -
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   219
  obtain m n where "P m \<noteq> \<zero>" and "Q n \<noteq> \<zero>"
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   220
    using assms(7,10) unfolding indexed_zero_def by blast
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   221
  hence "count m i = 0" and "count n i = 0"
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   222
    using assms(6,9) unfolding index_free_def by (meson count_inI)+ 
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   223
  with \<open>P m \<noteq> \<zero>\<close> and \<open>Q n \<noteq> \<zero>\<close> obtain m' n'
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   224
    where m': "count m' i = length Ps" "(indexed_eval_aux (Ps @ [ P ]) i) m' \<noteq> \<zero>"
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   225
      and n': "count n' i = length Qs" "(indexed_eval_aux (Qs @ [ Q ]) i) n' \<noteq> \<zero>"
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   226
    using exists_indexed_eval_aux_monomial[of P Ps m i 0]
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   227
          exists_indexed_eval_aux_monomial[of Q Qs n i 0] assms(1-5,8)
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   228
    by (metis (no_types, lifting) add.right_neutral)
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   229
  have "(indexed_eval_aux (Qs @ [ Q ]) i) m' \<noteq> \<zero>"
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   230
    using m'(2) assms(11) by simp
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   231
  with \<open>count m' i = length Ps\<close> have "length Ps \<le> length Qs"
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   232
    using indexed_eval_aux_monomial_degree_le[of "Qs @ [ Q ]" i m'] assms(3-4,8-9) by auto
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   233
  moreover have "(indexed_eval_aux (Ps @ [ P ]) i) n' \<noteq> \<zero>"
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   234
    using n'(2) assms(11) by simp
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   235
  with \<open>count n' i = length Qs\<close> have "length Qs \<le> length Ps"
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   236
    using indexed_eval_aux_monomial_degree_le[of "Ps @ [ P ]" i n'] assms(1-2,5-6) by auto
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   237
  ultimately have same_len: "length (Ps @ [ P ]) = length (Qs @ [ Q ])"
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   238
    by simp
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   239
  thus "Ps = Qs" and "P = Q"
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   240
    using indexed_eval_aux_is_inj[of "Ps @ [ P ]" i "Qs @ [ Q ]"] assms(1-6,8-9,11) by auto
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   241
qed
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   242
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   243
lemma (in ring) exists_indexed_eval_monomial:
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   244
  assumes "carrier_coeff P" and "list_all carrier_coeff Qs"
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   245
    and "P n \<noteq> \<zero>" and "list_all (\<lambda>Q. index_free Q i) Qs"
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   246
  obtains m where "count m i = length Qs + (count n i)" and "(indexed_eval (P # Qs) i) m \<noteq> \<zero>"
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   247
  using exists_indexed_eval_aux_monomial[OF assms(1) _ _ assms(3), of "rev Qs"] assms(2,4) by auto
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
corollary (in ring) exists_indexed_eval_monomial':
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   250
  assumes "carrier_coeff P" and "list_all carrier_coeff Qs"
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   251
    and "P \<noteq> indexed_const \<zero>" and "list_all (\<lambda>Q. index_free Q i) Qs"
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   252
  obtains m where "count m i \<ge> length Qs" and "(indexed_eval (P # Qs) i) m \<noteq> \<zero>"
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
  from \<open>P \<noteq> indexed_const \<zero>\<close> obtain n where "P n \<noteq> \<zero>"
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   255
    unfolding indexed_const_def by auto
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   256
  then obtain m where "count m i = length Qs + (count n i)" and "(indexed_eval (P # Qs) i) m \<noteq> \<zero>"
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   257
    using exists_indexed_eval_monomial[OF assms(1-2) _ assms(4)] by auto
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   258
  thus thesis
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   259
    using that by force
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   260
qed
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   261
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   262
lemma (in ring) indexed_eval_monomial_degree_le:
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   263
  assumes "list_all carrier_coeff Ps" and "list_all (\<lambda>P. index_free P i) Ps"
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   264
    and "(indexed_eval Ps i) m \<noteq> \<zero>" shows "count m i \<le> length Ps - 1"
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   265
  using indexed_eval_aux_monomial_degree_le[of "rev Ps"] assms by auto
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   266
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   267
lemma (in ring) indexed_eval_is_inj:
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   268
  assumes "list_all carrier_coeff Ps" and "list_all (\<lambda>P. index_free P i) Ps"
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   269
      and "list_all carrier_coeff Qs" and "list_all (\<lambda>Q. index_free Q i) Qs"
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   270
      and "carrier_coeff P" and "index_free P i" "P \<noteq> indexed_const \<zero>"
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   271
      and "carrier_coeff Q" and "index_free Q i" "Q \<noteq> indexed_const \<zero>"
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   272
    and "indexed_eval (P # Ps) i = indexed_eval (Q # Qs) i"
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   273
  shows "Ps = Qs" and "P = Q"
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   274
proof -
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   275
  have rev_cond:
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   276
    "list_all carrier_coeff (rev Ps)" "list_all (\<lambda>P. index_free P i) (rev Ps)"
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   277
    "list_all carrier_coeff (rev Qs)" "list_all (\<lambda>Q. index_free Q i) (rev Qs)"
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   278
    using assms(1-4) by auto
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   279
  show "Ps = Qs" and "P = Q"
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   280
    using indexed_eval_aux_is_inj'[OF rev_cond assms(5-10)] assms(11) by auto
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   281
qed
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   282
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   283
lemma (in ring) indexed_eval_inj_on_carrier:
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   284
  assumes "\<And>P. P \<in> carrier L \<Longrightarrow> carrier_coeff P" and "\<And>P. P \<in> carrier L \<Longrightarrow> index_free P i" and "\<zero>\<^bsub>L\<^esub> = indexed_const \<zero>"
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   285
  shows "inj_on (\<lambda>Ps. indexed_eval Ps i) (carrier (poly_ring L))"
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   286
proof -
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   287
  { fix Ps
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   288
    assume "Ps \<in> carrier (poly_ring L)" and "indexed_eval Ps i = indexed_const \<zero>"
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   289
    have "Ps = []"
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   290
    proof (rule ccontr)
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   291
      assume "Ps \<noteq> []"
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   292
      then obtain P' Ps' where Ps: "Ps = P' # Ps'"
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   293
        using list.exhaust by blast
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   294
      with \<open>Ps \<in> carrier (poly_ring L)\<close>
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   295
      have "P' \<noteq> indexed_const \<zero>" and "list_all carrier_coeff Ps" and "list_all (\<lambda>P. index_free P i) Ps"
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   296
        using assms unfolding sym[OF univ_poly_carrier[of L "carrier L"]] polynomial_def
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   297
        by (simp add: list.pred_set subset_code(1))+
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   298
      then obtain m where "(indexed_eval Ps i) m \<noteq> \<zero>"
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   299
        using exists_indexed_eval_monomial'[of P' Ps'] unfolding Ps by auto
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   300
      hence "indexed_eval Ps i \<noteq> indexed_const \<zero>"
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   301
        unfolding indexed_const_def by auto
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   302
      with \<open>indexed_eval Ps i = indexed_const \<zero>\<close> show False by simp
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   303
    qed } note aux_lemma = this
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   304
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   305
  show ?thesis
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   306
  proof (rule inj_onI)
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   307
    fix Ps Qs
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   308
    assume "Ps \<in> carrier (poly_ring L)" and "Qs \<in> carrier (poly_ring L)"
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   309
    show "indexed_eval Ps i = indexed_eval Qs i \<Longrightarrow> Ps = Qs"
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   310
    proof (cases)
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   311
      assume "Qs = []" and "indexed_eval Ps i = indexed_eval Qs i"
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   312
      with \<open>Ps \<in> carrier (poly_ring L)\<close> show "Ps = Qs"
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   313
        using aux_lemma by simp
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   314
    next
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   315
      assume "Qs \<noteq> []" and eq: "indexed_eval Ps i = indexed_eval Qs i"
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   316
      with \<open>Qs \<in> carrier (poly_ring L)\<close> have "Ps \<noteq> []"
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   317
        using aux_lemma by auto
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   318
      from \<open>Ps \<noteq> []\<close> and \<open>Qs \<noteq> []\<close> obtain P' Ps' Q' Qs' where Ps: "Ps = P' # Ps'" and Qs: "Qs = Q' # Qs'"
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   319
        using list.exhaust by metis
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   320
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   321
      from \<open>Ps \<in> carrier (poly_ring L)\<close> and \<open>Ps = P' # Ps'\<close>
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   322
      have "carrier_coeff P'" and "index_free P' i" "P' \<noteq> indexed_const \<zero>"
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   323
       and "list_all carrier_coeff Ps'" and "list_all (\<lambda>P. index_free P i) Ps'"
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   324
        using assms unfolding sym[OF univ_poly_carrier[of L "carrier L"]] polynomial_def
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   325
        by (simp add: list.pred_set subset_code(1))+
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   326
      moreover 
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   327
      from \<open>Qs \<in> carrier (poly_ring L)\<close> and \<open>Qs = Q' # Qs'\<close>
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   328
      have "carrier_coeff Q'" and "index_free Q' i" "Q' \<noteq> indexed_const \<zero>"
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   329
       and "list_all carrier_coeff Qs'" and "list_all (\<lambda>P. index_free P i) Qs'"
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   330
        using assms unfolding sym[OF univ_poly_carrier[of L "carrier L"]] polynomial_def
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   331
        by (simp add: list.pred_set subset_code(1))+
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   332
      ultimately show ?thesis
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   333
        using indexed_eval_is_inj[of Ps' i Qs' P' Q'] eq unfolding Ps Qs by auto
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   334
    qed
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   335
  qed
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   336
qed
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   337
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   338
70162
13b10ca71150 markup fixes
paulson <lp15@cam.ac.uk>
parents: 70160
diff changeset
   339
subsection \<open>Link with Weak Morphisms\<close>
70160
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   340
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   341
text \<open>We study some elements of the contradiction needed in the algebraic closure existence proof. \<close>
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
context ring
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   344
begin
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
lemma (in ring) indexed_padd_index_free:
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   347
  assumes "index_free P i" and "index_free Q i" shows "index_free (P \<Oplus> Q) i"
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   348
  using assms unfolding indexed_padd_def index_free_def by auto
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   349
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   350
lemma (in ring) indexed_pmult_index_free:
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   351
  assumes "index_free P j" and "i \<noteq> j" shows "index_free (P \<Otimes> i) j"
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   352
  using assms unfolding index_free_def indexed_pmult_def
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   353
  by (metis insert_DiffM insert_noteq_member)
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   354
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   355
lemma (in ring) indexed_eval_index_free:
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   356
  assumes "list_all (\<lambda>P. index_free P j) Ps" and "i \<noteq> j" shows "index_free (indexed_eval Ps i) j"
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   357
proof -
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   358
  { fix Ps assume "list_all (\<lambda>P. index_free P j) Ps" hence "index_free (indexed_eval_aux Ps i) j"
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   359
      using indexed_padd_index_free[OF indexed_pmult_index_free[OF _ assms(2)]]
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   360
      by (induct Ps) (auto simp add: indexed_zero_def index_free_def) }
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   361
  thus ?thesis
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   362
    using assms(1) by auto
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   363
qed
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   364
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   365
context
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   366
  fixes L :: "(('c multiset) \<Rightarrow> 'a) ring" and i :: 'c
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   367
  assumes hyps:
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   368
    \<comment> \<open>i\<close>   "field L"
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   369
    \<comment> \<open>ii\<close>  "\<And>P. P \<in> carrier L \<Longrightarrow> carrier_coeff P"
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   370
    \<comment> \<open>iii\<close> "\<And>P. P \<in> carrier L \<Longrightarrow> index_free P i"
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   371
    \<comment> \<open>iv\<close>  "\<zero>\<^bsub>L\<^esub> = indexed_const \<zero>"
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   372
begin
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   373
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   374
interpretation L: field L
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   375
  using \<open>field L\<close> .
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   376
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   377
interpretation UP: principal_domain "poly_ring L"
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   378
  using L.univ_poly_is_principal[OF L.carrier_is_subfield] .
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   379
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   380
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   381
abbreviation eval_pmod
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   382
  where "eval_pmod q \<equiv> (\<lambda>p. indexed_eval (L.pmod p q) i)"
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   383
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   384
abbreviation image_poly
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   385
  where "image_poly q \<equiv> image_ring (eval_pmod q) (poly_ring L)"
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   386
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
lemma indexed_eval_is_weak_ring_morphism:
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   389
  assumes "q \<in> carrier (poly_ring L)" shows "weak_ring_morphism (eval_pmod q) (PIdl\<^bsub>poly_ring L\<^esub> q) (poly_ring L)"
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   390
proof (rule weak_ring_morphismI)
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   391
  show "ideal (PIdl\<^bsub>poly_ring L\<^esub> q) (poly_ring L)"
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   392
    using UP.cgenideal_ideal[OF assms] .
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   393
next
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   394
  fix a b assume in_carrier: "a \<in> carrier (poly_ring L)" "b \<in> carrier (poly_ring L)"
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   395
  note ldiv_closed = in_carrier[THEN L.long_division_closed(2)[OF L.carrier_is_subfield _ assms]]
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   396
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   397
  have "(eval_pmod q) a = (eval_pmod q) b \<longleftrightarrow> L.pmod a q = L.pmod b q"
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   398
    using inj_onD[OF indexed_eval_inj_on_carrier[OF hyps(2-4)] _ ldiv_closed] by fastforce
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   399
  also have " ... \<longleftrightarrow> q pdivides\<^bsub>L\<^esub> (a \<ominus>\<^bsub>poly_ring L\<^esub> b)"
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   400
    unfolding L.same_pmod_iff_pdivides[OF L.carrier_is_subfield in_carrier assms] ..
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   401
  also have " ... \<longleftrightarrow> PIdl\<^bsub>poly_ring L\<^esub> (a \<ominus>\<^bsub>poly_ring L\<^esub> b) \<subseteq> PIdl\<^bsub>poly_ring L\<^esub> q"
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   402
    unfolding UP.to_contain_is_to_divide[OF assms UP.minus_closed[OF in_carrier]] pdivides_def ..
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   403
  also have " ... \<longleftrightarrow> a \<ominus>\<^bsub>poly_ring L\<^esub> b \<in> PIdl\<^bsub>poly_ring L\<^esub> q"
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   404
    unfolding UP.cgenideal_eq_genideal[OF assms] UP.cgenideal_eq_genideal[OF UP.minus_closed[OF in_carrier]]
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   405
              UP.Idl_subset_ideal'[OF UP.minus_closed[OF in_carrier] assms] ..
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   406
  finally show "(eval_pmod q) a = (eval_pmod q) b \<longleftrightarrow> a \<ominus>\<^bsub>poly_ring L\<^esub> b \<in> PIdl\<^bsub>poly_ring L\<^esub> q" .
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 eval_norm_eq_id:
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   410
  assumes "q \<in> carrier (poly_ring L)" and "degree q > 0" and "a \<in> carrier L"
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   411
  shows "((eval_pmod q) \<circ> (ring.poly_of_const L)) a = a"
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   412
proof (cases)
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   413
  assume "a = \<zero>\<^bsub>L\<^esub>" thus ?thesis
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   414
    using L.long_division_zero(2)[OF L.carrier_is_subfield assms(1)] hyps(4)
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   415
    unfolding ring.poly_of_const_def[OF L.ring_axioms] by auto
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   416
next
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   417
  assume "a \<noteq> \<zero>\<^bsub>L\<^esub>" then have in_carrier: "[ a ] \<in> carrier (poly_ring L)"
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   418
    using assms(3) unfolding sym[OF univ_poly_carrier[of L "carrier L"]] polynomial_def by simp
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   419
  from \<open>a \<noteq> \<zero>\<^bsub>L\<^esub>\<close> show ?thesis
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   420
    using L.pmod_const(2)[OF L.carrier_is_subfield in_carrier assms(1)] assms(2)
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   421
          indexed_padd_zero(2)[OF hyps(2)[OF assms(3)]]
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   422
    unfolding ring.poly_of_const_def[OF L.ring_axioms] by auto
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   423
qed
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   424
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   425
lemma image_poly_iso_incl:
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   426
  assumes "q \<in> carrier (poly_ring L)" and "degree q > 0" shows "id \<in> ring_hom L (image_poly q)"
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   427
proof -
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   428
  have "((eval_pmod q) \<circ> L.poly_of_const) \<in> ring_hom L (image_poly q)"
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   429
    using ring_hom_trans[OF L.canonical_embedding_is_hom[OF L.carrier_is_subring]
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   430
          UP.weak_ring_morphism_is_hom[OF indexed_eval_is_weak_ring_morphism[OF assms(1)]]]
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   431
    by simp
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   432
  thus ?thesis
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   433
    using eval_norm_eq_id[OF assms(1-2)] L.ring_hom_restrict[of _ "image_poly q" id] by auto
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   434
qed
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   435
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   436
lemma image_poly_is_field:
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   437
  assumes "q \<in> carrier (poly_ring L)" and "pirreducible\<^bsub>L\<^esub> (carrier L) q" shows "field (image_poly q)"
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   438
  using UP.image_ring_is_field[OF indexed_eval_is_weak_ring_morphism[OF assms(1)]] assms(2)
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   439
  unfolding sym[OF L.rupture_is_field_iff_pirreducible[OF L.carrier_is_subfield assms(1)]] rupture_def
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   440
  by simp
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   441
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   442
lemma image_poly_index_free:
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   443
  assumes "q \<in> carrier (poly_ring L)" and "P \<in> carrier (image_poly q)" and "\<not> index_free P j" "i \<noteq> j"
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   444
  obtains Q where "Q \<in> carrier L" and "\<not> index_free Q j"
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   445
proof -
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   446
  from \<open>P \<in> carrier (image_poly q)\<close> obtain p where p: "p \<in> carrier (poly_ring L)" and P: "P = (eval_pmod q) p"
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   447
    unfolding image_ring_carrier by blast
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   448
  from \<open>\<not> index_free P j\<close> have "\<not> list_all (\<lambda>P. index_free P j) (L.pmod p q)"
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   449
    using indexed_eval_index_free[OF _ assms(4), of "L.pmod p q"] unfolding sym[OF P] by auto
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   450
  then obtain Q where "Q \<in> set (L.pmod p q)" and "\<not> index_free Q j"
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   451
    unfolding list_all_iff by auto
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   452
  thus ?thesis
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   453
    using L.long_division_closed(2)[OF L.carrier_is_subfield p assms(1)] that
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   454
    unfolding sym[OF univ_poly_carrier[of L "carrier L"]] polynomial_def
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   455
    by auto
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   456
qed
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
lemma eval_pmod_var:
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   459
  assumes "indexed_const \<in> ring_hom R L" and "q \<in> carrier (poly_ring L)" and "degree q > 1"
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   460
  shows "(eval_pmod q) X\<^bsub>L\<^esub> = \<X>\<^bsub>i\<^esub>" and "\<X>\<^bsub>i\<^esub> \<in> carrier (image_poly q)"
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   461
proof -
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   462
  have "X\<^bsub>L\<^esub> = [ indexed_const \<one>, indexed_const \<zero> ]" and "X\<^bsub>L\<^esub> \<in> carrier (poly_ring L)"
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   463
    using ring_hom_one[OF assms(1)] hyps(4) L.var_closed(1) L.carrier_is_subring unfolding var_def by auto
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   464
  thus "(eval_pmod q) X\<^bsub>L\<^esub> = \<X>\<^bsub>i\<^esub>"
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   465
    using L.pmod_const(2)[OF L.carrier_is_subfield _ assms(2), of "X\<^bsub>L\<^esub>"] assms(3)
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   466
    by (auto simp add: indexed_pmult_def indexed_padd_def indexed_const_def indexed_var_def)
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   467
  with \<open>X\<^bsub>L\<^esub> \<in> carrier (poly_ring L)\<close> show "\<X>\<^bsub>i\<^esub> \<in> carrier (image_poly q)"
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   468
    using image_iff unfolding image_ring_carrier by fastforce
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   469
qed
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   470
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   471
lemma image_poly_eval_indexed_var:
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   472
  assumes "indexed_const \<in> ring_hom R L"
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   473
    and "q \<in> carrier (poly_ring L)" and "degree q > 1" and "pirreducible\<^bsub>L\<^esub> (carrier L) q"
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   474
  shows "(ring.eval (image_poly q)) q \<X>\<^bsub>i\<^esub> = \<zero>\<^bsub>image_poly q\<^esub>"
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   475
proof -
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   476
  let ?surj = "L.rupture_surj (carrier L) q"
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   477
  let ?Rupt = "Rupt\<^bsub>L\<^esub> (carrier L) q"
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   478
  let ?f = "eval_pmod q"
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   479
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   480
  interpret UP: ring "poly_ring L"
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   481
    using L.univ_poly_is_ring[OF L.carrier_is_subring] .
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   482
  from \<open>pirreducible\<^bsub>L\<^esub> (carrier L) q\<close> interpret Rupt: field ?Rupt
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   483
    using L.rupture_is_field_iff_pirreducible[OF L.carrier_is_subfield assms(2)] by simp
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   484
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   485
  have weak_morphism: "weak_ring_morphism ?f (PIdl\<^bsub>poly_ring L\<^esub> q) (poly_ring L)"
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   486
    using indexed_eval_is_weak_ring_morphism[OF assms(2)] .
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   487
  then interpret I: ideal "PIdl\<^bsub>poly_ring L\<^esub> q" "poly_ring L"
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   488
    using weak_ring_morphism.axioms(1) by auto
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   489
  interpret Hom: ring_hom_ring ?Rupt "image_poly q" "\<lambda>x. the_elem (?f ` x)"
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   490
    using ring_hom_ring.intro[OF I.quotient_is_ring UP.image_ring_is_ring[OF weak_morphism]]
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   491
          UP.weak_ring_morphism_is_iso[OF weak_morphism]
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   492
    unfolding ring_iso_def symmetric[OF ring_hom_ring_axioms_def] rupture_def
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   493
    by auto
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   494
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   495
  have "set q \<subseteq> carrier L" and lc: "q \<noteq> [] \<Longrightarrow> lead_coeff q \<in> carrier L - { \<zero>\<^bsub>L\<^esub> }"
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   496
    using assms(2) 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
   497
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   498
  have map_surj: "set (map (?surj \<circ> L.poly_of_const) q) \<subseteq> carrier ?Rupt"
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   499
  proof -
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   500
    have "L.poly_of_const a \<in> carrier (poly_ring L)" if "a \<in> carrier L" for a
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   501
      using that L.normalize_gives_polynomial[of "[ a ]"]
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   502
      unfolding univ_poly_carrier ring.poly_of_const_def[OF L.ring_axioms] by simp
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   503
    hence "(?surj \<circ> L.poly_of_const) a \<in> carrier ?Rupt" if "a \<in> carrier L" for a
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   504
      using ring_hom_memE(1)[OF L.rupture_surj_hom(1)[OF L.carrier_is_subring assms(2)]] that by simp
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   505
    with \<open>set q \<subseteq> carrier L\<close> show ?thesis
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   506
      by (induct q) (auto)
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   507
  qed
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   508
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   509
  have "?surj X\<^bsub>L\<^esub> \<in> carrier ?Rupt"
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   510
    using ring_hom_memE(1)[OF L.rupture_surj_hom(1)[OF _ assms(2)] L.var_closed(1)] L.carrier_is_subring by simp
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   511
  moreover have "map (\<lambda>x. the_elem (?f ` x)) (map (?surj \<circ> L.poly_of_const) q) = q"
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   512
  proof -
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   513
    define g where "g = (?surj \<circ> L.poly_of_const)"
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   514
    define f where "f = (\<lambda>x. the_elem (?f ` x))"
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   515
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   516
    have "the_elem (?f ` ((?surj \<circ> L.poly_of_const) a)) = ((eval_pmod q) \<circ> L.poly_of_const) a"
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   517
      if "a \<in> carrier L" for a
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   518
      using that L.normalize_gives_polynomial[of "[ a ]"] UP.weak_ring_morphism_range[OF weak_morphism]
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   519
      unfolding univ_poly_carrier ring.poly_of_const_def[OF L.ring_axioms] by auto
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   520
    hence "the_elem (?f ` ((?surj \<circ> L.poly_of_const) a)) = a" if "a \<in> carrier L" for a
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   521
      using eval_norm_eq_id[OF assms(2)] that assms(3) by simp
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   522
    hence "f (g a) = a" if "a \<in> carrier L" for a
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   523
      using that unfolding f_def g_def by simp
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   524
    with \<open>set q \<subseteq> carrier L\<close> have "map f (map g q) = q"
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   525
      by (induct q) (auto)
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   526
    thus ?thesis
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   527
      unfolding f_def g_def by simp
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   528
  qed
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   529
  moreover have "(\<lambda>x. the_elem (?f ` x)) (?surj X\<^bsub>L\<^esub>) = \<X>\<^bsub>i\<^esub>"
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   530
    using UP.weak_ring_morphism_range[OF weak_morphism L.var_closed(1)[OF L.carrier_is_subring]]
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   531
    unfolding eval_pmod_var(1)[OF assms(1-3)] by simp
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   532
  ultimately have "Hom.S.eval q \<X>\<^bsub>i\<^esub> = (\<lambda>x. the_elem (?f ` x)) (Rupt.eval (map (?surj \<circ> L.poly_of_const) q) (?surj X\<^bsub>L\<^esub>))"
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   533
    using Hom.eval_hom'[OF _ map_surj] by auto
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   534
  moreover have "\<zero>\<^bsub>?Rupt\<^esub> = ?surj \<zero>\<^bsub>poly_ring L\<^esub>"
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   535
    unfolding rupture_def FactRing_def by (simp add: I.a_rcos_const)
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   536
  hence "the_elem (?f ` \<zero>\<^bsub>?Rupt\<^esub>) = \<zero>\<^bsub>image_poly q\<^esub>"
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   537
    using UP.weak_ring_morphism_range[OF weak_morphism UP.zero_closed]
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   538
    unfolding image_ring_zero by simp 
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   539
  hence "(\<lambda>x. the_elem (?f ` x)) (Rupt.eval (map (?surj \<circ> L.poly_of_const) q) (?surj X\<^bsub>L\<^esub>)) = \<zero>\<^bsub>image_poly q\<^esub>"
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   540
    using L.polynomial_rupture[OF L.carrier_is_subring assms(2)] by simp
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   541
  ultimately show ?thesis
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   542
    by simp
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   543
qed
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   544
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   545
end (* of fixed L context. *)
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   546
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   547
end (* of ring context. *)
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   548
8e9100dcde52 Towards a proof of algebraic closure (NB not finished)
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   549
end