src/HOL/Algebra/UnivPoly.thy
author wenzelm
Fri, 13 Nov 2009 15:38:45 +0100
changeset 33659 2d7ab9458518
parent 33657 a4179bf442d1
child 34915 7894c7dab132
permissions -rw-r--r--
more "anti_sym" -> "antisym" (cf. a4179bf442d1);
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
     1
(*
14706
71590b7733b7 tuned document;
wenzelm
parents: 14666
diff changeset
     2
  Title:     HOL/Algebra/UnivPoly.thy
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
     3
  Author:    Clemens Ballarin, started 9 December 1996
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
     4
  Copyright: Clemens Ballarin
27717
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
     5
27933
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
     6
Contributions, in particular on long division, by Jesus Aransay.
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
     7
*)
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
     8
28823
dcbef866c9e2 tuned unfold_locales invocation
haftmann
parents: 27933
diff changeset
     9
theory UnivPoly
dcbef866c9e2 tuned unfold_locales invocation
haftmann
parents: 27933
diff changeset
    10
imports Module RingHom
dcbef866c9e2 tuned unfold_locales invocation
haftmann
parents: 27933
diff changeset
    11
begin
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
    12
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents: 20282
diff changeset
    13
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents: 20282
diff changeset
    14
section {* Univariate Polynomials *}
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
    15
14553
4740fc2da7bb Added brief intro text.
ballarin
parents: 14399
diff changeset
    16
text {*
14666
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
    17
  Polynomials are formalised as modules with additional operations for
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
    18
  extracting coefficients from polynomials and for obtaining monomials
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
    19
  from coefficients and exponents (record @{text "up_ring"}).  The
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
    20
  carrier set is a set of bounded functions from Nat to the
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
    21
  coefficient domain.  Bounded means that these functions return zero
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
    22
  above a certain bound (the degree).  There is a chapter on the
14706
71590b7733b7 tuned document;
wenzelm
parents: 14666
diff changeset
    23
  formalisation of polynomials in the PhD thesis \cite{Ballarin:1999},
71590b7733b7 tuned document;
wenzelm
parents: 14666
diff changeset
    24
  which was implemented with axiomatic type classes.  This was later
71590b7733b7 tuned document;
wenzelm
parents: 14666
diff changeset
    25
  ported to Locales.
14553
4740fc2da7bb Added brief intro text.
ballarin
parents: 14399
diff changeset
    26
*}
4740fc2da7bb Added brief intro text.
ballarin
parents: 14399
diff changeset
    27
14666
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
    28
13949
0ce528cd6f19 HOL-Algebra complete for release Isabelle2003 (modulo section headers).
ballarin
parents: 13940
diff changeset
    29
subsection {* The Constructor for Univariate Polynomials *}
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
    30
15095
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
    31
text {*
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
    32
  Functions with finite support.
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
    33
*}
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
    34
14666
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
    35
locale bound =
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
    36
  fixes z :: 'a
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
    37
    and n :: nat
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
    38
    and f :: "nat => 'a"
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
    39
  assumes bound: "!!m. n < m \<Longrightarrow> f m = z"
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
    40
14666
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
    41
declare bound.intro [intro!]
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
    42
  and bound.bound [dest]
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
    43
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
    44
lemma bound_below:
14666
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
    45
  assumes bound: "bound z m f" and nonzero: "f n \<noteq> z" shows "n \<le> m"
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
    46
proof (rule classical)
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
    47
  assume "~ ?thesis"
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
    48
  then have "m < n" by arith
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
    49
  with bound have "f n = z" ..
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
    50
  with nonzero show ?thesis by contradiction
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
    51
qed
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
    52
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
    53
record ('a, 'p) up_ring = "('a, 'p) module" +
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
    54
  monom :: "['a, nat] => 'p"
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
    55
  coeff :: "['p, nat] => 'a"
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
    56
27933
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
    57
definition up :: "('a, 'm) ring_scheme => (nat => 'a) set"
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
    58
  where up_def: "up R == {f. f \<in> UNIV -> carrier R & (EX n. bound \<zero>\<^bsub>R\<^esub> n f)}"
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
    59
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
    60
definition UP :: "('a, 'm) ring_scheme => ('a, nat => 'a) up_ring"
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
    61
  where UP_def: "UP R == (|
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
    62
   carrier = up R,
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
    63
   mult = (%p:up R. %q:up R. %n. \<Oplus>\<^bsub>R\<^esub>i \<in> {..n}. p i \<otimes>\<^bsub>R\<^esub> q (n-i)),
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
    64
   one = (%i. if i=0 then \<one>\<^bsub>R\<^esub> else \<zero>\<^bsub>R\<^esub>),
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
    65
   zero = (%i. \<zero>\<^bsub>R\<^esub>),
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
    66
   add = (%p:up R. %q:up R. %i. p i \<oplus>\<^bsub>R\<^esub> q i),
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
    67
   smult = (%a:carrier R. %p:up R. %i. a \<otimes>\<^bsub>R\<^esub> p i),
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
    68
   monom = (%a:carrier R. %n i. if i=n then a else \<zero>\<^bsub>R\<^esub>),
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
    69
   coeff = (%p:up R. %n. p n) |)"
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
    70
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
    71
text {*
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
    72
  Properties of the set of polynomials @{term up}.
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
    73
*}
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
    74
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
    75
lemma mem_upI [intro]:
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
    76
  "[| !!n. f n \<in> carrier R; EX n. bound (zero R) n f |] ==> f \<in> up R"
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
    77
  by (simp add: up_def Pi_def)
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
    78
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
    79
lemma mem_upD [dest]:
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
    80
  "f \<in> up R ==> f n \<in> carrier R"
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
    81
  by (simp add: up_def Pi_def)
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
    82
27717
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
    83
context ring
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
    84
begin
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
    85
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
    86
lemma bound_upD [dest]: "f \<in> up R ==> EX n. bound \<zero> n f" by (simp add: up_def)
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
    87
27717
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
    88
lemma up_one_closed: "(%n. if n = 0 then \<one> else \<zero>) \<in> up R" using up_def by force
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
    89
27717
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
    90
lemma up_smult_closed: "[| a \<in> carrier R; p \<in> up R |] ==> (%i. a \<otimes> p i) \<in> up R" by force
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
    91
27717
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
    92
lemma up_add_closed:
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
    93
  "[| p \<in> up R; q \<in> up R |] ==> (%i. p i \<oplus> q i) \<in> up R"
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
    94
proof
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
    95
  fix n
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
    96
  assume "p \<in> up R" and "q \<in> up R"
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
    97
  then show "p n \<oplus> q n \<in> carrier R"
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
    98
    by auto
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
    99
next
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   100
  assume UP: "p \<in> up R" "q \<in> up R"
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   101
  show "EX n. bound \<zero> n (%i. p i \<oplus> q i)"
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   102
  proof -
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   103
    from UP obtain n where boundn: "bound \<zero> n p" by fast
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   104
    from UP obtain m where boundm: "bound \<zero> m q" by fast
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   105
    have "bound \<zero> (max n m) (%i. p i \<oplus> q i)"
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   106
    proof
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   107
      fix i
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   108
      assume "max n m < i"
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   109
      with boundn and boundm and UP show "p i \<oplus> q i = \<zero>" by fastsimp
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   110
    qed
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   111
    then show ?thesis ..
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   112
  qed
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   113
qed
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   114
27717
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
   115
lemma up_a_inv_closed:
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   116
  "p \<in> up R ==> (%i. \<ominus> (p i)) \<in> up R"
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   117
proof
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   118
  assume R: "p \<in> up R"
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   119
  then obtain n where "bound \<zero> n p" by auto
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   120
  then have "bound \<zero> n (%i. \<ominus> p i)" by auto
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   121
  then show "EX n. bound \<zero> n (%i. \<ominus> p i)" by auto
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   122
qed auto
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   123
27933
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
   124
lemma up_minus_closed:
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
   125
  "[| p \<in> up R; q \<in> up R |] ==> (%i. p i \<ominus> q i) \<in> up R"
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
   126
  using mem_upD [of p R] mem_upD [of q R] up_add_closed up_a_inv_closed a_minus_def [of _ R]
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
   127
  by auto
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
   128
27717
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
   129
lemma up_mult_closed:
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   130
  "[| p \<in> up R; q \<in> up R |] ==>
14666
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
   131
  (%n. \<Oplus>i \<in> {..n}. p i \<otimes> q (n-i)) \<in> up R"
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   132
proof
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   133
  fix n
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   134
  assume "p \<in> up R" "q \<in> up R"
14666
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
   135
  then show "(\<Oplus>i \<in> {..n}. p i \<otimes> q (n-i)) \<in> carrier R"
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   136
    by (simp add: mem_upD  funcsetI)
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   137
next
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   138
  assume UP: "p \<in> up R" "q \<in> up R"
14666
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
   139
  show "EX n. bound \<zero> n (%n. \<Oplus>i \<in> {..n}. p i \<otimes> q (n-i))"
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   140
  proof -
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   141
    from UP obtain n where boundn: "bound \<zero> n p" by fast
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   142
    from UP obtain m where boundm: "bound \<zero> m q" by fast
14666
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
   143
    have "bound \<zero> (n + m) (%n. \<Oplus>i \<in> {..n}. p i \<otimes> q (n - i))"
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   144
    proof
14666
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
   145
      fix k assume bound: "n + m < k"
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   146
      {
14666
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
   147
        fix i
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
   148
        have "p i \<otimes> q (k-i) = \<zero>"
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
   149
        proof (cases "n < i")
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
   150
          case True
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
   151
          with boundn have "p i = \<zero>" by auto
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   152
          moreover from UP have "q (k-i) \<in> carrier R" by auto
14666
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
   153
          ultimately show ?thesis by simp
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
   154
        next
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
   155
          case False
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
   156
          with bound have "m < k-i" by arith
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
   157
          with boundm have "q (k-i) = \<zero>" by auto
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
   158
          moreover from UP have "p i \<in> carrier R" by auto
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
   159
          ultimately show ?thesis by simp
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
   160
        qed
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   161
      }
14666
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
   162
      then show "(\<Oplus>i \<in> {..k}. p i \<otimes> q (k-i)) = \<zero>"
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
   163
        by (simp add: Pi_def)
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   164
    qed
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   165
    then show ?thesis by fast
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   166
  qed
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   167
qed
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   168
27717
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
   169
end
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
   170
14666
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
   171
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents: 20282
diff changeset
   172
subsection {* Effect of Operations on Coefficients *}
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   173
19783
82f365a14960 Improved parameter management of locales.
ballarin
parents: 19582
diff changeset
   174
locale UP =
82f365a14960 Improved parameter management of locales.
ballarin
parents: 19582
diff changeset
   175
  fixes R (structure) and P (structure)
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   176
  defines P_def: "P == UP R"
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   177
29240
bb81c3709fb6 More porting to new locales.
ballarin
parents: 29237
diff changeset
   178
locale UP_ring = UP + R: ring R
27717
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
   179
29240
bb81c3709fb6 More porting to new locales.
ballarin
parents: 29237
diff changeset
   180
locale UP_cring = UP + R: cring R
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   181
29237
e90d9d51106b More porting to new locales.
ballarin
parents: 28823
diff changeset
   182
sublocale UP_cring < UP_ring
29240
bb81c3709fb6 More porting to new locales.
ballarin
parents: 29237
diff changeset
   183
  by intro_locales [1] (rule P_def)
27717
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
   184
29240
bb81c3709fb6 More porting to new locales.
ballarin
parents: 29237
diff changeset
   185
locale UP_domain = UP + R: "domain" R
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   186
29237
e90d9d51106b More porting to new locales.
ballarin
parents: 28823
diff changeset
   187
sublocale UP_domain < UP_cring
29240
bb81c3709fb6 More porting to new locales.
ballarin
parents: 29237
diff changeset
   188
  by intro_locales [1] (rule P_def)
27717
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
   189
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
   190
context UP
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
   191
begin
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
   192
30363
9b8d9b6ef803 proper local context for text with antiquotations;
wenzelm
parents: 29246
diff changeset
   193
text {*Temporarily declare @{thm P_def} as simp rule.*}
27717
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
   194
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
   195
declare P_def [simp]
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   196
27717
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
   197
lemma up_eqI:
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
   198
  assumes prem: "!!n. coeff P p n = coeff P q n" and R: "p \<in> carrier P" "q \<in> carrier P"
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
   199
  shows "p = q"
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
   200
proof
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
   201
  fix x
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
   202
  from prem and R show "p x = q x" by (simp add: UP_def)
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
   203
qed
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   204
27717
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
   205
lemma coeff_closed [simp]:
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
   206
  "p \<in> carrier P ==> coeff P p n \<in> carrier R" by (auto simp add: UP_def)
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
   207
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
   208
end
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
   209
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
   210
context UP_ring 
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
   211
begin
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
   212
27933
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
   213
(* Theorems generalised from commutative rings to rings by Jesus Aransay. *)
27717
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
   214
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
   215
lemma coeff_monom [simp]:
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
   216
  "a \<in> carrier R ==> coeff P (monom P a m) n = (if m=n then a else \<zero>)"
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   217
proof -
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   218
  assume R: "a \<in> carrier R"
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   219
  then have "(%n. if n = m then a else \<zero>) \<in> up R"
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   220
    using up_def by force
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   221
  with R show ?thesis by (simp add: UP_def)
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   222
qed
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   223
27717
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
   224
lemma coeff_zero [simp]: "coeff P \<zero>\<^bsub>P\<^esub> n = \<zero>" by (auto simp add: UP_def)
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   225
27717
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
   226
lemma coeff_one [simp]: "coeff P \<one>\<^bsub>P\<^esub> n = (if n=0 then \<one> else \<zero>)"
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   227
  using up_one_closed by (simp add: UP_def)
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   228
27717
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
   229
lemma coeff_smult [simp]:
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
   230
  "[| a \<in> carrier R; p \<in> carrier P |] ==> coeff P (a \<odot>\<^bsub>P\<^esub> p) n = a \<otimes> coeff P p n"
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   231
  by (simp add: UP_def up_smult_closed)
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   232
27717
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
   233
lemma coeff_add [simp]:
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
   234
  "[| p \<in> carrier P; q \<in> carrier P |] ==> coeff P (p \<oplus>\<^bsub>P\<^esub> q) n = coeff P p n \<oplus> coeff P q n"
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   235
  by (simp add: UP_def up_add_closed)
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   236
27717
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
   237
lemma coeff_mult [simp]:
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
   238
  "[| p \<in> carrier P; q \<in> carrier P |] ==> coeff P (p \<otimes>\<^bsub>P\<^esub> q) n = (\<Oplus>i \<in> {..n}. coeff P p i \<otimes> coeff P q (n-i))"
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   239
  by (simp add: UP_def up_mult_closed)
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   240
27717
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
   241
end
14666
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
   242
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents: 20282
diff changeset
   243
27717
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
   244
subsection {* Polynomials Form a Ring. *}
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
   245
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
   246
context UP_ring
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
   247
begin
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   248
14666
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
   249
text {* Operations are closed over @{term P}. *}
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   250
27717
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
   251
lemma UP_mult_closed [simp]:
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
   252
  "[| p \<in> carrier P; q \<in> carrier P |] ==> p \<otimes>\<^bsub>P\<^esub> q \<in> carrier P" by (simp add: UP_def up_mult_closed)
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   253
27717
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
   254
lemma UP_one_closed [simp]:
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
   255
  "\<one>\<^bsub>P\<^esub> \<in> carrier P" by (simp add: UP_def up_one_closed)
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   256
27717
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
   257
lemma UP_zero_closed [intro, simp]:
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
   258
  "\<zero>\<^bsub>P\<^esub> \<in> carrier P" by (auto simp add: UP_def)
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   259
27717
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
   260
lemma UP_a_closed [intro, simp]:
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
   261
  "[| p \<in> carrier P; q \<in> carrier P |] ==> p \<oplus>\<^bsub>P\<^esub> q \<in> carrier P" by (simp add: UP_def up_add_closed)
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   262
27717
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
   263
lemma monom_closed [simp]:
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
   264
  "a \<in> carrier R ==> monom P a n \<in> carrier P" by (auto simp add: UP_def up_def Pi_def)
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   265
27717
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
   266
lemma UP_smult_closed [simp]:
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
   267
  "[| a \<in> carrier R; p \<in> carrier P |] ==> a \<odot>\<^bsub>P\<^esub> p \<in> carrier P" by (simp add: UP_def up_smult_closed)
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
   268
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
   269
end
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   270
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   271
declare (in UP) P_def [simp del]
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   272
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   273
text {* Algebraic ring properties *}
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   274
27717
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
   275
context UP_ring
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
   276
begin
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   277
27717
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
   278
lemma UP_a_assoc:
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
   279
  assumes R: "p \<in> carrier P" "q \<in> carrier P" "r \<in> carrier P"
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
   280
  shows "(p \<oplus>\<^bsub>P\<^esub> q) \<oplus>\<^bsub>P\<^esub> r = p \<oplus>\<^bsub>P\<^esub> (q \<oplus>\<^bsub>P\<^esub> r)" by (rule up_eqI, simp add: a_assoc R, simp_all add: R)
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
   281
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
   282
lemma UP_l_zero [simp]:
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   283
  assumes R: "p \<in> carrier P"
27717
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
   284
  shows "\<zero>\<^bsub>P\<^esub> \<oplus>\<^bsub>P\<^esub> p = p" by (rule up_eqI, simp_all add: R)
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   285
27717
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
   286
lemma UP_l_neg_ex:
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   287
  assumes R: "p \<in> carrier P"
15095
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
   288
  shows "EX q : carrier P. q \<oplus>\<^bsub>P\<^esub> p = \<zero>\<^bsub>P\<^esub>"
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   289
proof -
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   290
  let ?q = "%i. \<ominus> (p i)"
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   291
  from R have closed: "?q \<in> carrier P"
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   292
    by (simp add: UP_def P_def up_a_inv_closed)
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   293
  from R have coeff: "!!n. coeff P ?q n = \<ominus> (coeff P p n)"
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   294
    by (simp add: UP_def P_def up_a_inv_closed)
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   295
  show ?thesis
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   296
  proof
15095
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
   297
    show "?q \<oplus>\<^bsub>P\<^esub> p = \<zero>\<^bsub>P\<^esub>"
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   298
      by (auto intro!: up_eqI simp add: R closed coeff R.l_neg)
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   299
  qed (rule closed)
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   300
qed
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   301
27717
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
   302
lemma UP_a_comm:
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   303
  assumes R: "p \<in> carrier P" "q \<in> carrier P"
27717
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
   304
  shows "p \<oplus>\<^bsub>P\<^esub> q = q \<oplus>\<^bsub>P\<^esub> p" by (rule up_eqI, simp add: a_comm R, simp_all add: R)
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
   305
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
   306
lemma UP_m_assoc:
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   307
  assumes R: "p \<in> carrier P" "q \<in> carrier P" "r \<in> carrier P"
15095
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
   308
  shows "(p \<otimes>\<^bsub>P\<^esub> q) \<otimes>\<^bsub>P\<^esub> r = p \<otimes>\<^bsub>P\<^esub> (q \<otimes>\<^bsub>P\<^esub> r)"
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   309
proof (rule up_eqI)
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   310
  fix n
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   311
  {
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   312
    fix k and a b c :: "nat=>'a"
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   313
    assume R: "a \<in> UNIV -> carrier R" "b \<in> UNIV -> carrier R"
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   314
      "c \<in> UNIV -> carrier R"
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   315
    then have "k <= n ==>
14666
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
   316
      (\<Oplus>j \<in> {..k}. (\<Oplus>i \<in> {..j}. a i \<otimes> b (j-i)) \<otimes> c (n-j)) =
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
   317
      (\<Oplus>j \<in> {..k}. a j \<otimes> (\<Oplus>i \<in> {..k-j}. b i \<otimes> c (n-j-i)))"
19582
a669c98b9c24 get rid of 'concl is';
wenzelm
parents: 17094
diff changeset
   318
      (is "_ \<Longrightarrow> ?eq k")
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   319
    proof (induct k)
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   320
      case 0 then show ?case by (simp add: Pi_def m_assoc)
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   321
    next
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   322
      case (Suc k)
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   323
      then have "k <= n" by arith
23350
50c5b0912a0c tuned proofs: avoid implicit prems;
wenzelm
parents: 22931
diff changeset
   324
      from this R have "?eq k" by (rule Suc)
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   325
      with R show ?case
14666
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
   326
        by (simp cong: finsum_cong
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   327
             add: Suc_diff_le Pi_def l_distr r_distr m_assoc)
27717
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
   328
           (simp cong: finsum_cong add: Pi_def a_ac finsum_ldistr m_assoc)
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   329
    qed
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   330
  }
15095
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
   331
  with R show "coeff P ((p \<otimes>\<^bsub>P\<^esub> q) \<otimes>\<^bsub>P\<^esub> r) n = coeff P (p \<otimes>\<^bsub>P\<^esub> (q \<otimes>\<^bsub>P\<^esub> r)) n"
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   332
    by (simp add: Pi_def)
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   333
qed (simp_all add: R)
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   334
27717
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
   335
lemma UP_r_one [simp]:
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
   336
  assumes R: "p \<in> carrier P" shows "p \<otimes>\<^bsub>P\<^esub> \<one>\<^bsub>P\<^esub> = p"
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
   337
proof (rule up_eqI)
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
   338
  fix n
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
   339
  show "coeff P (p \<otimes>\<^bsub>P\<^esub> \<one>\<^bsub>P\<^esub>) n = coeff P p n"
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
   340
  proof (cases n)
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
   341
    case 0 
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
   342
    {
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
   343
      with R show ?thesis by simp
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
   344
    }
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
   345
  next
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
   346
    case Suc
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
   347
    {
27933
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
   348
      (*JE: in the locale UP_cring the proof was solved only with "by (simp del: finsum_Suc add: finsum_Suc2 Pi_def)", but I did not get it to work here*)
27717
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
   349
      fix nn assume Succ: "n = Suc nn"
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
   350
      have "coeff P (p \<otimes>\<^bsub>P\<^esub> \<one>\<^bsub>P\<^esub>) (Suc nn) = coeff P p (Suc nn)"
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
   351
      proof -
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
   352
        have "coeff P (p \<otimes>\<^bsub>P\<^esub> \<one>\<^bsub>P\<^esub>) (Suc nn) = (\<Oplus>i\<in>{..Suc nn}. coeff P p i \<otimes> (if Suc nn \<le> i then \<one> else \<zero>))" using R by simp
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
   353
        also have "\<dots> = coeff P p (Suc nn) \<otimes> (if Suc nn \<le> Suc nn then \<one> else \<zero>) \<oplus> (\<Oplus>i\<in>{..nn}. coeff P p i \<otimes> (if Suc nn \<le> i then \<one> else \<zero>))"
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
   354
          using finsum_Suc [of "(\<lambda>i::nat. coeff P p i \<otimes> (if Suc nn \<le> i then \<one> else \<zero>))" "nn"] unfolding Pi_def using R by simp
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
   355
        also have "\<dots> = coeff P p (Suc nn) \<otimes> (if Suc nn \<le> Suc nn then \<one> else \<zero>)"
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
   356
        proof -
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
   357
          have "(\<Oplus>i\<in>{..nn}. coeff P p i \<otimes> (if Suc nn \<le> i then \<one> else \<zero>)) = (\<Oplus>i\<in>{..nn}. \<zero>)"
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
   358
            using finsum_cong [of "{..nn}" "{..nn}" "(\<lambda>i::nat. coeff P p i \<otimes> (if Suc nn \<le> i then \<one> else \<zero>))" "(\<lambda>i::nat. \<zero>)"] using R 
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
   359
            unfolding Pi_def by simp
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
   360
          also have "\<dots> = \<zero>" by simp
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
   361
          finally show ?thesis using r_zero R by simp
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
   362
        qed
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
   363
        also have "\<dots> = coeff P p (Suc nn)" using R by simp
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
   364
        finally show ?thesis by simp
27717
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
   365
      qed
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
   366
      then show ?thesis using Succ by simp
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
   367
    }
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
   368
  qed
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
   369
qed (simp_all add: R)
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
   370
  
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
   371
lemma UP_l_one [simp]:
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   372
  assumes R: "p \<in> carrier P"
15095
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
   373
  shows "\<one>\<^bsub>P\<^esub> \<otimes>\<^bsub>P\<^esub> p = p"
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   374
proof (rule up_eqI)
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   375
  fix n
15095
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
   376
  show "coeff P (\<one>\<^bsub>P\<^esub> \<otimes>\<^bsub>P\<^esub> p) n = coeff P p n"
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   377
  proof (cases n)
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   378
    case 0 with R show ?thesis by simp
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   379
  next
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   380
    case Suc with R show ?thesis
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   381
      by (simp del: finsum_Suc add: finsum_Suc2 Pi_def)
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   382
  qed
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   383
qed (simp_all add: R)
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   384
27717
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
   385
lemma UP_l_distr:
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   386
  assumes R: "p \<in> carrier P" "q \<in> carrier P" "r \<in> carrier P"
15095
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
   387
  shows "(p \<oplus>\<^bsub>P\<^esub> q) \<otimes>\<^bsub>P\<^esub> r = (p \<otimes>\<^bsub>P\<^esub> r) \<oplus>\<^bsub>P\<^esub> (q \<otimes>\<^bsub>P\<^esub> r)"
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   388
  by (rule up_eqI) (simp add: l_distr R Pi_def, simp_all add: R)
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   389
27717
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
   390
lemma UP_r_distr:
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
   391
  assumes R: "p \<in> carrier P" "q \<in> carrier P" "r \<in> carrier P"
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
   392
  shows "r \<otimes>\<^bsub>P\<^esub> (p \<oplus>\<^bsub>P\<^esub> q) = (r \<otimes>\<^bsub>P\<^esub> p) \<oplus>\<^bsub>P\<^esub> (r \<otimes>\<^bsub>P\<^esub> q)"
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
   393
  by (rule up_eqI) (simp add: r_distr R Pi_def, simp_all add: R)
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
   394
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
   395
theorem UP_ring: "ring P"
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
   396
  by (auto intro!: ringI abelian_groupI monoidI UP_a_assoc)
27933
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
   397
    (auto intro: UP_a_comm UP_l_neg_ex UP_m_assoc UP_l_distr UP_r_distr)
27717
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
   398
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
   399
end
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
   400
27933
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
   401
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
   402
subsection {* Polynomials Form a Commutative Ring. *}
27717
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
   403
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
   404
context UP_cring
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
   405
begin
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
   406
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
   407
lemma UP_m_comm:
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
   408
  assumes R1: "p \<in> carrier P" and R2: "q \<in> carrier P" shows "p \<otimes>\<^bsub>P\<^esub> q = q \<otimes>\<^bsub>P\<^esub> p"
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   409
proof (rule up_eqI)
14666
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
   410
  fix n
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   411
  {
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   412
    fix k and a b :: "nat=>'a"
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   413
    assume R: "a \<in> UNIV -> carrier R" "b \<in> UNIV -> carrier R"
14666
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
   414
    then have "k <= n ==>
27717
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
   415
      (\<Oplus>i \<in> {..k}. a i \<otimes> b (n-i)) = (\<Oplus>i \<in> {..k}. a (k-i) \<otimes> b (i+n-k))"
19582
a669c98b9c24 get rid of 'concl is';
wenzelm
parents: 17094
diff changeset
   416
      (is "_ \<Longrightarrow> ?eq k")
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   417
    proof (induct k)
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   418
      case 0 then show ?case by (simp add: Pi_def)
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   419
    next
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   420
      case (Suc k) then show ?case
15944
9b00875e21f7 from simplesubst to new subst
paulson
parents: 15763
diff changeset
   421
        by (subst (2) finsum_Suc2) (simp add: Pi_def a_comm)+
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   422
    qed
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   423
  }
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   424
  note l = this
27717
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
   425
  from R1 R2 show "coeff P (p \<otimes>\<^bsub>P\<^esub> q) n =  coeff P (q \<otimes>\<^bsub>P\<^esub> p) n"
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
   426
    unfolding coeff_mult [OF R1 R2, of n] 
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
   427
    unfolding coeff_mult [OF R2 R1, of n] 
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
   428
    using l [of "(\<lambda>i. coeff P p i)" "(\<lambda>i. coeff P q i)" "n"] by (simp add: Pi_def m_comm)
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
   429
qed (simp_all add: R1 R2)
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   430
27717
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
   431
subsection{*Polynomials over a commutative ring for a commutative ring*}
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
   432
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
   433
theorem UP_cring:
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
   434
  "cring P" using UP_ring unfolding cring_def by (auto intro!: comm_monoidI UP_m_assoc UP_m_comm)
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   435
27717
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
   436
end
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
   437
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
   438
context UP_ring
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
   439
begin
14399
dc677b35e54f New lemmas about inversion of restricted functions.
ballarin
parents: 13975
diff changeset
   440
27717
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
   441
lemma UP_a_inv_closed [intro, simp]:
15095
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
   442
  "p \<in> carrier P ==> \<ominus>\<^bsub>P\<^esub> p \<in> carrier P"
27717
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
   443
  by (rule abelian_group.a_inv_closed [OF ring.is_abelian_group [OF UP_ring]])
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   444
27717
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
   445
lemma coeff_a_inv [simp]:
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   446
  assumes R: "p \<in> carrier P"
15095
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
   447
  shows "coeff P (\<ominus>\<^bsub>P\<^esub> p) n = \<ominus> (coeff P p n)"
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   448
proof -
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   449
  from R coeff_closed UP_a_inv_closed have
15095
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
   450
    "coeff P (\<ominus>\<^bsub>P\<^esub> p) n = \<ominus> coeff P p n \<oplus> (coeff P p n \<oplus> coeff P (\<ominus>\<^bsub>P\<^esub> p) n)"
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   451
    by algebra
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   452
  also from R have "... =  \<ominus> (coeff P p n)"
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   453
    by (simp del: coeff_add add: coeff_add [THEN sym]
14399
dc677b35e54f New lemmas about inversion of restricted functions.
ballarin
parents: 13975
diff changeset
   454
      abelian_group.r_neg [OF ring.is_abelian_group [OF UP_ring]])
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   455
  finally show ?thesis .
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   456
qed
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   457
27717
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
   458
end
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   459
29240
bb81c3709fb6 More porting to new locales.
ballarin
parents: 29237
diff changeset
   460
sublocale UP_ring < P: ring P using UP_ring .
bb81c3709fb6 More porting to new locales.
ballarin
parents: 29237
diff changeset
   461
sublocale UP_cring < P: cring P using UP_cring .
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   462
14666
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
   463
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents: 20282
diff changeset
   464
subsection {* Polynomials Form an Algebra *}
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   465
27717
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
   466
context UP_ring
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
   467
begin
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
   468
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
   469
lemma UP_smult_l_distr:
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   470
  "[| a \<in> carrier R; b \<in> carrier R; p \<in> carrier P |] ==>
15095
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
   471
  (a \<oplus> b) \<odot>\<^bsub>P\<^esub> p = a \<odot>\<^bsub>P\<^esub> p \<oplus>\<^bsub>P\<^esub> b \<odot>\<^bsub>P\<^esub> p"
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   472
  by (rule up_eqI) (simp_all add: R.l_distr)
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   473
27717
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
   474
lemma UP_smult_r_distr:
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   475
  "[| a \<in> carrier R; p \<in> carrier P; q \<in> carrier P |] ==>
15095
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
   476
  a \<odot>\<^bsub>P\<^esub> (p \<oplus>\<^bsub>P\<^esub> q) = a \<odot>\<^bsub>P\<^esub> p \<oplus>\<^bsub>P\<^esub> a \<odot>\<^bsub>P\<^esub> q"
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   477
  by (rule up_eqI) (simp_all add: R.r_distr)
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   478
27717
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
   479
lemma UP_smult_assoc1:
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   480
      "[| a \<in> carrier R; b \<in> carrier R; p \<in> carrier P |] ==>
15095
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
   481
      (a \<otimes> b) \<odot>\<^bsub>P\<^esub> p = a \<odot>\<^bsub>P\<^esub> (b \<odot>\<^bsub>P\<^esub> p)"
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   482
  by (rule up_eqI) (simp_all add: R.m_assoc)
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   483
27717
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
   484
lemma UP_smult_zero [simp]:
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
   485
      "p \<in> carrier P ==> \<zero> \<odot>\<^bsub>P\<^esub> p = \<zero>\<^bsub>P\<^esub>"
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
   486
  by (rule up_eqI) simp_all
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
   487
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
   488
lemma UP_smult_one [simp]:
15095
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
   489
      "p \<in> carrier P ==> \<one> \<odot>\<^bsub>P\<^esub> p = p"
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   490
  by (rule up_eqI) simp_all
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   491
27717
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
   492
lemma UP_smult_assoc2:
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   493
  "[| a \<in> carrier R; p \<in> carrier P; q \<in> carrier P |] ==>
15095
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
   494
  (a \<odot>\<^bsub>P\<^esub> p) \<otimes>\<^bsub>P\<^esub> q = a \<odot>\<^bsub>P\<^esub> (p \<otimes>\<^bsub>P\<^esub> q)"
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   495
  by (rule up_eqI) (simp_all add: R.finsum_rdistr R.m_assoc Pi_def)
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   496
27717
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
   497
end
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
   498
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   499
text {*
17094
7a3c2efecffe Use interpretation in locales.
ballarin
parents: 16639
diff changeset
   500
  Interpretation of lemmas from @{term algebra}.
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   501
*}
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   502
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   503
lemma (in cring) cring:
28823
dcbef866c9e2 tuned unfold_locales invocation
haftmann
parents: 27933
diff changeset
   504
  "cring R" ..
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   505
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   506
lemma (in UP_cring) UP_algebra:
27717
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
   507
  "algebra R P" by (auto intro!: algebraI R.cring UP_cring UP_smult_l_distr UP_smult_r_distr
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   508
    UP_smult_assoc1 UP_smult_assoc2)
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   509
29237
e90d9d51106b More porting to new locales.
ballarin
parents: 28823
diff changeset
   510
sublocale UP_cring < algebra R P using UP_algebra .
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   511
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   512
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents: 20282
diff changeset
   513
subsection {* Further Lemmas Involving Monomials *}
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   514
27717
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
   515
context UP_ring
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
   516
begin
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   517
27717
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
   518
lemma monom_zero [simp]:
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
   519
  "monom P \<zero> n = \<zero>\<^bsub>P\<^esub>" by (simp add: UP_def P_def)
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
   520
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
   521
lemma monom_mult_is_smult:
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   522
  assumes R: "a \<in> carrier R" "p \<in> carrier P"
15095
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
   523
  shows "monom P a 0 \<otimes>\<^bsub>P\<^esub> p = a \<odot>\<^bsub>P\<^esub> p"
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   524
proof (rule up_eqI)
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   525
  fix n
27717
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
   526
  show "coeff P (monom P a 0 \<otimes>\<^bsub>P\<^esub> p) n = coeff P (a \<odot>\<^bsub>P\<^esub> p) n"
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   527
  proof (cases n)
27717
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
   528
    case 0 with R show ?thesis by simp
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   529
  next
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   530
    case Suc with R show ?thesis
27717
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
   531
      using R.finsum_Suc2 by (simp del: R.finsum_Suc add: R.r_null Pi_def)
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   532
  qed
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   533
qed (simp_all add: R)
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   534
27717
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
   535
lemma monom_one [simp]:
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
   536
  "monom P \<one> 0 = \<one>\<^bsub>P\<^esub>"
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
   537
  by (rule up_eqI) simp_all
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
   538
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
   539
lemma monom_add [simp]:
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   540
  "[| a \<in> carrier R; b \<in> carrier R |] ==>
15095
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
   541
  monom P (a \<oplus> b) n = monom P a n \<oplus>\<^bsub>P\<^esub> monom P b n"
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   542
  by (rule up_eqI) simp_all
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   543
27717
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
   544
lemma monom_one_Suc:
15095
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
   545
  "monom P \<one> (Suc n) = monom P \<one> n \<otimes>\<^bsub>P\<^esub> monom P \<one> 1"
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   546
proof (rule up_eqI)
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   547
  fix k
15095
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
   548
  show "coeff P (monom P \<one> (Suc n)) k = coeff P (monom P \<one> n \<otimes>\<^bsub>P\<^esub> monom P \<one> 1) k"
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   549
  proof (cases "k = Suc n")
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   550
    case True show ?thesis
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   551
    proof -
26934
c1ae80a58341 avoid undeclared variables within proofs;
wenzelm
parents: 26202
diff changeset
   552
      fix m
14666
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
   553
      from True have less_add_diff:
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
   554
        "!!i. [| n < i; i <= n + m |] ==> n + m - i < m" by arith
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   555
      from True have "coeff P (monom P \<one> (Suc n)) k = \<one>" by simp
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   556
      also from True
15045
d59f7e2e18d3 Moved to new m<..<n syntax for set intervals.
nipkow
parents: 14963
diff changeset
   557
      have "... = (\<Oplus>i \<in> {..<n} \<union> {n}. coeff P (monom P \<one> n) i \<otimes>
14666
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
   558
        coeff P (monom P \<one> 1) (k - i))"
17094
7a3c2efecffe Use interpretation in locales.
ballarin
parents: 16639
diff changeset
   559
        by (simp cong: R.finsum_cong add: Pi_def)
14666
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
   560
      also have "... = (\<Oplus>i \<in>  {..n}. coeff P (monom P \<one> n) i \<otimes>
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
   561
        coeff P (monom P \<one> 1) (k - i))"
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
   562
        by (simp only: ivl_disj_un_singleton)
15095
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
   563
      also from True
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
   564
      have "... = (\<Oplus>i \<in> {..n} \<union> {n<..k}. coeff P (monom P \<one> n) i \<otimes>
14666
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
   565
        coeff P (monom P \<one> 1) (k - i))"
17094
7a3c2efecffe Use interpretation in locales.
ballarin
parents: 16639
diff changeset
   566
        by (simp cong: R.finsum_cong add: R.finsum_Un_disjoint ivl_disj_int_one
14666
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
   567
          order_less_imp_not_eq Pi_def)
15095
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
   568
      also from True have "... = coeff P (monom P \<one> n \<otimes>\<^bsub>P\<^esub> monom P \<one> 1) k"
14666
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
   569
        by (simp add: ivl_disj_un_one)
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   570
      finally show ?thesis .
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   571
    qed
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   572
  next
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   573
    case False
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   574
    note neq = False
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   575
    let ?s =
14666
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
   576
      "\<lambda>i. (if n = i then \<one> else \<zero>) \<otimes> (if Suc 0 = k - i then \<one> else \<zero>)"
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   577
    from neq have "coeff P (monom P \<one> (Suc n)) k = \<zero>" by simp
14666
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
   578
    also have "... = (\<Oplus>i \<in> {..k}. ?s i)"
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   579
    proof -
15095
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
   580
      have f1: "(\<Oplus>i \<in> {..<n}. ?s i) = \<zero>"
17094
7a3c2efecffe Use interpretation in locales.
ballarin
parents: 16639
diff changeset
   581
        by (simp cong: R.finsum_cong add: Pi_def)
14666
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
   582
      from neq have f2: "(\<Oplus>i \<in> {n}. ?s i) = \<zero>"
20432
07ec57376051 lin_arith_prover: splitting reverted because of performance loss
webertj
parents: 20318
diff changeset
   583
        by (simp cong: R.finsum_cong add: Pi_def) arith
15045
d59f7e2e18d3 Moved to new m<..<n syntax for set intervals.
nipkow
parents: 14963
diff changeset
   584
      have f3: "n < k ==> (\<Oplus>i \<in> {n<..k}. ?s i) = \<zero>"
17094
7a3c2efecffe Use interpretation in locales.
ballarin
parents: 16639
diff changeset
   585
        by (simp cong: R.finsum_cong add: order_less_imp_not_eq Pi_def)
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   586
      show ?thesis
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   587
      proof (cases "k < n")
17094
7a3c2efecffe Use interpretation in locales.
ballarin
parents: 16639
diff changeset
   588
        case True then show ?thesis by (simp cong: R.finsum_cong add: Pi_def)
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   589
      next
14666
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
   590
        case False then have n_le_k: "n <= k" by arith
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
   591
        show ?thesis
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
   592
        proof (cases "n = k")
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
   593
          case True
15045
d59f7e2e18d3 Moved to new m<..<n syntax for set intervals.
nipkow
parents: 14963
diff changeset
   594
          then have "\<zero> = (\<Oplus>i \<in> {..<n} \<union> {n}. ?s i)"
32456
341c83339aeb tuned the simp rules for Int involving insert and intervals.
nipkow
parents: 32436
diff changeset
   595
            by (simp cong: R.finsum_cong add: Pi_def)
14666
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
   596
          also from True have "... = (\<Oplus>i \<in> {..k}. ?s i)"
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
   597
            by (simp only: ivl_disj_un_singleton)
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
   598
          finally show ?thesis .
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
   599
        next
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
   600
          case False with n_le_k have n_less_k: "n < k" by arith
15045
d59f7e2e18d3 Moved to new m<..<n syntax for set intervals.
nipkow
parents: 14963
diff changeset
   601
          with neq have "\<zero> = (\<Oplus>i \<in> {..<n} \<union> {n}. ?s i)"
32456
341c83339aeb tuned the simp rules for Int involving insert and intervals.
nipkow
parents: 32436
diff changeset
   602
            by (simp add: R.finsum_Un_disjoint f1 f2 Pi_def del: Un_insert_right)
14666
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
   603
          also have "... = (\<Oplus>i \<in> {..n}. ?s i)"
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
   604
            by (simp only: ivl_disj_un_singleton)
15045
d59f7e2e18d3 Moved to new m<..<n syntax for set intervals.
nipkow
parents: 14963
diff changeset
   605
          also from n_less_k neq have "... = (\<Oplus>i \<in> {..n} \<union> {n<..k}. ?s i)"
17094
7a3c2efecffe Use interpretation in locales.
ballarin
parents: 16639
diff changeset
   606
            by (simp add: R.finsum_Un_disjoint f3 ivl_disj_int_one Pi_def)
14666
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
   607
          also from n_less_k have "... = (\<Oplus>i \<in> {..k}. ?s i)"
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
   608
            by (simp only: ivl_disj_un_one)
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
   609
          finally show ?thesis .
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
   610
        qed
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   611
      qed
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   612
    qed
15095
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
   613
    also have "... = coeff P (monom P \<one> n \<otimes>\<^bsub>P\<^esub> monom P \<one> 1) k" by simp
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   614
    finally show ?thesis .
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   615
  qed
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   616
qed (simp_all)
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   617
27717
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
   618
lemma monom_one_Suc2:
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
   619
  "monom P \<one> (Suc n) = monom P \<one> 1 \<otimes>\<^bsub>P\<^esub> monom P \<one> n"
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
   620
proof (induct n)
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
   621
  case 0 show ?case by simp
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
   622
next
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
   623
  case Suc
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
   624
  {
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
   625
    fix k:: nat
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
   626
    assume hypo: "monom P \<one> (Suc k) = monom P \<one> 1 \<otimes>\<^bsub>P\<^esub> monom P \<one> k"
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
   627
    then show "monom P \<one> (Suc (Suc k)) = monom P \<one> 1 \<otimes>\<^bsub>P\<^esub> monom P \<one> (Suc k)"
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
   628
    proof -
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
   629
      have lhs: "monom P \<one> (Suc (Suc k)) = monom P \<one> 1 \<otimes>\<^bsub>P\<^esub> monom P \<one> k \<otimes>\<^bsub>P\<^esub> monom P \<one> 1"
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
   630
        unfolding monom_one_Suc [of "Suc k"] unfolding hypo ..
27717
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
   631
      note cl = monom_closed [OF R.one_closed, of 1]
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
   632
      note clk = monom_closed [OF R.one_closed, of k]
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
   633
      have rhs: "monom P \<one> 1 \<otimes>\<^bsub>P\<^esub> monom P \<one> (Suc k) = monom P \<one> 1 \<otimes>\<^bsub>P\<^esub> monom P \<one> k \<otimes>\<^bsub>P\<^esub> monom P \<one> 1"
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
   634
        unfolding monom_one_Suc [of k] unfolding sym [OF m_assoc  [OF cl clk cl]] ..
27717
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
   635
      from lhs rhs show ?thesis by simp
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
   636
    qed
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
   637
  }
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
   638
qed
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
   639
30363
9b8d9b6ef803 proper local context for text with antiquotations;
wenzelm
parents: 29246
diff changeset
   640
text{*The following corollary follows from lemmas @{thm "monom_one_Suc"} 
9b8d9b6ef803 proper local context for text with antiquotations;
wenzelm
parents: 29246
diff changeset
   641
  and @{thm "monom_one_Suc2"}, and is trivial in @{term UP_cring}*}
27717
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
   642
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
   643
corollary monom_one_comm: shows "monom P \<one> k \<otimes>\<^bsub>P\<^esub> monom P \<one> 1 = monom P \<one> 1 \<otimes>\<^bsub>P\<^esub> monom P \<one> k"
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
   644
  unfolding monom_one_Suc [symmetric] monom_one_Suc2 [symmetric] ..
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
   645
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
   646
lemma monom_mult_smult:
15095
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
   647
  "[| a \<in> carrier R; b \<in> carrier R |] ==> monom P (a \<otimes> b) n = a \<odot>\<^bsub>P\<^esub> monom P b n"
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   648
  by (rule up_eqI) simp_all
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   649
27717
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
   650
lemma monom_one_mult:
15095
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
   651
  "monom P \<one> (n + m) = monom P \<one> n \<otimes>\<^bsub>P\<^esub> monom P \<one> m"
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   652
proof (induct n)
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   653
  case 0 show ?case by simp
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   654
next
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   655
  case Suc then show ?case
27717
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
   656
    unfolding add_Suc unfolding monom_one_Suc unfolding Suc.hyps
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
   657
    using m_assoc monom_one_comm [of m] by simp
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   658
qed
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   659
27717
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
   660
lemma monom_one_mult_comm: "monom P \<one> n \<otimes>\<^bsub>P\<^esub> monom P \<one> m = monom P \<one> m \<otimes>\<^bsub>P\<^esub> monom P \<one> n"
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
   661
  unfolding monom_one_mult [symmetric] by (rule up_eqI) simp_all
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
   662
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
   663
lemma monom_mult [simp]:
27933
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
   664
  assumes a_in_R: "a \<in> carrier R" and b_in_R: "b \<in> carrier R"
15095
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
   665
  shows "monom P (a \<otimes> b) (n + m) = monom P a n \<otimes>\<^bsub>P\<^esub> monom P b m"
27933
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
   666
proof (rule up_eqI)
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
   667
  fix k 
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
   668
  show "coeff P (monom P (a \<otimes> b) (n + m)) k = coeff P (monom P a n \<otimes>\<^bsub>P\<^esub> monom P b m) k"
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
   669
  proof (cases "n + m = k")
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
   670
    case True 
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
   671
    {
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
   672
      show ?thesis
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
   673
        unfolding True [symmetric]
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
   674
          coeff_mult [OF monom_closed [OF a_in_R, of n] monom_closed [OF b_in_R, of m], of "n + m"] 
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
   675
          coeff_monom [OF a_in_R, of n] coeff_monom [OF b_in_R, of m]
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
   676
        using R.finsum_cong [of "{.. n + m}" "{.. n + m}" "(\<lambda>i. (if n = i then a else \<zero>) \<otimes> (if m = n + m - i then b else \<zero>))" 
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
   677
          "(\<lambda>i. if n = i then a \<otimes> b else \<zero>)"]
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
   678
          a_in_R b_in_R
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
   679
        unfolding simp_implies_def
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
   680
        using R.finsum_singleton [of n "{.. n + m}" "(\<lambda>i. a \<otimes> b)"]
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
   681
        unfolding Pi_def by auto
27933
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
   682
    }
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
   683
  next
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
   684
    case False
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
   685
    {
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
   686
      show ?thesis
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
   687
        unfolding coeff_monom [OF R.m_closed [OF a_in_R b_in_R], of "n + m" k] apply (simp add: False)
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
   688
        unfolding coeff_mult [OF monom_closed [OF a_in_R, of n] monom_closed [OF b_in_R, of m], of k]
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
   689
        unfolding coeff_monom [OF a_in_R, of n] unfolding coeff_monom [OF b_in_R, of m] using False
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
   690
        using R.finsum_cong [of "{..k}" "{..k}" "(\<lambda>i. (if n = i then a else \<zero>) \<otimes> (if m = k - i then b else \<zero>))" "(\<lambda>i. \<zero>)"]
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
   691
        unfolding Pi_def simp_implies_def using a_in_R b_in_R by force
27933
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
   692
    }
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
   693
  qed
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
   694
qed (simp_all add: a_in_R b_in_R)
27717
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
   695
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
   696
lemma monom_a_inv [simp]:
15095
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
   697
  "a \<in> carrier R ==> monom P (\<ominus> a) n = \<ominus>\<^bsub>P\<^esub> monom P a n"
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   698
  by (rule up_eqI) simp_all
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   699
27717
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
   700
lemma monom_inj:
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   701
  "inj_on (%a. monom P a n) (carrier R)"
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   702
proof (rule inj_onI)
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   703
  fix x y
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   704
  assume R: "x \<in> carrier R" "y \<in> carrier R" and eq: "monom P x n = monom P y n"
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   705
  then have "coeff P (monom P x n) n = coeff P (monom P y n) n" by simp
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   706
  with R show "x = y" by simp
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   707
qed
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   708
27717
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
   709
end
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
   710
17094
7a3c2efecffe Use interpretation in locales.
ballarin
parents: 16639
diff changeset
   711
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents: 20282
diff changeset
   712
subsection {* The Degree Function *}
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   713
27933
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
   714
definition deg :: "[('a, 'm) ring_scheme, nat => 'a] => nat"
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
   715
  where "deg R p == LEAST n. bound \<zero>\<^bsub>R\<^esub> n (coeff (UP R) p)"
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   716
27717
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
   717
context UP_ring
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
   718
begin
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
   719
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
   720
lemma deg_aboveI:
14666
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
   721
  "[| (!!m. n < m ==> coeff P p m = \<zero>); p \<in> carrier P |] ==> deg R p <= n"
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   722
  by (unfold deg_def P_def) (fast intro: Least_le)
15095
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
   723
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   724
(*
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   725
lemma coeff_bound_ex: "EX n. bound n (coeff p)"
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   726
proof -
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   727
  have "(%n. coeff p n) : UP" by (simp add: coeff_def Rep_UP)
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   728
  then obtain n where "bound n (coeff p)" by (unfold UP_def) fast
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   729
  then show ?thesis ..
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   730
qed
14666
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
   731
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   732
lemma bound_coeff_obtain:
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   733
  assumes prem: "(!!n. bound n (coeff p) ==> P)" shows "P"
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   734
proof -
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   735
  have "(%n. coeff p n) : UP" by (simp add: coeff_def Rep_UP)
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   736
  then obtain n where "bound n (coeff p)" by (unfold UP_def) fast
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   737
  with prem show P .
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   738
qed
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   739
*)
15095
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
   740
27717
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
   741
lemma deg_aboveD:
23350
50c5b0912a0c tuned proofs: avoid implicit prems;
wenzelm
parents: 22931
diff changeset
   742
  assumes "deg R p < m" and "p \<in> carrier P"
50c5b0912a0c tuned proofs: avoid implicit prems;
wenzelm
parents: 22931
diff changeset
   743
  shows "coeff P p m = \<zero>"
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   744
proof -
23350
50c5b0912a0c tuned proofs: avoid implicit prems;
wenzelm
parents: 22931
diff changeset
   745
  from `p \<in> carrier P` obtain n where "bound \<zero> n (coeff P p)"
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   746
    by (auto simp add: UP_def P_def)
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   747
  then have "bound \<zero> (deg R p) (coeff P p)"
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   748
    by (auto simp: deg_def P_def dest: LeastI)
23350
50c5b0912a0c tuned proofs: avoid implicit prems;
wenzelm
parents: 22931
diff changeset
   749
  from this and `deg R p < m` show ?thesis ..
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   750
qed
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   751
27717
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
   752
lemma deg_belowI:
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   753
  assumes non_zero: "n ~= 0 ==> coeff P p n ~= \<zero>"
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   754
    and R: "p \<in> carrier P"
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   755
  shows "n <= deg R p"
14666
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
   756
-- {* Logically, this is a slightly stronger version of
15095
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
   757
   @{thm [source] deg_aboveD} *}
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   758
proof (cases "n=0")
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   759
  case True then show ?thesis by simp
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   760
next
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   761
  case False then have "coeff P p n ~= \<zero>" by (rule non_zero)
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   762
  then have "~ deg R p < n" by (fast dest: deg_aboveD intro: R)
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   763
  then show ?thesis by arith
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   764
qed
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   765
27717
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
   766
lemma lcoeff_nonzero_deg:
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   767
  assumes deg: "deg R p ~= 0" and R: "p \<in> carrier P"
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   768
  shows "coeff P p (deg R p) ~= \<zero>"
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   769
proof -
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   770
  from R obtain m where "deg R p <= m" and m_coeff: "coeff P p m ~= \<zero>"
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   771
  proof -
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   772
    have minus: "!!(n::nat) m. n ~= 0 ==> (n - Suc 0 < m) = (n <= m)"
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   773
      by arith
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   774
    from deg have "deg R p - 1 < (LEAST n. bound \<zero> n (coeff P p))"
27717
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
   775
      by (unfold deg_def P_def) simp
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   776
    then have "~ bound \<zero> (deg R p - 1) (coeff P p)" by (rule not_less_Least)
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   777
    then have "EX m. deg R p - 1 < m & coeff P p m ~= \<zero>"
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   778
      by (unfold bound_def) fast
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   779
    then have "EX m. deg R p <= m & coeff P p m ~= \<zero>" by (simp add: deg minus)
23350
50c5b0912a0c tuned proofs: avoid implicit prems;
wenzelm
parents: 22931
diff changeset
   780
    then show ?thesis by (auto intro: that)
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   781
  qed
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   782
  with deg_belowI R have "deg R p = m" by fastsimp
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   783
  with m_coeff show ?thesis by simp
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   784
qed
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   785
27717
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
   786
lemma lcoeff_nonzero_nonzero:
15095
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
   787
  assumes deg: "deg R p = 0" and nonzero: "p ~= \<zero>\<^bsub>P\<^esub>" and R: "p \<in> carrier P"
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   788
  shows "coeff P p 0 ~= \<zero>"
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   789
proof -
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   790
  have "EX m. coeff P p m ~= \<zero>"
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   791
  proof (rule classical)
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   792
    assume "~ ?thesis"
15095
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
   793
    with R have "p = \<zero>\<^bsub>P\<^esub>" by (auto intro: up_eqI)
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   794
    with nonzero show ?thesis by contradiction
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   795
  qed
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   796
  then obtain m where coeff: "coeff P p m ~= \<zero>" ..
23350
50c5b0912a0c tuned proofs: avoid implicit prems;
wenzelm
parents: 22931
diff changeset
   797
  from this and R have "m <= deg R p" by (rule deg_belowI)
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   798
  then have "m = 0" by (simp add: deg)
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   799
  with coeff show ?thesis by simp
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   800
qed
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   801
27717
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
   802
lemma lcoeff_nonzero:
15095
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
   803
  assumes neq: "p ~= \<zero>\<^bsub>P\<^esub>" and R: "p \<in> carrier P"
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   804
  shows "coeff P p (deg R p) ~= \<zero>"
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   805
proof (cases "deg R p = 0")
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   806
  case True with neq R show ?thesis by (simp add: lcoeff_nonzero_nonzero)
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   807
next
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   808
  case False with neq R show ?thesis by (simp add: lcoeff_nonzero_deg)
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   809
qed
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   810
27717
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
   811
lemma deg_eqI:
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   812
  "[| !!m. n < m ==> coeff P p m = \<zero>;
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   813
      !!n. n ~= 0 ==> coeff P p n ~= \<zero>; p \<in> carrier P |] ==> deg R p = n"
33657
a4179bf442d1 renamed lemmas "anti_sym" -> "antisym"
nipkow
parents: 32960
diff changeset
   814
by (fast intro: le_antisym deg_aboveI deg_belowI)
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   815
17094
7a3c2efecffe Use interpretation in locales.
ballarin
parents: 16639
diff changeset
   816
text {* Degree and polynomial operations *}
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   817
27717
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
   818
lemma deg_add [simp]:
32436
10cd49e0c067 Turned "x <= y ==> sup x y = y" (and relatives) into simp rules
nipkow
parents: 30729
diff changeset
   819
  "p \<in> carrier P \<Longrightarrow> q \<in> carrier P \<Longrightarrow>
10cd49e0c067 Turned "x <= y ==> sup x y = y" (and relatives) into simp rules
nipkow
parents: 30729
diff changeset
   820
  deg R (p \<oplus>\<^bsub>P\<^esub> q) <= max (deg R p) (deg R q)"
10cd49e0c067 Turned "x <= y ==> sup x y = y" (and relatives) into simp rules
nipkow
parents: 30729
diff changeset
   821
by(rule deg_aboveI)(simp_all add: deg_aboveD)
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   822
27717
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
   823
lemma deg_monom_le:
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   824
  "a \<in> carrier R ==> deg R (monom P a n) <= n"
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   825
  by (intro deg_aboveI) simp_all
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   826
27717
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
   827
lemma deg_monom [simp]:
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   828
  "[| a ~= \<zero>; a \<in> carrier R |] ==> deg R (monom P a n) = n"
33657
a4179bf442d1 renamed lemmas "anti_sym" -> "antisym"
nipkow
parents: 32960
diff changeset
   829
  by (fastsimp intro: le_antisym deg_aboveI deg_belowI)
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   830
27717
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
   831
lemma deg_const [simp]:
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   832
  assumes R: "a \<in> carrier R" shows "deg R (monom P a 0) = 0"
33657
a4179bf442d1 renamed lemmas "anti_sym" -> "antisym"
nipkow
parents: 32960
diff changeset
   833
proof (rule le_antisym)
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   834
  show "deg R (monom P a 0) <= 0" by (rule deg_aboveI) (simp_all add: R)
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   835
next
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   836
  show "0 <= deg R (monom P a 0)" by (rule deg_belowI) (simp_all add: R)
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   837
qed
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   838
27717
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
   839
lemma deg_zero [simp]:
15095
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
   840
  "deg R \<zero>\<^bsub>P\<^esub> = 0"
33657
a4179bf442d1 renamed lemmas "anti_sym" -> "antisym"
nipkow
parents: 32960
diff changeset
   841
proof (rule le_antisym)
15095
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
   842
  show "deg R \<zero>\<^bsub>P\<^esub> <= 0" by (rule deg_aboveI) simp_all
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   843
next
15095
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
   844
  show "0 <= deg R \<zero>\<^bsub>P\<^esub>" by (rule deg_belowI) simp_all
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   845
qed
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   846
27717
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
   847
lemma deg_one [simp]:
15095
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
   848
  "deg R \<one>\<^bsub>P\<^esub> = 0"
33657
a4179bf442d1 renamed lemmas "anti_sym" -> "antisym"
nipkow
parents: 32960
diff changeset
   849
proof (rule le_antisym)
15095
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
   850
  show "deg R \<one>\<^bsub>P\<^esub> <= 0" by (rule deg_aboveI) simp_all
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   851
next
15095
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
   852
  show "0 <= deg R \<one>\<^bsub>P\<^esub>" by (rule deg_belowI) simp_all
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   853
qed
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   854
27717
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
   855
lemma deg_uminus [simp]:
15095
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
   856
  assumes R: "p \<in> carrier P" shows "deg R (\<ominus>\<^bsub>P\<^esub> p) = deg R p"
33657
a4179bf442d1 renamed lemmas "anti_sym" -> "antisym"
nipkow
parents: 32960
diff changeset
   857
proof (rule le_antisym)
15095
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
   858
  show "deg R (\<ominus>\<^bsub>P\<^esub> p) <= deg R p" by (simp add: deg_aboveI deg_aboveD R)
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   859
next
15095
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
   860
  show "deg R p <= deg R (\<ominus>\<^bsub>P\<^esub> p)"
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   861
    by (simp add: deg_belowI lcoeff_nonzero_deg
17094
7a3c2efecffe Use interpretation in locales.
ballarin
parents: 16639
diff changeset
   862
      inj_on_iff [OF R.a_inv_inj, of _ "\<zero>", simplified] R)
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   863
qed
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   864
27717
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
   865
text{*The following lemma is later \emph{overwritten} by the most
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
   866
  specific one for domains, @{text deg_smult}.*}
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
   867
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
   868
lemma deg_smult_ring [simp]:
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   869
  "[| a \<in> carrier R; p \<in> carrier P |] ==>
15095
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
   870
  deg R (a \<odot>\<^bsub>P\<^esub> p) <= (if a = \<zero> then 0 else deg R p)"
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   871
  by (cases "a = \<zero>") (simp add: deg_aboveI deg_aboveD)+
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   872
27717
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
   873
end
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
   874
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
   875
context UP_domain
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
   876
begin
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
   877
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
   878
lemma deg_smult [simp]:
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   879
  assumes R: "a \<in> carrier R" "p \<in> carrier P"
15095
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
   880
  shows "deg R (a \<odot>\<^bsub>P\<^esub> p) = (if a = \<zero> then 0 else deg R p)"
33657
a4179bf442d1 renamed lemmas "anti_sym" -> "antisym"
nipkow
parents: 32960
diff changeset
   881
proof (rule le_antisym)
15095
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
   882
  show "deg R (a \<odot>\<^bsub>P\<^esub> p) <= (if a = \<zero> then 0 else deg R p)"
23350
50c5b0912a0c tuned proofs: avoid implicit prems;
wenzelm
parents: 22931
diff changeset
   883
    using R by (rule deg_smult_ring)
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   884
next
15095
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
   885
  show "(if a = \<zero> then 0 else deg R p) <= deg R (a \<odot>\<^bsub>P\<^esub> p)"
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   886
  proof (cases "a = \<zero>")
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   887
  qed (simp, simp add: deg_belowI lcoeff_nonzero_deg integral_iff R)
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   888
qed
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   889
27717
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
   890
end
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
   891
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
   892
context UP_ring
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
   893
begin
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
   894
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
   895
lemma deg_mult_ring:
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   896
  assumes R: "p \<in> carrier P" "q \<in> carrier P"
15095
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
   897
  shows "deg R (p \<otimes>\<^bsub>P\<^esub> q) <= deg R p + deg R q"
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   898
proof (rule deg_aboveI)
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   899
  fix m
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   900
  assume boundm: "deg R p + deg R q < m"
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   901
  {
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   902
    fix k i
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   903
    assume boundk: "deg R p + deg R q < k"
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   904
    then have "coeff P p i \<otimes> coeff P q (k - i) = \<zero>"
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   905
    proof (cases "deg R p < i")
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   906
      case True then show ?thesis by (simp add: deg_aboveD R)
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   907
    next
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   908
      case False with boundk have "deg R q < k - i" by arith
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   909
      then show ?thesis by (simp add: deg_aboveD R)
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   910
    qed
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   911
  }
15095
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
   912
  with boundm R show "coeff P (p \<otimes>\<^bsub>P\<^esub> q) m = \<zero>" by simp
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   913
qed (simp add: R)
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   914
27717
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
   915
end
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
   916
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
   917
context UP_domain
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
   918
begin
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
   919
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
   920
lemma deg_mult [simp]:
15095
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
   921
  "[| p ~= \<zero>\<^bsub>P\<^esub>; q ~= \<zero>\<^bsub>P\<^esub>; p \<in> carrier P; q \<in> carrier P |] ==>
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
   922
  deg R (p \<otimes>\<^bsub>P\<^esub> q) = deg R p + deg R q"
33657
a4179bf442d1 renamed lemmas "anti_sym" -> "antisym"
nipkow
parents: 32960
diff changeset
   923
proof (rule le_antisym)
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   924
  assume "p \<in> carrier P" " q \<in> carrier P"
27717
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
   925
  then show "deg R (p \<otimes>\<^bsub>P\<^esub> q) <= deg R p + deg R q" by (rule deg_mult_ring)
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   926
next
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   927
  let ?s = "(%i. coeff P p i \<otimes> coeff P q (deg R p + deg R q - i))"
15095
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
   928
  assume R: "p \<in> carrier P" "q \<in> carrier P" and nz: "p ~= \<zero>\<^bsub>P\<^esub>" "q ~= \<zero>\<^bsub>P\<^esub>"
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   929
  have less_add_diff: "!!(k::nat) n m. k < n ==> m < n + m - k" by arith
15095
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
   930
  show "deg R p + deg R q <= deg R (p \<otimes>\<^bsub>P\<^esub> q)"
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   931
  proof (rule deg_belowI, simp add: R)
15095
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
   932
    have "(\<Oplus>i \<in> {.. deg R p + deg R q}. ?s i)
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
   933
      = (\<Oplus>i \<in> {..< deg R p} \<union> {deg R p .. deg R p + deg R q}. ?s i)"
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   934
      by (simp only: ivl_disj_un_one)
15095
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
   935
    also have "... = (\<Oplus>i \<in> {deg R p .. deg R p + deg R q}. ?s i)"
17094
7a3c2efecffe Use interpretation in locales.
ballarin
parents: 16639
diff changeset
   936
      by (simp cong: R.finsum_cong add: R.finsum_Un_disjoint ivl_disj_int_one
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   937
        deg_aboveD less_add_diff R Pi_def)
15095
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
   938
    also have "...= (\<Oplus>i \<in> {deg R p} \<union> {deg R p <.. deg R p + deg R q}. ?s i)"
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   939
      by (simp only: ivl_disj_un_singleton)
14666
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
   940
    also have "... = coeff P p (deg R p) \<otimes> coeff P q (deg R q)"
32456
341c83339aeb tuned the simp rules for Int involving insert and intervals.
nipkow
parents: 32436
diff changeset
   941
      by (simp cong: R.finsum_cong add: deg_aboveD R Pi_def)
15095
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
   942
    finally have "(\<Oplus>i \<in> {.. deg R p + deg R q}. ?s i)
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   943
      = coeff P p (deg R p) \<otimes> coeff P q (deg R q)" .
15095
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
   944
    with nz show "(\<Oplus>i \<in> {.. deg R p + deg R q}. ?s i) ~= \<zero>"
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   945
      by (simp add: integral_iff lcoeff_nonzero R)
27717
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
   946
  qed (simp add: R)
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
   947
qed
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
   948
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
   949
end
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   950
27717
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
   951
text{*The following lemmas also can be lifted to @{term UP_ring}.*}
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
   952
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
   953
context UP_ring
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
   954
begin
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
   955
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
   956
lemma coeff_finsum:
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   957
  assumes fin: "finite A"
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   958
  shows "p \<in> A -> carrier P ==>
15095
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
   959
    coeff P (finsum P p A) k = (\<Oplus>i \<in> A. coeff P (p i) k)"
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   960
  using fin by induct (auto simp: Pi_def)
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   961
27717
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
   962
lemma up_repr:
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   963
  assumes R: "p \<in> carrier P"
15095
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
   964
  shows "(\<Oplus>\<^bsub>P\<^esub> i \<in> {..deg R p}. monom P (coeff P p i) i) = p"
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   965
proof (rule up_eqI)
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   966
  let ?s = "(%i. monom P (coeff P p i) i)"
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   967
  fix k
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   968
  from R have RR: "!!i. (if i = k then coeff P p i else \<zero>) \<in> carrier R"
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   969
    by simp
15095
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
   970
  show "coeff P (\<Oplus>\<^bsub>P\<^esub> i \<in> {..deg R p}. ?s i) k = coeff P p k"
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   971
  proof (cases "k <= deg R p")
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   972
    case True
15095
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
   973
    hence "coeff P (\<Oplus>\<^bsub>P\<^esub> i \<in> {..deg R p}. ?s i) k =
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
   974
          coeff P (\<Oplus>\<^bsub>P\<^esub> i \<in> {..k} \<union> {k<..deg R p}. ?s i) k"
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   975
      by (simp only: ivl_disj_un_one)
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   976
    also from True
15095
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
   977
    have "... = coeff P (\<Oplus>\<^bsub>P\<^esub> i \<in> {..k}. ?s i) k"
17094
7a3c2efecffe Use interpretation in locales.
ballarin
parents: 16639
diff changeset
   978
      by (simp cong: R.finsum_cong add: R.finsum_Un_disjoint
14666
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
   979
        ivl_disj_int_one order_less_imp_not_eq2 coeff_finsum R RR Pi_def)
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   980
    also
15095
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
   981
    have "... = coeff P (\<Oplus>\<^bsub>P\<^esub> i \<in> {..<k} \<union> {k}. ?s i) k"
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   982
      by (simp only: ivl_disj_un_singleton)
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   983
    also have "... = coeff P p k"
32456
341c83339aeb tuned the simp rules for Int involving insert and intervals.
nipkow
parents: 32436
diff changeset
   984
      by (simp cong: R.finsum_cong add: coeff_finsum deg_aboveD R RR Pi_def)
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   985
    finally show ?thesis .
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   986
  next
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   987
    case False
15095
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
   988
    hence "coeff P (\<Oplus>\<^bsub>P\<^esub> i \<in> {..deg R p}. ?s i) k =
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
   989
          coeff P (\<Oplus>\<^bsub>P\<^esub> i \<in> {..<deg R p} \<union> {deg R p}. ?s i) k"
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   990
      by (simp only: ivl_disj_un_singleton)
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   991
    also from False have "... = coeff P p k"
32456
341c83339aeb tuned the simp rules for Int involving insert and intervals.
nipkow
parents: 32436
diff changeset
   992
      by (simp cong: R.finsum_cong add: coeff_finsum deg_aboveD R Pi_def)
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   993
    finally show ?thesis .
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   994
  qed
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   995
qed (simp_all add: R Pi_def)
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   996
27717
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
   997
lemma up_repr_le:
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   998
  "[| deg R p <= n; p \<in> carrier P |] ==>
15095
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
   999
  (\<Oplus>\<^bsub>P\<^esub> i \<in> {..n}. monom P (coeff P p i) i) = p"
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1000
proof -
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1001
  let ?s = "(%i. monom P (coeff P p i) i)"
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1002
  assume R: "p \<in> carrier P" and "deg R p <= n"
15095
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
  1003
  then have "finsum P ?s {..n} = finsum P ?s ({..deg R p} \<union> {deg R p<..n})"
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1004
    by (simp only: ivl_disj_un_one)
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1005
  also have "... = finsum P ?s {..deg R p}"
17094
7a3c2efecffe Use interpretation in locales.
ballarin
parents: 16639
diff changeset
  1006
    by (simp cong: P.finsum_cong add: P.finsum_Un_disjoint ivl_disj_int_one
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1007
      deg_aboveD R Pi_def)
23350
50c5b0912a0c tuned proofs: avoid implicit prems;
wenzelm
parents: 22931
diff changeset
  1008
  also have "... = p" using R by (rule up_repr)
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1009
  finally show ?thesis .
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1010
qed
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1011
27717
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
  1012
end
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
  1013
17094
7a3c2efecffe Use interpretation in locales.
ballarin
parents: 16639
diff changeset
  1014
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents: 20282
diff changeset
  1015
subsection {* Polynomials over Integral Domains *}
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1016
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1017
lemma domainI:
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1018
  assumes cring: "cring R"
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1019
    and one_not_zero: "one R ~= zero R"
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1020
    and integral: "!!a b. [| mult R a b = zero R; a \<in> carrier R;
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1021
      b \<in> carrier R |] ==> a = zero R | b = zero R"
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1022
  shows "domain R"
27714
27b4d7c01f8b Tuned (for the sake of a meaningless log entry).
ballarin
parents: 27611
diff changeset
  1023
  by (auto intro!: domain.intro domain_axioms.intro cring.axioms assms
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1024
    del: disjCI)
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1025
27717
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
  1026
context UP_domain
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
  1027
begin
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
  1028
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
  1029
lemma UP_one_not_zero:
15095
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
  1030
  "\<one>\<^bsub>P\<^esub> ~= \<zero>\<^bsub>P\<^esub>"
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1031
proof
15095
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
  1032
  assume "\<one>\<^bsub>P\<^esub> = \<zero>\<^bsub>P\<^esub>"
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
  1033
  hence "coeff P \<one>\<^bsub>P\<^esub> 0 = (coeff P \<zero>\<^bsub>P\<^esub> 0)" by simp
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1034
  hence "\<one> = \<zero>" by simp
27717
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
  1035
  with R.one_not_zero show "False" by contradiction
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1036
qed
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1037
27717
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
  1038
lemma UP_integral:
15095
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
  1039
  "[| p \<otimes>\<^bsub>P\<^esub> q = \<zero>\<^bsub>P\<^esub>; p \<in> carrier P; q \<in> carrier P |] ==> p = \<zero>\<^bsub>P\<^esub> | q = \<zero>\<^bsub>P\<^esub>"
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1040
proof -
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1041
  fix p q
15095
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
  1042
  assume pq: "p \<otimes>\<^bsub>P\<^esub> q = \<zero>\<^bsub>P\<^esub>" and R: "p \<in> carrier P" "q \<in> carrier P"
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
  1043
  show "p = \<zero>\<^bsub>P\<^esub> | q = \<zero>\<^bsub>P\<^esub>"
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1044
  proof (rule classical)
15095
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
  1045
    assume c: "~ (p = \<zero>\<^bsub>P\<^esub> | q = \<zero>\<^bsub>P\<^esub>)"
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
  1046
    with R have "deg R p + deg R q = deg R (p \<otimes>\<^bsub>P\<^esub> q)" by simp
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1047
    also from pq have "... = 0" by simp
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1048
    finally have "deg R p + deg R q = 0" .
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1049
    then have f1: "deg R p = 0 & deg R q = 0" by simp
15095
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
  1050
    from f1 R have "p = (\<Oplus>\<^bsub>P\<^esub> i \<in> {..0}. monom P (coeff P p i) i)"
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1051
      by (simp only: up_repr_le)
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1052
    also from R have "... = monom P (coeff P p 0) 0" by simp
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1053
    finally have p: "p = monom P (coeff P p 0) 0" .
15095
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
  1054
    from f1 R have "q = (\<Oplus>\<^bsub>P\<^esub> i \<in> {..0}. monom P (coeff P q i) i)"
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1055
      by (simp only: up_repr_le)
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1056
    also from R have "... = monom P (coeff P q 0) 0" by simp
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1057
    finally have q: "q = monom P (coeff P q 0) 0" .
15095
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
  1058
    from R have "coeff P p 0 \<otimes> coeff P q 0 = coeff P (p \<otimes>\<^bsub>P\<^esub> q) 0" by simp
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1059
    also from pq have "... = \<zero>" by simp
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1060
    finally have "coeff P p 0 \<otimes> coeff P q 0 = \<zero>" .
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1061
    with R have "coeff P p 0 = \<zero> | coeff P q 0 = \<zero>"
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1062
      by (simp add: R.integral_iff)
15095
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
  1063
    with p q show "p = \<zero>\<^bsub>P\<^esub> | q = \<zero>\<^bsub>P\<^esub>" by fastsimp
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1064
  qed
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1065
qed
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1066
27717
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
  1067
theorem UP_domain:
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1068
  "domain P"
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1069
  by (auto intro!: domainI UP_cring UP_one_not_zero UP_integral del: disjCI)
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1070
27717
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
  1071
end
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
  1072
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1073
text {*
17094
7a3c2efecffe Use interpretation in locales.
ballarin
parents: 16639
diff changeset
  1074
  Interpretation of theorems from @{term domain}.
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1075
*}
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1076
29237
e90d9d51106b More porting to new locales.
ballarin
parents: 28823
diff changeset
  1077
sublocale UP_domain < "domain" P
19984
29bb4659f80a Method intro_locales replaced by intro_locales and unfold_locales.
ballarin
parents: 19931
diff changeset
  1078
  by intro_locales (rule domain.axioms UP_domain)+
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1079
14666
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
  1080
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents: 20282
diff changeset
  1081
subsection {* The Evaluation Homomorphism and Universal Property*}
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1082
14666
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
  1083
(* alternative congruence rule (possibly more efficient)
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
  1084
lemma (in abelian_monoid) finsum_cong2:
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
  1085
  "[| !!i. i \<in> A ==> f i \<in> carrier G = True; A = B;
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
  1086
  !!i. i \<in> B ==> f i = g i |] ==> finsum G f A = finsum G g B"
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
  1087
  sorry*)
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
  1088
27717
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
  1089
lemma (in abelian_monoid) boundD_carrier:
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
  1090
  "[| bound \<zero> n f; n < m |] ==> f m \<in> carrier G"
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
  1091
  by auto
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
  1092
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
  1093
context ring
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
  1094
begin
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
  1095
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
  1096
theorem diagonal_sum:
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1097
  "[| f \<in> {..n + m::nat} -> carrier R; g \<in> {..n + m} -> carrier R |] ==>
14666
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
  1098
  (\<Oplus>k \<in> {..n + m}. \<Oplus>i \<in> {..k}. f i \<otimes> g (k - i)) =
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
  1099
  (\<Oplus>k \<in> {..n + m}. \<Oplus>i \<in> {..n + m - k}. f k \<otimes> g i)"
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1100
proof -
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1101
  assume Rf: "f \<in> {..n + m} -> carrier R" and Rg: "g \<in> {..n + m} -> carrier R"
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1102
  {
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1103
    fix j
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1104
    have "j <= n + m ==>
14666
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
  1105
      (\<Oplus>k \<in> {..j}. \<Oplus>i \<in> {..k}. f i \<otimes> g (k - i)) =
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
  1106
      (\<Oplus>k \<in> {..j}. \<Oplus>i \<in> {..j - k}. f k \<otimes> g i)"
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1107
    proof (induct j)
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1108
      case 0 from Rf Rg show ?case by (simp add: Pi_def)
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1109
    next
14666
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
  1110
      case (Suc j)
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1111
      have R6: "!!i k. [| k <= j; i <= Suc j - k |] ==> g i \<in> carrier R"
20217
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19984
diff changeset
  1112
        using Suc by (auto intro!: funcset_mem [OF Rg])
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1113
      have R8: "!!i k. [| k <= Suc j; i <= k |] ==> g (k - i) \<in> carrier R"
20217
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19984
diff changeset
  1114
        using Suc by (auto intro!: funcset_mem [OF Rg])
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1115
      have R9: "!!i k. [| k <= Suc j |] ==> f k \<in> carrier R"
14666
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
  1116
        using Suc by (auto intro!: funcset_mem [OF Rf])
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1117
      have R10: "!!i k. [| k <= Suc j; i <= Suc j - k |] ==> g i \<in> carrier R"
20217
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19984
diff changeset
  1118
        using Suc by (auto intro!: funcset_mem [OF Rg])
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1119
      have R11: "g 0 \<in> carrier R"
14666
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
  1120
        using Suc by (auto intro!: funcset_mem [OF Rg])
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1121
      from Suc show ?case
14666
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
  1122
        by (simp cong: finsum_cong add: Suc_diff_le a_ac
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
  1123
          Pi_def R6 R8 R9 R10 R11)
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1124
    qed
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1125
  }
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1126
  then show ?thesis by fast
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1127
qed
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1128
27717
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
  1129
theorem cauchy_product:
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1130
  assumes bf: "bound \<zero> n f" and bg: "bound \<zero> m g"
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1131
    and Rf: "f \<in> {..n} -> carrier R" and Rg: "g \<in> {..m} -> carrier R"
14666
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
  1132
  shows "(\<Oplus>k \<in> {..n + m}. \<Oplus>i \<in> {..k}. f i \<otimes> g (k - i)) =
17094
7a3c2efecffe Use interpretation in locales.
ballarin
parents: 16639
diff changeset
  1133
    (\<Oplus>i \<in> {..n}. f i) \<otimes> (\<Oplus>i \<in> {..m}. g i)"      (* State reverse direction? *)
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1134
proof -
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1135
  have f: "!!x. f x \<in> carrier R"
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1136
  proof -
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1137
    fix x
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1138
    show "f x \<in> carrier R"
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1139
      using Rf bf boundD_carrier by (cases "x <= n") (auto simp: Pi_def)
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1140
  qed
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1141
  have g: "!!x. g x \<in> carrier R"
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1142
  proof -
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1143
    fix x
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1144
    show "g x \<in> carrier R"
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1145
      using Rg bg boundD_carrier by (cases "x <= m") (auto simp: Pi_def)
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1146
  qed
14666
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
  1147
  from f g have "(\<Oplus>k \<in> {..n + m}. \<Oplus>i \<in> {..k}. f i \<otimes> g (k - i)) =
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
  1148
      (\<Oplus>k \<in> {..n + m}. \<Oplus>i \<in> {..n + m - k}. f k \<otimes> g i)"
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1149
    by (simp add: diagonal_sum Pi_def)
15045
d59f7e2e18d3 Moved to new m<..<n syntax for set intervals.
nipkow
parents: 14963
diff changeset
  1150
  also have "... = (\<Oplus>k \<in> {..n} \<union> {n<..n + m}. \<Oplus>i \<in> {..n + m - k}. f k \<otimes> g i)"
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1151
    by (simp only: ivl_disj_un_one)
14666
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
  1152
  also from f g have "... = (\<Oplus>k \<in> {..n}. \<Oplus>i \<in> {..n + m - k}. f k \<otimes> g i)"
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1153
    by (simp cong: finsum_cong
14666
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
  1154
      add: bound.bound [OF bf] finsum_Un_disjoint ivl_disj_int_one Pi_def)
15095
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
  1155
  also from f g
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
  1156
  have "... = (\<Oplus>k \<in> {..n}. \<Oplus>i \<in> {..m} \<union> {m<..n + m - k}. f k \<otimes> g i)"
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1157
    by (simp cong: finsum_cong add: ivl_disj_un_one le_add_diff Pi_def)
14666
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
  1158
  also from f g have "... = (\<Oplus>k \<in> {..n}. \<Oplus>i \<in> {..m}. f k \<otimes> g i)"
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1159
    by (simp cong: finsum_cong
14666
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
  1160
      add: bound.bound [OF bg] finsum_Un_disjoint ivl_disj_int_one Pi_def)
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
  1161
  also from f g have "... = (\<Oplus>i \<in> {..n}. f i) \<otimes> (\<Oplus>i \<in> {..m}. g i)"
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1162
    by (simp add: finsum_ldistr diagonal_sum Pi_def,
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1163
      simp cong: finsum_cong add: finsum_rdistr Pi_def)
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1164
  finally show ?thesis .
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1165
qed
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1166
27717
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
  1167
end
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
  1168
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
  1169
lemma (in UP_ring) const_ring_hom:
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1170
  "(%a. monom P a 0) \<in> ring_hom R P"
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1171
  by (auto intro!: ring_hom_memI intro: up_eqI simp: monom_mult_is_smult)
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1172
27933
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1173
definition
15095
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
  1174
  eval :: "[('a, 'm) ring_scheme, ('b, 'n) ring_scheme,
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
  1175
           'a => 'b, 'b, nat => 'a] => 'b"
27933
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1176
  where "eval R S phi s == \<lambda>p \<in> carrier (UP R).
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1177
    \<Oplus>\<^bsub>S\<^esub>i \<in> {..deg R p}. phi (coeff (UP R) p i) \<otimes>\<^bsub>S\<^esub> s (^)\<^bsub>S\<^esub> i"
15095
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
  1178
27717
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
  1179
context UP
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
  1180
begin
14666
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
  1181
27717
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
  1182
lemma eval_on_carrier:
19783
82f365a14960 Improved parameter management of locales.
ballarin
parents: 19582
diff changeset
  1183
  fixes S (structure)
17094
7a3c2efecffe Use interpretation in locales.
ballarin
parents: 16639
diff changeset
  1184
  shows "p \<in> carrier P ==>
7a3c2efecffe Use interpretation in locales.
ballarin
parents: 16639
diff changeset
  1185
  eval R S phi s p = (\<Oplus>\<^bsub>S\<^esub> i \<in> {..deg R p}. phi (coeff P p i) \<otimes>\<^bsub>S\<^esub> s (^)\<^bsub>S\<^esub> i)"
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1186
  by (unfold eval_def, fold P_def) simp
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1187
27717
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
  1188
lemma eval_extensional:
17094
7a3c2efecffe Use interpretation in locales.
ballarin
parents: 16639
diff changeset
  1189
  "eval R S phi p \<in> extensional (carrier P)"
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1190
  by (unfold eval_def, fold P_def) simp
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1191
27717
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
  1192
end
17094
7a3c2efecffe Use interpretation in locales.
ballarin
parents: 16639
diff changeset
  1193
7a3c2efecffe Use interpretation in locales.
ballarin
parents: 16639
diff changeset
  1194
text {* The universal property of the polynomial ring *}
7a3c2efecffe Use interpretation in locales.
ballarin
parents: 16639
diff changeset
  1195
29240
bb81c3709fb6 More porting to new locales.
ballarin
parents: 29237
diff changeset
  1196
locale UP_pre_univ_prop = ring_hom_cring + UP_cring
bb81c3709fb6 More porting to new locales.
ballarin
parents: 29237
diff changeset
  1197
bb81c3709fb6 More porting to new locales.
ballarin
parents: 29237
diff changeset
  1198
(* FIXME print_locale ring_hom_cring fails *)
17094
7a3c2efecffe Use interpretation in locales.
ballarin
parents: 16639
diff changeset
  1199
19783
82f365a14960 Improved parameter management of locales.
ballarin
parents: 19582
diff changeset
  1200
locale UP_univ_prop = UP_pre_univ_prop +
82f365a14960 Improved parameter management of locales.
ballarin
parents: 19582
diff changeset
  1201
  fixes s and Eval
17094
7a3c2efecffe Use interpretation in locales.
ballarin
parents: 16639
diff changeset
  1202
  assumes indet_img_carrier [simp, intro]: "s \<in> carrier S"
7a3c2efecffe Use interpretation in locales.
ballarin
parents: 16639
diff changeset
  1203
  defines Eval_def: "Eval == eval R S h s"
7a3c2efecffe Use interpretation in locales.
ballarin
parents: 16639
diff changeset
  1204
27717
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
  1205
text{*JE: I have moved the following lemma from Ring.thy and lifted then to the locale @{term ring_hom_ring} from @{term ring_hom_cring}.*}
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
  1206
text{*JE: I was considering using it in @{text eval_ring_hom}, but that property does not hold for non commutative rings, so 
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
  1207
  maybe it is not that necessary.*}
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
  1208
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
  1209
lemma (in ring_hom_ring) hom_finsum [simp]:
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
  1210
  "[| finite A; f \<in> A -> carrier R |] ==>
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
  1211
  h (finsum R f A) = finsum S (h o f) A"
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
  1212
proof (induct set: finite)
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
  1213
  case empty then show ?case by simp
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
  1214
next
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
  1215
  case insert then show ?case by (simp add: Pi_def)
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
  1216
qed
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
  1217
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
  1218
context UP_pre_univ_prop
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
  1219
begin
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
  1220
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
  1221
theorem eval_ring_hom:
17094
7a3c2efecffe Use interpretation in locales.
ballarin
parents: 16639
diff changeset
  1222
  assumes S: "s \<in> carrier S"
7a3c2efecffe Use interpretation in locales.
ballarin
parents: 16639
diff changeset
  1223
  shows "eval R S h s \<in> ring_hom P S"
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1224
proof (rule ring_hom_memI)
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1225
  fix p
17094
7a3c2efecffe Use interpretation in locales.
ballarin
parents: 16639
diff changeset
  1226
  assume R: "p \<in> carrier P"
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1227
  then show "eval R S h s p \<in> carrier S"
17094
7a3c2efecffe Use interpretation in locales.
ballarin
parents: 16639
diff changeset
  1228
    by (simp only: eval_on_carrier) (simp add: S Pi_def)
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1229
next
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1230
  fix p q
17094
7a3c2efecffe Use interpretation in locales.
ballarin
parents: 16639
diff changeset
  1231
  assume R: "p \<in> carrier P" "q \<in> carrier P"
15095
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
  1232
  then show "eval R S h s (p \<oplus>\<^bsub>P\<^esub> q) = eval R S h s p \<oplus>\<^bsub>S\<^esub> eval R S h s q"
17094
7a3c2efecffe Use interpretation in locales.
ballarin
parents: 16639
diff changeset
  1233
  proof (simp only: eval_on_carrier P.a_closed)
7a3c2efecffe Use interpretation in locales.
ballarin
parents: 16639
diff changeset
  1234
    from S R have
15095
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
  1235
      "(\<Oplus>\<^bsub>S \<^esub>i\<in>{..deg R (p \<oplus>\<^bsub>P\<^esub> q)}. h (coeff P (p \<oplus>\<^bsub>P\<^esub> q) i) \<otimes>\<^bsub>S\<^esub> s (^)\<^bsub>S\<^esub> i) =
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
  1236
      (\<Oplus>\<^bsub>S \<^esub>i\<in>{..deg R (p \<oplus>\<^bsub>P\<^esub> q)} \<union> {deg R (p \<oplus>\<^bsub>P\<^esub> q)<..max (deg R p) (deg R q)}.
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
  1237
        h (coeff P (p \<oplus>\<^bsub>P\<^esub> q) i) \<otimes>\<^bsub>S\<^esub> s (^)\<^bsub>S\<^esub> i)"
17094
7a3c2efecffe Use interpretation in locales.
ballarin
parents: 16639
diff changeset
  1238
      by (simp cong: S.finsum_cong
27717
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
  1239
        add: deg_aboveD S.finsum_Un_disjoint ivl_disj_int_one Pi_def del: coeff_add)
17094
7a3c2efecffe Use interpretation in locales.
ballarin
parents: 16639
diff changeset
  1240
    also from R have "... =
15095
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
  1241
        (\<Oplus>\<^bsub>S\<^esub> i \<in> {..max (deg R p) (deg R q)}.
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
  1242
          h (coeff P (p \<oplus>\<^bsub>P\<^esub> q) i) \<otimes>\<^bsub>S\<^esub> s (^)\<^bsub>S\<^esub> i)"
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1243
      by (simp add: ivl_disj_un_one)
17094
7a3c2efecffe Use interpretation in locales.
ballarin
parents: 16639
diff changeset
  1244
    also from R S have "... =
15095
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
  1245
      (\<Oplus>\<^bsub>S\<^esub>i\<in>{..max (deg R p) (deg R q)}. h (coeff P p i) \<otimes>\<^bsub>S\<^esub> s (^)\<^bsub>S\<^esub> i) \<oplus>\<^bsub>S\<^esub>
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
  1246
      (\<Oplus>\<^bsub>S\<^esub>i\<in>{..max (deg R p) (deg R q)}. h (coeff P q i) \<otimes>\<^bsub>S\<^esub> s (^)\<^bsub>S\<^esub> i)"
17094
7a3c2efecffe Use interpretation in locales.
ballarin
parents: 16639
diff changeset
  1247
      by (simp cong: S.finsum_cong
7a3c2efecffe Use interpretation in locales.
ballarin
parents: 16639
diff changeset
  1248
        add: S.l_distr deg_aboveD ivl_disj_int_one Pi_def)
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1249
    also have "... =
15095
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
  1250
        (\<Oplus>\<^bsub>S\<^esub> i \<in> {..deg R p} \<union> {deg R p<..max (deg R p) (deg R q)}.
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
  1251
          h (coeff P p i) \<otimes>\<^bsub>S\<^esub> s (^)\<^bsub>S\<^esub> i) \<oplus>\<^bsub>S\<^esub>
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
  1252
        (\<Oplus>\<^bsub>S\<^esub> i \<in> {..deg R q} \<union> {deg R q<..max (deg R p) (deg R q)}.
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
  1253
          h (coeff P q i) \<otimes>\<^bsub>S\<^esub> s (^)\<^bsub>S\<^esub> i)"
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1254
      by (simp only: ivl_disj_un_one le_maxI1 le_maxI2)
17094
7a3c2efecffe Use interpretation in locales.
ballarin
parents: 16639
diff changeset
  1255
    also from R S have "... =
15095
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
  1256
      (\<Oplus>\<^bsub>S\<^esub> i \<in> {..deg R p}. h (coeff P p i) \<otimes>\<^bsub>S\<^esub> s (^)\<^bsub>S\<^esub> i) \<oplus>\<^bsub>S\<^esub>
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
  1257
      (\<Oplus>\<^bsub>S\<^esub> i \<in> {..deg R q}. h (coeff P q i) \<otimes>\<^bsub>S\<^esub> s (^)\<^bsub>S\<^esub> i)"
17094
7a3c2efecffe Use interpretation in locales.
ballarin
parents: 16639
diff changeset
  1258
      by (simp cong: S.finsum_cong
7a3c2efecffe Use interpretation in locales.
ballarin
parents: 16639
diff changeset
  1259
        add: deg_aboveD S.finsum_Un_disjoint ivl_disj_int_one Pi_def)
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1260
    finally show
15095
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
  1261
      "(\<Oplus>\<^bsub>S\<^esub>i \<in> {..deg R (p \<oplus>\<^bsub>P\<^esub> q)}. h (coeff P (p \<oplus>\<^bsub>P\<^esub> q) i) \<otimes>\<^bsub>S\<^esub> s (^)\<^bsub>S\<^esub> i) =
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
  1262
      (\<Oplus>\<^bsub>S\<^esub>i \<in> {..deg R p}. h (coeff P p i) \<otimes>\<^bsub>S\<^esub> s (^)\<^bsub>S\<^esub> i) \<oplus>\<^bsub>S\<^esub>
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
  1263
      (\<Oplus>\<^bsub>S\<^esub>i \<in> {..deg R q}. h (coeff P q i) \<otimes>\<^bsub>S\<^esub> s (^)\<^bsub>S\<^esub> i)" .
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1264
  qed
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1265
next
17094
7a3c2efecffe Use interpretation in locales.
ballarin
parents: 16639
diff changeset
  1266
  show "eval R S h s \<one>\<^bsub>P\<^esub> = \<one>\<^bsub>S\<^esub>"
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1267
    by (simp only: eval_on_carrier UP_one_closed) simp
27717
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
  1268
next
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
  1269
  fix p q
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
  1270
  assume R: "p \<in> carrier P" "q \<in> carrier P"
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
  1271
  then show "eval R S h s (p \<otimes>\<^bsub>P\<^esub> q) = eval R S h s p \<otimes>\<^bsub>S\<^esub> eval R S h s q"
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
  1272
  proof (simp only: eval_on_carrier UP_mult_closed)
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
  1273
    from R S have
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
  1274
      "(\<Oplus>\<^bsub>S\<^esub> i \<in> {..deg R (p \<otimes>\<^bsub>P\<^esub> q)}. h (coeff P (p \<otimes>\<^bsub>P\<^esub> q) i) \<otimes>\<^bsub>S\<^esub> s (^)\<^bsub>S\<^esub> i) =
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
  1275
      (\<Oplus>\<^bsub>S\<^esub> i \<in> {..deg R (p \<otimes>\<^bsub>P\<^esub> q)} \<union> {deg R (p \<otimes>\<^bsub>P\<^esub> q)<..deg R p + deg R q}.
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
  1276
        h (coeff P (p \<otimes>\<^bsub>P\<^esub> q) i) \<otimes>\<^bsub>S\<^esub> s (^)\<^bsub>S\<^esub> i)"
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
  1277
      by (simp cong: S.finsum_cong
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
  1278
        add: deg_aboveD S.finsum_Un_disjoint ivl_disj_int_one Pi_def
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
  1279
        del: coeff_mult)
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
  1280
    also from R have "... =
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
  1281
      (\<Oplus>\<^bsub>S\<^esub> i \<in> {..deg R p + deg R q}. h (coeff P (p \<otimes>\<^bsub>P\<^esub> q) i) \<otimes>\<^bsub>S\<^esub> s (^)\<^bsub>S\<^esub> i)"
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
  1282
      by (simp only: ivl_disj_un_one deg_mult_ring)
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
  1283
    also from R S have "... =
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
  1284
      (\<Oplus>\<^bsub>S\<^esub> i \<in> {..deg R p + deg R q}.
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
  1285
         \<Oplus>\<^bsub>S\<^esub> k \<in> {..i}.
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
  1286
           h (coeff P p k) \<otimes>\<^bsub>S\<^esub> h (coeff P q (i - k)) \<otimes>\<^bsub>S\<^esub>
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
  1287
           (s (^)\<^bsub>S\<^esub> k \<otimes>\<^bsub>S\<^esub> s (^)\<^bsub>S\<^esub> (i - k)))"
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
  1288
      by (simp cong: S.finsum_cong add: S.nat_pow_mult Pi_def
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
  1289
        S.m_ac S.finsum_rdistr)
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
  1290
    also from R S have "... =
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
  1291
      (\<Oplus>\<^bsub>S\<^esub> i\<in>{..deg R p}. h (coeff P p i) \<otimes>\<^bsub>S\<^esub> s (^)\<^bsub>S\<^esub> i) \<otimes>\<^bsub>S\<^esub>
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
  1292
      (\<Oplus>\<^bsub>S\<^esub> i\<in>{..deg R q}. h (coeff P q i) \<otimes>\<^bsub>S\<^esub> s (^)\<^bsub>S\<^esub> i)"
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
  1293
      by (simp add: S.cauchy_product [THEN sym] bound.intro deg_aboveD S.m_ac
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
  1294
        Pi_def)
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
  1295
    finally show
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
  1296
      "(\<Oplus>\<^bsub>S\<^esub> i \<in> {..deg R (p \<otimes>\<^bsub>P\<^esub> q)}. h (coeff P (p \<otimes>\<^bsub>P\<^esub> q) i) \<otimes>\<^bsub>S\<^esub> s (^)\<^bsub>S\<^esub> i) =
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
  1297
      (\<Oplus>\<^bsub>S\<^esub> i \<in> {..deg R p}. h (coeff P p i) \<otimes>\<^bsub>S\<^esub> s (^)\<^bsub>S\<^esub> i) \<otimes>\<^bsub>S\<^esub>
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
  1298
      (\<Oplus>\<^bsub>S\<^esub> i \<in> {..deg R q}. h (coeff P q i) \<otimes>\<^bsub>S\<^esub> s (^)\<^bsub>S\<^esub> i)" .
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
  1299
  qed
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1300
qed
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1301
21502
7f3ea2b3bab6 prefer antiquotations over LaTeX macros;
wenzelm
parents: 20432
diff changeset
  1302
text {*
7f3ea2b3bab6 prefer antiquotations over LaTeX macros;
wenzelm
parents: 20432
diff changeset
  1303
  The following lemma could be proved in @{text UP_cring} with the additional
7f3ea2b3bab6 prefer antiquotations over LaTeX macros;
wenzelm
parents: 20432
diff changeset
  1304
  assumption that @{text h} is closed. *}
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1305
17094
7a3c2efecffe Use interpretation in locales.
ballarin
parents: 16639
diff changeset
  1306
lemma (in UP_pre_univ_prop) eval_const:
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1307
  "[| s \<in> carrier S; r \<in> carrier R |] ==> eval R S h s (monom P r 0) = h r"
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1308
  by (simp only: eval_on_carrier monom_closed) simp
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1309
27717
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
  1310
text {* Further properties of the evaluation homomorphism. *}
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
  1311
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1312
text {* The following proof is complicated by the fact that in arbitrary
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1313
  rings one might have @{term "one R = zero R"}. *}
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1314
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1315
(* TODO: simplify by cases "one R = zero R" *)
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1316
17094
7a3c2efecffe Use interpretation in locales.
ballarin
parents: 16639
diff changeset
  1317
lemma (in UP_pre_univ_prop) eval_monom1:
7a3c2efecffe Use interpretation in locales.
ballarin
parents: 16639
diff changeset
  1318
  assumes S: "s \<in> carrier S"
7a3c2efecffe Use interpretation in locales.
ballarin
parents: 16639
diff changeset
  1319
  shows "eval R S h s (monom P \<one> 1) = s"
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1320
proof (simp only: eval_on_carrier monom_closed R.one_closed)
17094
7a3c2efecffe Use interpretation in locales.
ballarin
parents: 16639
diff changeset
  1321
   from S have
15095
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
  1322
    "(\<Oplus>\<^bsub>S\<^esub> i\<in>{..deg R (monom P \<one> 1)}. h (coeff P (monom P \<one> 1) i) \<otimes>\<^bsub>S\<^esub> s (^)\<^bsub>S\<^esub> i) =
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
  1323
    (\<Oplus>\<^bsub>S\<^esub> i\<in>{..deg R (monom P \<one> 1)} \<union> {deg R (monom P \<one> 1)<..1}.
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
  1324
      h (coeff P (monom P \<one> 1) i) \<otimes>\<^bsub>S\<^esub> s (^)\<^bsub>S\<^esub> i)"
17094
7a3c2efecffe Use interpretation in locales.
ballarin
parents: 16639
diff changeset
  1325
    by (simp cong: S.finsum_cong del: coeff_monom
7a3c2efecffe Use interpretation in locales.
ballarin
parents: 16639
diff changeset
  1326
      add: deg_aboveD S.finsum_Un_disjoint ivl_disj_int_one Pi_def)
14666
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
  1327
  also have "... =
15095
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
  1328
    (\<Oplus>\<^bsub>S\<^esub> i \<in> {..1}. h (coeff P (monom P \<one> 1) i) \<otimes>\<^bsub>S\<^esub> s (^)\<^bsub>S\<^esub> i)"
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1329
    by (simp only: ivl_disj_un_one deg_monom_le R.one_closed)
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1330
  also have "... = s"
15095
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
  1331
  proof (cases "s = \<zero>\<^bsub>S\<^esub>")
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1332
    case True then show ?thesis by (simp add: Pi_def)
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1333
  next
17094
7a3c2efecffe Use interpretation in locales.
ballarin
parents: 16639
diff changeset
  1334
    case False then show ?thesis by (simp add: S Pi_def)
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1335
  qed
15095
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
  1336
  finally show "(\<Oplus>\<^bsub>S\<^esub> i \<in> {..deg R (monom P \<one> 1)}.
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
  1337
    h (coeff P (monom P \<one> 1) i) \<otimes>\<^bsub>S\<^esub> s (^)\<^bsub>S\<^esub> i) = s" .
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1338
qed
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1339
27717
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
  1340
end
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
  1341
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
  1342
text {* Interpretation of ring homomorphism lemmas. *}
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
  1343
29237
e90d9d51106b More porting to new locales.
ballarin
parents: 28823
diff changeset
  1344
sublocale UP_univ_prop < ring_hom_cring P S Eval
27717
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
  1345
  apply (unfold Eval_def)
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
  1346
  apply intro_locales
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
  1347
  apply (rule ring_hom_cring.axioms)
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
  1348
  apply (rule ring_hom_cring.intro)
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
  1349
  apply unfold_locales
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
  1350
  apply (rule eval_ring_hom)
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
  1351
  apply rule
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
  1352
  done
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
  1353
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1354
lemma (in UP_cring) monom_pow:
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1355
  assumes R: "a \<in> carrier R"
15095
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
  1356
  shows "(monom P a n) (^)\<^bsub>P\<^esub> m = monom P (a (^) m) (n * m)"
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1357
proof (induct m)
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1358
  case 0 from R show ?case by simp
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1359
next
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1360
  case Suc with R show ?case
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1361
    by (simp del: monom_mult add: monom_mult [THEN sym] add_commute)
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1362
qed
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1363
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1364
lemma (in ring_hom_cring) hom_pow [simp]:
15095
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
  1365
  "x \<in> carrier R ==> h (x (^) n) = h x (^)\<^bsub>S\<^esub> (n::nat)"
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1366
  by (induct n) simp_all
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1367
17094
7a3c2efecffe Use interpretation in locales.
ballarin
parents: 16639
diff changeset
  1368
lemma (in UP_univ_prop) Eval_monom:
7a3c2efecffe Use interpretation in locales.
ballarin
parents: 16639
diff changeset
  1369
  "r \<in> carrier R ==> Eval (monom P r n) = h r \<otimes>\<^bsub>S\<^esub> s (^)\<^bsub>S\<^esub> n"
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1370
proof -
17094
7a3c2efecffe Use interpretation in locales.
ballarin
parents: 16639
diff changeset
  1371
  assume R: "r \<in> carrier R"
7a3c2efecffe Use interpretation in locales.
ballarin
parents: 16639
diff changeset
  1372
  from R have "Eval (monom P r n) = Eval (monom P r 0 \<otimes>\<^bsub>P\<^esub> (monom P \<one> 1) (^)\<^bsub>P\<^esub> n)"
7a3c2efecffe Use interpretation in locales.
ballarin
parents: 16639
diff changeset
  1373
    by (simp del: monom_mult add: monom_mult [THEN sym] monom_pow)
15095
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
  1374
  also
17094
7a3c2efecffe Use interpretation in locales.
ballarin
parents: 16639
diff changeset
  1375
  from R eval_monom1 [where s = s, folded Eval_def]
7a3c2efecffe Use interpretation in locales.
ballarin
parents: 16639
diff changeset
  1376
  have "... = h r \<otimes>\<^bsub>S\<^esub> s (^)\<^bsub>S\<^esub> n"
7a3c2efecffe Use interpretation in locales.
ballarin
parents: 16639
diff changeset
  1377
    by (simp add: eval_const [where s = s, folded Eval_def])
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1378
  finally show ?thesis .
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1379
qed
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1380
17094
7a3c2efecffe Use interpretation in locales.
ballarin
parents: 16639
diff changeset
  1381
lemma (in UP_pre_univ_prop) eval_monom:
7a3c2efecffe Use interpretation in locales.
ballarin
parents: 16639
diff changeset
  1382
  assumes R: "r \<in> carrier R" and S: "s \<in> carrier S"
7a3c2efecffe Use interpretation in locales.
ballarin
parents: 16639
diff changeset
  1383
  shows "eval R S h s (monom P r n) = h r \<otimes>\<^bsub>S\<^esub> s (^)\<^bsub>S\<^esub> n"
15095
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
  1384
proof -
29237
e90d9d51106b More porting to new locales.
ballarin
parents: 28823
diff changeset
  1385
  interpret UP_univ_prop R S h P s "eval R S h s"
26202
51f8a696cd8d explicit referencing of background facts;
wenzelm
parents: 23350
diff changeset
  1386
    using UP_pre_univ_prop_axioms P_def R S
22931
11cc1ccad58e tuned proofs;
wenzelm
parents: 21502
diff changeset
  1387
    by (auto intro: UP_univ_prop.intro UP_univ_prop_axioms.intro)
17094
7a3c2efecffe Use interpretation in locales.
ballarin
parents: 16639
diff changeset
  1388
  from R
7a3c2efecffe Use interpretation in locales.
ballarin
parents: 16639
diff changeset
  1389
  show ?thesis by (rule Eval_monom)
7a3c2efecffe Use interpretation in locales.
ballarin
parents: 16639
diff changeset
  1390
qed
7a3c2efecffe Use interpretation in locales.
ballarin
parents: 16639
diff changeset
  1391
7a3c2efecffe Use interpretation in locales.
ballarin
parents: 16639
diff changeset
  1392
lemma (in UP_univ_prop) Eval_smult:
7a3c2efecffe Use interpretation in locales.
ballarin
parents: 16639
diff changeset
  1393
  "[| r \<in> carrier R; p \<in> carrier P |] ==> Eval (r \<odot>\<^bsub>P\<^esub> p) = h r \<otimes>\<^bsub>S\<^esub> Eval p"
7a3c2efecffe Use interpretation in locales.
ballarin
parents: 16639
diff changeset
  1394
proof -
7a3c2efecffe Use interpretation in locales.
ballarin
parents: 16639
diff changeset
  1395
  assume R: "r \<in> carrier R" and P: "p \<in> carrier P"
7a3c2efecffe Use interpretation in locales.
ballarin
parents: 16639
diff changeset
  1396
  then show ?thesis
7a3c2efecffe Use interpretation in locales.
ballarin
parents: 16639
diff changeset
  1397
    by (simp add: monom_mult_is_smult [THEN sym]
7a3c2efecffe Use interpretation in locales.
ballarin
parents: 16639
diff changeset
  1398
      eval_const [where s = s, folded Eval_def])
15095
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
  1399
qed
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1400
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1401
lemma ring_hom_cringI:
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1402
  assumes "cring R"
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1403
    and "cring S"
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1404
    and "h \<in> ring_hom R S"
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1405
  shows "ring_hom_cring R S h"
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1406
  by (fast intro: ring_hom_cring.intro ring_hom_cring_axioms.intro
27714
27b4d7c01f8b Tuned (for the sake of a meaningless log entry).
ballarin
parents: 27611
diff changeset
  1407
    cring.axioms assms)
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1408
27717
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
  1409
context UP_pre_univ_prop
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
  1410
begin
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
  1411
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
  1412
lemma UP_hom_unique:
27611
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 26934
diff changeset
  1413
  assumes "ring_hom_cring P S Phi"
17094
7a3c2efecffe Use interpretation in locales.
ballarin
parents: 16639
diff changeset
  1414
  assumes Phi: "Phi (monom P \<one> (Suc 0)) = s"
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1415
      "!!r. r \<in> carrier R ==> Phi (monom P r 0) = h r"
27611
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 26934
diff changeset
  1416
  assumes "ring_hom_cring P S Psi"
17094
7a3c2efecffe Use interpretation in locales.
ballarin
parents: 16639
diff changeset
  1417
  assumes Psi: "Psi (monom P \<one> (Suc 0)) = s"
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1418
      "!!r. r \<in> carrier R ==> Psi (monom P r 0) = h r"
17094
7a3c2efecffe Use interpretation in locales.
ballarin
parents: 16639
diff changeset
  1419
    and P: "p \<in> carrier P" and S: "s \<in> carrier S"
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1420
  shows "Phi p = Psi p"
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1421
proof -
29237
e90d9d51106b More porting to new locales.
ballarin
parents: 28823
diff changeset
  1422
  interpret ring_hom_cring P S Phi by fact
e90d9d51106b More porting to new locales.
ballarin
parents: 28823
diff changeset
  1423
  interpret ring_hom_cring P S Psi by fact
15095
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
  1424
  have "Phi p =
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
  1425
      Phi (\<Oplus>\<^bsub>P \<^esub>i \<in> {..deg R p}. monom P (coeff P p i) 0 \<otimes>\<^bsub>P\<^esub> monom P \<one> 1 (^)\<^bsub>P\<^esub> i)"
17094
7a3c2efecffe Use interpretation in locales.
ballarin
parents: 16639
diff changeset
  1426
    by (simp add: up_repr P monom_mult [THEN sym] monom_pow del: monom_mult)
15696
1da4ce092c0b First release of interpretation commands.
ballarin
parents: 15596
diff changeset
  1427
  also
1da4ce092c0b First release of interpretation commands.
ballarin
parents: 15596
diff changeset
  1428
  have "... =
15095
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
  1429
      Psi (\<Oplus>\<^bsub>P \<^esub>i\<in>{..deg R p}. monom P (coeff P p i) 0 \<otimes>\<^bsub>P\<^esub> monom P \<one> 1 (^)\<^bsub>P\<^esub> i)"
17094
7a3c2efecffe Use interpretation in locales.
ballarin
parents: 16639
diff changeset
  1430
    by (simp add: Phi Psi P Pi_def comp_def)
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1431
  also have "... = Psi p"
17094
7a3c2efecffe Use interpretation in locales.
ballarin
parents: 16639
diff changeset
  1432
    by (simp add: up_repr P monom_mult [THEN sym] monom_pow del: monom_mult)
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1433
  finally show ?thesis .
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1434
qed
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1435
27717
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
  1436
lemma ring_homD:
17094
7a3c2efecffe Use interpretation in locales.
ballarin
parents: 16639
diff changeset
  1437
  assumes Phi: "Phi \<in> ring_hom P S"
7a3c2efecffe Use interpretation in locales.
ballarin
parents: 16639
diff changeset
  1438
  shows "ring_hom_cring P S Phi"
7a3c2efecffe Use interpretation in locales.
ballarin
parents: 16639
diff changeset
  1439
proof (rule ring_hom_cring.intro)
7a3c2efecffe Use interpretation in locales.
ballarin
parents: 16639
diff changeset
  1440
  show "ring_hom_cring_axioms P S Phi"
7a3c2efecffe Use interpretation in locales.
ballarin
parents: 16639
diff changeset
  1441
  by (rule ring_hom_cring_axioms.intro) (rule Phi)
19984
29bb4659f80a Method intro_locales replaced by intro_locales and unfold_locales.
ballarin
parents: 19931
diff changeset
  1442
qed unfold_locales
17094
7a3c2efecffe Use interpretation in locales.
ballarin
parents: 16639
diff changeset
  1443
27717
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
  1444
theorem UP_universal_property:
17094
7a3c2efecffe Use interpretation in locales.
ballarin
parents: 16639
diff changeset
  1445
  assumes S: "s \<in> carrier S"
7a3c2efecffe Use interpretation in locales.
ballarin
parents: 16639
diff changeset
  1446
  shows "EX! Phi. Phi \<in> ring_hom P S \<inter> extensional (carrier P) &
14666
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
  1447
    Phi (monom P \<one> 1) = s &
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1448
    (ALL r : carrier R. Phi (monom P r 0) = h r)"
17094
7a3c2efecffe Use interpretation in locales.
ballarin
parents: 16639
diff changeset
  1449
  using S eval_monom1
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1450
  apply (auto intro: eval_ring_hom eval_const eval_extensional)
14666
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
  1451
  apply (rule extensionalityI)
17094
7a3c2efecffe Use interpretation in locales.
ballarin
parents: 16639
diff changeset
  1452
  apply (auto intro: UP_hom_unique ring_homD)
14666
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
  1453
  done
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1454
27717
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
  1455
end
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
  1456
27933
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1457
text{*JE: The following lemma was added by me; it might be even lifted to a simpler locale*}
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1458
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1459
context monoid
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1460
begin
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1461
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1462
lemma nat_pow_eone[simp]: assumes x_in_G: "x \<in> carrier G" shows "x (^) (1::nat) = x"
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1463
  using nat_pow_Suc [of x 0] unfolding nat_pow_0 [of x] unfolding l_one [OF x_in_G] by simp
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1464
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1465
end
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1466
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1467
context UP_ring
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1468
begin
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1469
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1470
abbreviation lcoeff :: "(nat =>'a) => 'a" where "lcoeff p == coeff P p (deg R p)"
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1471
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1472
lemma lcoeff_nonzero2: assumes p_in_R: "p \<in> carrier P" and p_not_zero: "p \<noteq> \<zero>\<^bsub>P\<^esub>" shows "lcoeff p \<noteq> \<zero>" 
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1473
  using lcoeff_nonzero [OF p_not_zero p_in_R] .
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1474
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1475
subsection{*The long division algorithm: some previous facts.*}
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1476
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1477
lemma coeff_minus [simp]:
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1478
  assumes p: "p \<in> carrier P" and q: "q \<in> carrier P" shows "coeff P (p \<ominus>\<^bsub>P\<^esub> q) n = coeff P p n \<ominus> coeff P q n" 
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1479
  unfolding a_minus_def [OF p q] unfolding coeff_add [OF p a_inv_closed [OF q]] unfolding coeff_a_inv [OF q]
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1480
  using coeff_closed [OF p, of n] using coeff_closed [OF q, of n] by algebra
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1481
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1482
lemma lcoeff_closed [simp]: assumes p: "p \<in> carrier P" shows "lcoeff p \<in> carrier R"
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1483
  using coeff_closed [OF p, of "deg R p"] by simp
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1484
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1485
lemma deg_smult_decr: assumes a_in_R: "a \<in> carrier R" and f_in_P: "f \<in> carrier P" shows "deg R (a \<odot>\<^bsub>P\<^esub> f) \<le> deg R f"
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1486
  using deg_smult_ring [OF a_in_R f_in_P] by (cases "a = \<zero>", auto)
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1487
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1488
lemma coeff_monom_mult: assumes R: "c \<in> carrier R" and P: "p \<in> carrier P" 
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1489
  shows "coeff P (monom P c n \<otimes>\<^bsub>P\<^esub> p) (m + n) = c \<otimes> (coeff P p m)"
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1490
proof -
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1491
  have "coeff P (monom P c n \<otimes>\<^bsub>P\<^esub> p) (m + n) = (\<Oplus>i\<in>{..m + n}. (if n = i then c else \<zero>) \<otimes> coeff P p (m + n - i))"
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1492
    unfolding coeff_mult [OF monom_closed [OF R, of n] P, of "m + n"] unfolding coeff_monom [OF R, of n] by simp
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1493
  also have "(\<Oplus>i\<in>{..m + n}. (if n = i then c else \<zero>) \<otimes> coeff P p (m + n - i)) = 
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1494
    (\<Oplus>i\<in>{..m + n}. (if n = i then c \<otimes> coeff P p (m + n - i) else \<zero>))"
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1495
    using  R.finsum_cong [of "{..m + n}" "{..m + n}" "(\<lambda>i::nat. (if n = i then c else \<zero>) \<otimes> coeff P p (m + n - i))" 
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1496
      "(\<lambda>i::nat. (if n = i then c \<otimes> coeff P p (m + n - i) else \<zero>))"]
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1497
    using coeff_closed [OF P] unfolding Pi_def simp_implies_def using R by auto
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1498
  also have "\<dots> = c \<otimes> coeff P p m" using R.finsum_singleton [of n "{..m + n}" "(\<lambda>i. c \<otimes> coeff P p (m + n - i))"]
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1499
    unfolding Pi_def using coeff_closed [OF P] using P R by auto
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1500
  finally show ?thesis by simp
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1501
qed
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1502
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1503
lemma deg_lcoeff_cancel: 
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1504
  assumes p_in_P: "p \<in> carrier P" and q_in_P: "q \<in> carrier P" and r_in_P: "r \<in> carrier P" 
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1505
  and deg_r_nonzero: "deg R r \<noteq> 0"
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1506
  and deg_R_p: "deg R p \<le> deg R r" and deg_R_q: "deg R q \<le> deg R r" 
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1507
  and coeff_R_p_eq_q: "coeff P p (deg R r) = \<ominus>\<^bsub>R\<^esub> (coeff P q (deg R r))"
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1508
  shows "deg R (p \<oplus>\<^bsub>P\<^esub> q) < deg R r"
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1509
proof -
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1510
  have deg_le: "deg R (p \<oplus>\<^bsub>P\<^esub> q) \<le> deg R r"
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1511
  proof (rule deg_aboveI)
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1512
    fix m
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1513
    assume deg_r_le: "deg R r < m"
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1514
    show "coeff P (p \<oplus>\<^bsub>P\<^esub> q) m = \<zero>"
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1515
    proof -
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1516
      have slp: "deg R p < m" and "deg R q < m" using deg_R_p deg_R_q using deg_r_le by auto
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1517
      then have max_sl: "max (deg R p) (deg R q) < m" by simp
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1518
      then have "deg R (p \<oplus>\<^bsub>P\<^esub> q) < m" using deg_add [OF p_in_P q_in_P] by arith
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1519
      with deg_R_p deg_R_q show ?thesis using coeff_add [OF p_in_P q_in_P, of m]
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  1520
        using deg_aboveD [of "p \<oplus>\<^bsub>P\<^esub> q" m] using p_in_P q_in_P by simp 
27933
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1521
    qed
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1522
  qed (simp add: p_in_P q_in_P)
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1523
  moreover have deg_ne: "deg R (p \<oplus>\<^bsub>P\<^esub> q) \<noteq> deg R r"
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1524
  proof (rule ccontr)
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1525
    assume nz: "\<not> deg R (p \<oplus>\<^bsub>P\<^esub> q) \<noteq> deg R r" then have deg_eq: "deg R (p \<oplus>\<^bsub>P\<^esub> q) = deg R r" by simp
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1526
    from deg_r_nonzero have r_nonzero: "r \<noteq> \<zero>\<^bsub>P\<^esub>" by (cases "r = \<zero>\<^bsub>P\<^esub>", simp_all)
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1527
    have "coeff P (p \<oplus>\<^bsub>P\<^esub> q) (deg R r) = \<zero>\<^bsub>R\<^esub>" using coeff_add [OF p_in_P q_in_P, of "deg R r"] using coeff_R_p_eq_q
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1528
      using coeff_closed [OF p_in_P, of "deg R r"] coeff_closed [OF q_in_P, of "deg R r"] by algebra
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1529
    with lcoeff_nonzero [OF r_nonzero r_in_P]  and deg_eq show False using lcoeff_nonzero [of "p \<oplus>\<^bsub>P\<^esub> q"] using p_in_P q_in_P
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1530
      using deg_r_nonzero by (cases "p \<oplus>\<^bsub>P\<^esub> q \<noteq> \<zero>\<^bsub>P\<^esub>", auto)
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1531
  qed
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1532
  ultimately show ?thesis by simp
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1533
qed
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1534
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1535
lemma monom_deg_mult: 
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1536
  assumes f_in_P: "f \<in> carrier P" and g_in_P: "g \<in> carrier P" and deg_le: "deg R g \<le> deg R f"
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1537
  and a_in_R: "a \<in> carrier R"
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1538
  shows "deg R (g \<otimes>\<^bsub>P\<^esub> monom P a (deg R f - deg R g)) \<le> deg R f"
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1539
  using deg_mult_ring [OF g_in_P monom_closed [OF a_in_R, of "deg R f - deg R g"]]
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1540
  apply (cases "a = \<zero>") using g_in_P apply simp 
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1541
  using deg_monom [OF _ a_in_R, of "deg R f - deg R g"] using deg_le by simp
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1542
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1543
lemma deg_zero_impl_monom:
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1544
  assumes f_in_P: "f \<in> carrier P" and deg_f: "deg R f = 0" 
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1545
  shows "f = monom P (coeff P f 0) 0"
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1546
  apply (rule up_eqI) using coeff_monom [OF coeff_closed [OF f_in_P], of 0 0]
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1547
  using f_in_P deg_f using deg_aboveD [of f _] by auto
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1548
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1549
end
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1550
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1551
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1552
subsection {* The long division proof for commutative rings *}
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1553
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1554
context UP_cring
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1555
begin
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1556
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1557
lemma exI3: assumes exist: "Pred x y z" 
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1558
  shows "\<exists> x y z. Pred x y z"
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1559
  using exist by blast
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1560
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1561
text {* Jacobson's Theorem 2.14 *}
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1562
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1563
lemma long_div_theorem: 
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1564
  assumes g_in_P [simp]: "g \<in> carrier P" and f_in_P [simp]: "f \<in> carrier P"
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1565
  and g_not_zero: "g \<noteq> \<zero>\<^bsub>P\<^esub>"
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1566
  shows "\<exists> q r (k::nat). (q \<in> carrier P) \<and> (r \<in> carrier P) \<and> (lcoeff g)(^)\<^bsub>R\<^esub>k \<odot>\<^bsub>P\<^esub> f = g \<otimes>\<^bsub>P\<^esub> q \<oplus>\<^bsub>P\<^esub> r \<and> (r = \<zero>\<^bsub>P\<^esub> | deg R r < deg R g)"
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1567
proof -
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1568
  let ?pred = "(\<lambda> q r (k::nat).
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1569
    (q \<in> carrier P) \<and> (r \<in> carrier P) \<and> (lcoeff g)(^)\<^bsub>R\<^esub>k \<odot>\<^bsub>P\<^esub> f = g \<otimes>\<^bsub>P\<^esub> q \<oplus>\<^bsub>P\<^esub> r \<and> (r = \<zero>\<^bsub>P\<^esub> | deg R r < deg R g))"
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1570
    and ?lg = "lcoeff g"
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1571
  show ?thesis
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1572
    (*JE: we distinguish some particular cases where the solution is almost direct.*)
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1573
  proof (cases "deg R f < deg R g")
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1574
    case True     
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1575
      (*JE: if the degree of f is smaller than the one of g the solution is straightforward.*)
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1576
      (* CB: avoid exI3 *)
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1577
      have "?pred \<zero>\<^bsub>P\<^esub> f 0" using True by force
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1578
      then show ?thesis by fast
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1579
  next
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1580
    case False then have deg_g_le_deg_f: "deg R g \<le> deg R f" by simp
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1581
    {
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1582
      (*JE: we now apply the induction hypothesis with some additional facts required*)
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1583
      from f_in_P deg_g_le_deg_f show ?thesis
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1584
      proof (induct n \<equiv> "deg R f" arbitrary: "f" rule: nat_less_induct)
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  1585
        fix n f
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  1586
        assume hypo: "\<forall>m<n. \<forall>x. x \<in> carrier P \<longrightarrow>
27933
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1587
          deg R g \<le> deg R x \<longrightarrow> 
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  1588
          m = deg R x \<longrightarrow>
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  1589
          (\<exists>q r (k::nat). q \<in> carrier P \<and> r \<in> carrier P \<and> lcoeff g (^) k \<odot>\<^bsub>P\<^esub> x = g \<otimes>\<^bsub>P\<^esub> q \<oplus>\<^bsub>P\<^esub> r & (r = \<zero>\<^bsub>P\<^esub> | deg R r < deg R g))"
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  1590
          and prem: "n = deg R f" and f_in_P [simp]: "f \<in> carrier P"
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  1591
          and deg_g_le_deg_f: "deg R g \<le> deg R f"
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  1592
        let ?k = "1::nat" and ?r = "(g \<otimes>\<^bsub>P\<^esub> (monom P (lcoeff f) (deg R f - deg R g))) \<oplus>\<^bsub>P\<^esub> \<ominus>\<^bsub>P\<^esub> (lcoeff g \<odot>\<^bsub>P\<^esub> f)"
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  1593
          and ?q = "monom P (lcoeff f) (deg R f - deg R g)"
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  1594
        show "\<exists> q r (k::nat). q \<in> carrier P \<and> r \<in> carrier P \<and> lcoeff g (^) k \<odot>\<^bsub>P\<^esub> f = g \<otimes>\<^bsub>P\<^esub> q \<oplus>\<^bsub>P\<^esub> r & (r = \<zero>\<^bsub>P\<^esub> | deg R r < deg R g)"
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  1595
        proof -
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  1596
          (*JE: we first extablish the existence of a triple satisfying the previous equation. 
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  1597
            Then we will have to prove the second part of the predicate.*)
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  1598
          have exist: "lcoeff g (^) ?k \<odot>\<^bsub>P\<^esub> f = g \<otimes>\<^bsub>P\<^esub> ?q \<oplus>\<^bsub>P\<^esub> \<ominus>\<^bsub>P\<^esub> ?r"
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  1599
            using minus_add
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  1600
            using sym [OF a_assoc [of "g \<otimes>\<^bsub>P\<^esub> ?q" "\<ominus>\<^bsub>P\<^esub> (g \<otimes>\<^bsub>P\<^esub> ?q)" "lcoeff g \<odot>\<^bsub>P\<^esub> f"]]
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  1601
            using r_neg by auto
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  1602
          show ?thesis
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  1603
          proof (cases "deg R (\<ominus>\<^bsub>P\<^esub> ?r) < deg R g")
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  1604
            (*JE: if the degree of the remainder satisfies the statement property we are done*)
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  1605
            case True
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  1606
            {
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  1607
              show ?thesis
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  1608
              proof (rule exI3 [of _ ?q "\<ominus>\<^bsub>P\<^esub> ?r" ?k], intro conjI)
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  1609
                show "lcoeff g (^) ?k \<odot>\<^bsub>P\<^esub> f = g \<otimes>\<^bsub>P\<^esub> ?q \<oplus>\<^bsub>P\<^esub> \<ominus>\<^bsub>P\<^esub> ?r" using exist by simp
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  1610
                show "\<ominus>\<^bsub>P\<^esub> ?r = \<zero>\<^bsub>P\<^esub> \<or> deg R (\<ominus>\<^bsub>P\<^esub> ?r) < deg R g" using True by simp
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  1611
              qed (simp_all)
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  1612
            }
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  1613
          next
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  1614
            case False note n_deg_r_l_deg_g = False
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  1615
            {
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  1616
              (*JE: otherwise, we verify the conditions of the induction hypothesis.*)
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  1617
              show ?thesis
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  1618
              proof (cases "deg R f = 0")
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  1619
                (*JE: the solutions are different if the degree of f is zero or not*)
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  1620
                case True
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  1621
                {
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  1622
                  have deg_g: "deg R g = 0" using True using deg_g_le_deg_f by simp
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  1623
                  have "lcoeff g (^) (1::nat) \<odot>\<^bsub>P\<^esub> f = g \<otimes>\<^bsub>P\<^esub> f \<oplus>\<^bsub>P\<^esub> \<zero>\<^bsub>P\<^esub>"
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  1624
                    unfolding deg_g apply simp
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  1625
                    unfolding sym [OF monom_mult_is_smult [OF coeff_closed [OF g_in_P, of 0] f_in_P]]
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  1626
                    using deg_zero_impl_monom [OF g_in_P deg_g] by simp
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  1627
                  then show ?thesis using f_in_P by blast
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  1628
                }
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  1629
              next
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  1630
                case False note deg_f_nzero = False
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  1631
                {
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  1632
                  (*JE: now it only remains the case where the induction hypothesis can be used.*)
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  1633
                  (*JE: we first prove that the degree of the remainder is smaller than the one of f*)
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  1634
                  have deg_remainder_l_f: "deg R (\<ominus>\<^bsub>P\<^esub> ?r) < n"
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  1635
                  proof -
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  1636
                    have "deg R (\<ominus>\<^bsub>P\<^esub> ?r) = deg R ?r" using deg_uminus [of ?r] by simp
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  1637
                    also have "\<dots> < deg R f"
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  1638
                    proof (rule deg_lcoeff_cancel)
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  1639
                      show "deg R (\<ominus>\<^bsub>P\<^esub> (lcoeff g \<odot>\<^bsub>P\<^esub> f)) \<le> deg R f"
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  1640
                        using deg_smult_ring [of "lcoeff g" f] using prem
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  1641
                        using lcoeff_nonzero2 [OF g_in_P g_not_zero] by simp
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  1642
                      show "deg R (g \<otimes>\<^bsub>P\<^esub> ?q) \<le> deg R f"
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  1643
                        using monom_deg_mult [OF _ g_in_P, of f "lcoeff f"] and deg_g_le_deg_f
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  1644
                        by simp
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  1645
                      show "coeff P (g \<otimes>\<^bsub>P\<^esub> ?q) (deg R f) = \<ominus> coeff P (\<ominus>\<^bsub>P\<^esub> (lcoeff g \<odot>\<^bsub>P\<^esub> f)) (deg R f)"
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  1646
                        unfolding coeff_mult [OF g_in_P monom_closed [OF lcoeff_closed [OF f_in_P], of "deg R f - deg R g"], of "deg R f"]
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  1647
                        unfolding coeff_monom [OF lcoeff_closed [OF f_in_P], of "(deg R f - deg R g)"]
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  1648
                        using R.finsum_cong' [of "{..deg R f}" "{..deg R f}" 
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  1649
                          "(\<lambda>i. coeff P g i \<otimes> (if deg R f - deg R g = deg R f - i then lcoeff f else \<zero>))" 
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  1650
                          "(\<lambda>i. if deg R g = i then coeff P g i \<otimes> lcoeff f else \<zero>)"]
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  1651
                        using R.finsum_singleton [of "deg R g" "{.. deg R f}" "(\<lambda>i. coeff P g i \<otimes> lcoeff f)"]
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  1652
                        unfolding Pi_def using deg_g_le_deg_f by force
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  1653
                    qed (simp_all add: deg_f_nzero)
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  1654
                    finally show "deg R (\<ominus>\<^bsub>P\<^esub> ?r) < n" unfolding prem .
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  1655
                  qed
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  1656
                  moreover have "\<ominus>\<^bsub>P\<^esub> ?r \<in> carrier P" by simp
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  1657
                  moreover obtain m where deg_rem_eq_m: "deg R (\<ominus>\<^bsub>P\<^esub> ?r) = m" by auto
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  1658
                  moreover have "deg R g \<le> deg R (\<ominus>\<^bsub>P\<^esub> ?r)" using n_deg_r_l_deg_g by simp
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  1659
                    (*JE: now, by applying the induction hypothesis, we obtain new quotient, remainder and exponent.*)
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  1660
                  ultimately obtain q' r' k'
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  1661
                    where rem_desc: "lcoeff g (^) (k'::nat) \<odot>\<^bsub>P\<^esub> (\<ominus>\<^bsub>P\<^esub> ?r) = g \<otimes>\<^bsub>P\<^esub> q' \<oplus>\<^bsub>P\<^esub> r'"and rem_deg: "(r' = \<zero>\<^bsub>P\<^esub> \<or> deg R r' < deg R g)"
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  1662
                    and q'_in_carrier: "q' \<in> carrier P" and r'_in_carrier: "r' \<in> carrier P"
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  1663
                    using hypo by blast
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  1664
                      (*JE: we now prove that the new quotient, remainder and exponent can be used to get 
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  1665
                      the quotient, remainder and exponent of the long division theorem*)
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  1666
                  show ?thesis
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  1667
                  proof (rule exI3 [of _ "((lcoeff g (^) k') \<odot>\<^bsub>P\<^esub> ?q \<oplus>\<^bsub>P\<^esub> q')" r' "Suc k'"], intro conjI)
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  1668
                    show "(lcoeff g (^) (Suc k')) \<odot>\<^bsub>P\<^esub> f = g \<otimes>\<^bsub>P\<^esub> ((lcoeff g (^) k') \<odot>\<^bsub>P\<^esub> ?q \<oplus>\<^bsub>P\<^esub> q') \<oplus>\<^bsub>P\<^esub> r'"
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  1669
                    proof -
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  1670
                      have "(lcoeff g (^) (Suc k')) \<odot>\<^bsub>P\<^esub> f = (lcoeff g (^) k') \<odot>\<^bsub>P\<^esub> (g \<otimes>\<^bsub>P\<^esub> ?q \<oplus>\<^bsub>P\<^esub> \<ominus>\<^bsub>P\<^esub> ?r)" 
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  1671
                        using smult_assoc1 exist by simp
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  1672
                      also have "\<dots> = (lcoeff g (^) k') \<odot>\<^bsub>P\<^esub> (g \<otimes>\<^bsub>P\<^esub> ?q) \<oplus>\<^bsub>P\<^esub> ((lcoeff g (^) k') \<odot>\<^bsub>P\<^esub> ( \<ominus>\<^bsub>P\<^esub> ?r))"
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  1673
                        using UP_smult_r_distr by simp
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  1674
                      also have "\<dots> = (lcoeff g (^) k') \<odot>\<^bsub>P\<^esub> (g \<otimes>\<^bsub>P\<^esub> ?q) \<oplus>\<^bsub>P\<^esub> (g \<otimes>\<^bsub>P\<^esub> q' \<oplus>\<^bsub>P\<^esub> r')"
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  1675
                        using rem_desc by simp
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  1676
                      also have "\<dots> = (lcoeff g (^) k') \<odot>\<^bsub>P\<^esub> (g \<otimes>\<^bsub>P\<^esub> ?q) \<oplus>\<^bsub>P\<^esub> g \<otimes>\<^bsub>P\<^esub> q' \<oplus>\<^bsub>P\<^esub> r'"
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  1677
                        using sym [OF a_assoc [of "lcoeff g (^) k' \<odot>\<^bsub>P\<^esub> (g \<otimes>\<^bsub>P\<^esub> ?q)" "g \<otimes>\<^bsub>P\<^esub> q'" "r'"]]
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  1678
                        using q'_in_carrier r'_in_carrier by simp
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  1679
                      also have "\<dots> = (lcoeff g (^) k') \<odot>\<^bsub>P\<^esub> (?q \<otimes>\<^bsub>P\<^esub> g) \<oplus>\<^bsub>P\<^esub> q' \<otimes>\<^bsub>P\<^esub> g \<oplus>\<^bsub>P\<^esub> r'"
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  1680
                        using q'_in_carrier by (auto simp add: m_comm)
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  1681
                      also have "\<dots> = (((lcoeff g (^) k') \<odot>\<^bsub>P\<^esub> ?q) \<otimes>\<^bsub>P\<^esub> g) \<oplus>\<^bsub>P\<^esub> q' \<otimes>\<^bsub>P\<^esub> g \<oplus>\<^bsub>P\<^esub> r'" 
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  1682
                        using smult_assoc2 q'_in_carrier by auto
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  1683
                      also have "\<dots> = ((lcoeff g (^) k') \<odot>\<^bsub>P\<^esub> ?q \<oplus>\<^bsub>P\<^esub> q') \<otimes>\<^bsub>P\<^esub> g \<oplus>\<^bsub>P\<^esub> r'"
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  1684
                        using sym [OF l_distr] and q'_in_carrier by auto
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  1685
                      finally show ?thesis using m_comm q'_in_carrier by auto
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  1686
                    qed
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  1687
                  qed (simp_all add: rem_deg q'_in_carrier r'_in_carrier)
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  1688
                }
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  1689
              qed
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  1690
            }
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  1691
          qed
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  1692
        qed
27933
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1693
      qed
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1694
    }
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1695
  qed
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1696
qed
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1697
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1698
end
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1699
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1700
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1701
text {*The remainder theorem as corollary of the long division theorem.*}
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1702
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1703
context UP_cring
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1704
begin
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1705
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1706
lemma deg_minus_monom:
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1707
  assumes a: "a \<in> carrier R"
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1708
  and R_not_trivial: "(carrier R \<noteq> {\<zero>})"
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1709
  shows "deg R (monom P \<one>\<^bsub>R\<^esub> 1 \<ominus>\<^bsub>P\<^esub> monom P a 0) = 1"
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1710
  (is "deg R ?g = 1")
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1711
proof -
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1712
  have "deg R ?g \<le> 1"
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1713
  proof (rule deg_aboveI)
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1714
    fix m
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1715
    assume "(1::nat) < m" 
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1716
    then show "coeff P ?g m = \<zero>" 
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1717
      using coeff_minus using a by auto algebra
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1718
  qed (simp add: a)
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1719
  moreover have "deg R ?g \<ge> 1"
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1720
  proof (rule deg_belowI)
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1721
    show "coeff P ?g 1 \<noteq> \<zero>"
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1722
      using a using R.carrier_one_not_zero R_not_trivial by simp algebra
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1723
  qed (simp add: a)
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1724
  ultimately show ?thesis by simp
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1725
qed
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1726
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1727
lemma lcoeff_monom:
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1728
  assumes a: "a \<in> carrier R" and R_not_trivial: "(carrier R \<noteq> {\<zero>})"
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1729
  shows "lcoeff (monom P \<one>\<^bsub>R\<^esub> 1 \<ominus>\<^bsub>P\<^esub> monom P a 0) = \<one>"
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1730
  using deg_minus_monom [OF a R_not_trivial]
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1731
  using coeff_minus a by auto algebra
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1732
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1733
lemma deg_nzero_nzero:
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1734
  assumes deg_p_nzero: "deg R p \<noteq> 0"
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1735
  shows "p \<noteq> \<zero>\<^bsub>P\<^esub>"
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1736
  using deg_zero deg_p_nzero by auto
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1737
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1738
lemma deg_monom_minus:
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1739
  assumes a: "a \<in> carrier R"
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1740
  and R_not_trivial: "carrier R \<noteq> {\<zero>}"
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1741
  shows "deg R (monom P \<one>\<^bsub>R\<^esub> 1 \<ominus>\<^bsub>P\<^esub> monom P a 0) = 1"
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1742
  (is "deg R ?g = 1")
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1743
proof -
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1744
  have "deg R ?g \<le> 1"
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1745
  proof (rule deg_aboveI)
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1746
    fix m::nat assume "1 < m" then show "coeff P ?g m = \<zero>" 
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1747
      using coeff_minus [OF monom_closed [OF R.one_closed, of 1] monom_closed [OF a, of 0], of m] 
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1748
      using coeff_monom [OF R.one_closed, of 1 m] using coeff_monom [OF a, of 0 m] by auto algebra
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1749
  qed (simp add: a)
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1750
  moreover have "1 \<le> deg R ?g"
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1751
  proof (rule deg_belowI)
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1752
    show "coeff P ?g 1 \<noteq> \<zero>" 
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1753
      using coeff_minus [OF monom_closed [OF R.one_closed, of 1] monom_closed [OF a, of 0], of 1]
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1754
      using coeff_monom [OF R.one_closed, of 1 1] using coeff_monom [OF a, of 0 1] 
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1755
      using R_not_trivial using R.carrier_one_not_zero
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1756
      by auto algebra
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1757
  qed (simp add: a)
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1758
  ultimately show ?thesis by simp
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1759
qed
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1760
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1761
lemma eval_monom_expr:
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1762
  assumes a: "a \<in> carrier R"
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1763
  shows "eval R R id a (monom P \<one>\<^bsub>R\<^esub> 1 \<ominus>\<^bsub>P\<^esub> monom P a 0) = \<zero>"
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1764
  (is "eval R R id a ?g = _")
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1765
proof -
29246
3593802c9cf1 More porting to new locales.
ballarin
parents: 29240
diff changeset
  1766
  interpret UP_pre_univ_prop R R id proof qed simp
27933
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1767
  have eval_ring_hom: "eval R R id a \<in> ring_hom P R" using eval_ring_hom [OF a] by simp
29237
e90d9d51106b More porting to new locales.
ballarin
parents: 28823
diff changeset
  1768
  interpret ring_hom_cring P R "eval R R id a" proof qed (simp add: eval_ring_hom)
27933
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1769
  have mon1_closed: "monom P \<one>\<^bsub>R\<^esub> 1 \<in> carrier P" 
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1770
    and mon0_closed: "monom P a 0 \<in> carrier P" 
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1771
    and min_mon0_closed: "\<ominus>\<^bsub>P\<^esub> monom P a 0 \<in> carrier P"
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1772
    using a R.a_inv_closed by auto
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1773
  have "eval R R id a ?g = eval R R id a (monom P \<one> 1) \<ominus> eval R R id a (monom P a 0)"
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1774
    unfolding P.minus_eq [OF mon1_closed mon0_closed]
29246
3593802c9cf1 More porting to new locales.
ballarin
parents: 29240
diff changeset
  1775
    unfolding hom_add [OF mon1_closed min_mon0_closed]
3593802c9cf1 More porting to new locales.
ballarin
parents: 29240
diff changeset
  1776
    unfolding hom_a_inv [OF mon0_closed] 
27933
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1777
    using R.minus_eq [symmetric] mon1_closed mon0_closed by auto
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1778
  also have "\<dots> = a \<ominus> a"
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1779
    using eval_monom [OF R.one_closed a, of 1] using eval_monom [OF a a, of 0] using a by simp
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1780
  also have "\<dots> = \<zero>"
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1781
    using a by algebra
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1782
  finally show ?thesis by simp
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1783
qed
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1784
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1785
lemma remainder_theorem_exist:
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1786
  assumes f: "f \<in> carrier P" and a: "a \<in> carrier R"
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1787
  and R_not_trivial: "carrier R \<noteq> {\<zero>}"
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1788
  shows "\<exists> q r. (q \<in> carrier P) \<and> (r \<in> carrier P) \<and> f = (monom P \<one>\<^bsub>R\<^esub> 1 \<ominus>\<^bsub>P\<^esub> monom P a 0) \<otimes>\<^bsub>P\<^esub> q \<oplus>\<^bsub>P\<^esub> r \<and> (deg R r = 0)"
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1789
  (is "\<exists> q r. (q \<in> carrier P) \<and> (r \<in> carrier P) \<and> f = ?g \<otimes>\<^bsub>P\<^esub> q \<oplus>\<^bsub>P\<^esub> r \<and> (deg R r = 0)")
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1790
proof -
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1791
  let ?g = "monom P \<one>\<^bsub>R\<^esub> 1 \<ominus>\<^bsub>P\<^esub> monom P a 0"
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1792
  from deg_minus_monom [OF a R_not_trivial]
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1793
  have deg_g_nzero: "deg R ?g \<noteq> 0" by simp
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1794
  have "\<exists>q r (k::nat). q \<in> carrier P \<and> r \<in> carrier P \<and>
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1795
    lcoeff ?g (^) k \<odot>\<^bsub>P\<^esub> f = ?g \<otimes>\<^bsub>P\<^esub> q \<oplus>\<^bsub>P\<^esub> r \<and> (r = \<zero>\<^bsub>P\<^esub> \<or> deg R r < deg R ?g)"
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1796
    using long_div_theorem [OF _ f deg_nzero_nzero [OF deg_g_nzero]] a
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1797
    by auto
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1798
  then show ?thesis
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1799
    unfolding lcoeff_monom [OF a R_not_trivial]
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1800
    unfolding deg_monom_minus [OF a R_not_trivial]
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1801
    using smult_one [OF f] using deg_zero by force
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1802
qed
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1803
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1804
lemma remainder_theorem_expression:
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1805
  assumes f [simp]: "f \<in> carrier P" and a [simp]: "a \<in> carrier R"
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1806
  and q [simp]: "q \<in> carrier P" and r [simp]: "r \<in> carrier P"
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1807
  and R_not_trivial: "carrier R \<noteq> {\<zero>}"
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1808
  and f_expr: "f = (monom P \<one>\<^bsub>R\<^esub> 1 \<ominus>\<^bsub>P\<^esub> monom P a 0) \<otimes>\<^bsub>P\<^esub> q \<oplus>\<^bsub>P\<^esub> r"
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1809
  (is "f = ?g \<otimes>\<^bsub>P\<^esub> q \<oplus>\<^bsub>P\<^esub> r" is "f = ?gq \<oplus>\<^bsub>P\<^esub> r")
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1810
    and deg_r_0: "deg R r = 0"
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1811
    shows "r = monom P (eval R R id a f) 0"
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1812
proof -
29237
e90d9d51106b More porting to new locales.
ballarin
parents: 28823
diff changeset
  1813
  interpret UP_pre_univ_prop R R id P proof qed simp
27933
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1814
  have eval_ring_hom: "eval R R id a \<in> ring_hom P R"
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1815
    using eval_ring_hom [OF a] by simp
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1816
  have "eval R R id a f = eval R R id a ?gq \<oplus>\<^bsub>R\<^esub> eval R R id a r"
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1817
    unfolding f_expr using ring_hom_add [OF eval_ring_hom] by auto
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1818
  also have "\<dots> = ((eval R R id a ?g) \<otimes> (eval R R id a q)) \<oplus>\<^bsub>R\<^esub> eval R R id a r"
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1819
    using ring_hom_mult [OF eval_ring_hom] by auto
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1820
  also have "\<dots> = \<zero> \<oplus> eval R R id a r"
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1821
    unfolding eval_monom_expr [OF a] using eval_ring_hom 
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1822
    unfolding ring_hom_def using q unfolding Pi_def by simp
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1823
  also have "\<dots> = eval R R id a r"
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1824
    using eval_ring_hom unfolding ring_hom_def using r unfolding Pi_def by simp
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1825
  finally have eval_eq: "eval R R id a f = eval R R id a r" by simp
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1826
  from deg_zero_impl_monom [OF r deg_r_0]
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1827
  have "r = monom P (coeff P r 0) 0" by simp
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1828
  with eval_const [OF a, of "coeff P r 0"] eval_eq 
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1829
  show ?thesis by auto
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1830
qed
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1831
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1832
corollary remainder_theorem:
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1833
  assumes f [simp]: "f \<in> carrier P" and a [simp]: "a \<in> carrier R"
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1834
  and R_not_trivial: "carrier R \<noteq> {\<zero>}"
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1835
  shows "\<exists> q r. (q \<in> carrier P) \<and> (r \<in> carrier P) \<and> 
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1836
     f = (monom P \<one>\<^bsub>R\<^esub> 1 \<ominus>\<^bsub>P\<^esub> monom P a 0) \<otimes>\<^bsub>P\<^esub> q \<oplus>\<^bsub>P\<^esub> monom P (eval R R id a f) 0"
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1837
  (is "\<exists> q r. (q \<in> carrier P) \<and> (r \<in> carrier P) \<and> f = ?g \<otimes>\<^bsub>P\<^esub> q \<oplus>\<^bsub>P\<^esub> monom P (eval R R id a f) 0")
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1838
proof -
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1839
  from remainder_theorem_exist [OF f a R_not_trivial]
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1840
  obtain q r
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1841
    where q_r: "q \<in> carrier P \<and> r \<in> carrier P \<and> f = ?g \<otimes>\<^bsub>P\<^esub> q \<oplus>\<^bsub>P\<^esub> r"
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1842
    and deg_r: "deg R r = 0" by force
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1843
  with remainder_theorem_expression [OF f a _ _ R_not_trivial, of q r]
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1844
  show ?thesis by auto
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1845
qed
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1846
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1847
end
4b867f6a65d3 Theorem on polynomial division and lemmas.
ballarin
parents: 27717
diff changeset
  1848
17094
7a3c2efecffe Use interpretation in locales.
ballarin
parents: 16639
diff changeset
  1849
20318
0e0ea63fe768 Restructured algebra library, added ideals and quotient rings.
ballarin
parents: 20282
diff changeset
  1850
subsection {* Sample Application of Evaluation Homomorphism *}
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1851
17094
7a3c2efecffe Use interpretation in locales.
ballarin
parents: 16639
diff changeset
  1852
lemma UP_pre_univ_propI:
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1853
  assumes "cring R"
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1854
    and "cring S"
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1855
    and "h \<in> ring_hom R S"
19931
fb32b43e7f80 Restructured locales with predicates: import is now an interpretation.
ballarin
parents: 19783
diff changeset
  1856
  shows "UP_pre_univ_prop R S h"
23350
50c5b0912a0c tuned proofs: avoid implicit prems;
wenzelm
parents: 22931
diff changeset
  1857
  using assms
19931
fb32b43e7f80 Restructured locales with predicates: import is now an interpretation.
ballarin
parents: 19783
diff changeset
  1858
  by (auto intro!: UP_pre_univ_prop.intro ring_hom_cring.intro
fb32b43e7f80 Restructured locales with predicates: import is now an interpretation.
ballarin
parents: 19783
diff changeset
  1859
    ring_hom_cring_axioms.intro UP_cring.intro)
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1860
27717
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
  1861
definition  INTEG :: "int ring"
21bbd410ba04 Generalised polynomial lemmas from cring to ring.
ballarin
parents: 27714
diff changeset
  1862
  where INTEG_def: "INTEG == (| carrier = UNIV, mult = op *, one = 1, zero = 0, add = op + |)"
13975
c8e9a89883ce Small changes for release Isabelle 2003.
ballarin
parents: 13949
diff changeset
  1863
15095
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
  1864
lemma INTEG_cring:
13975
c8e9a89883ce Small changes for release Isabelle 2003.
ballarin
parents: 13949
diff changeset
  1865
  "cring INTEG"
c8e9a89883ce Small changes for release Isabelle 2003.
ballarin
parents: 13949
diff changeset
  1866
  by (unfold INTEG_def) (auto intro!: cringI abelian_groupI comm_monoidI
c8e9a89883ce Small changes for release Isabelle 2003.
ballarin
parents: 13949
diff changeset
  1867
    zadd_zminus_inverse2 zadd_zmult_distrib)
c8e9a89883ce Small changes for release Isabelle 2003.
ballarin
parents: 13949
diff changeset
  1868
15095
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
  1869
lemma INTEG_id_eval:
17094
7a3c2efecffe Use interpretation in locales.
ballarin
parents: 16639
diff changeset
  1870
  "UP_pre_univ_prop INTEG INTEG id"
7a3c2efecffe Use interpretation in locales.
ballarin
parents: 16639
diff changeset
  1871
  by (fast intro: UP_pre_univ_propI INTEG_cring id_ring_hom)
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1872
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1873
text {*
17094
7a3c2efecffe Use interpretation in locales.
ballarin
parents: 16639
diff changeset
  1874
  Interpretation now enables to import all theorems and lemmas
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1875
  valid in the context of homomorphisms between @{term INTEG} and @{term
15095
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
  1876
  "UP INTEG"} globally.
14666
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
  1877
*}
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1878
30729
461ee3e49ad3 interpretation/interpret: prefixes are mandatory by default;
wenzelm
parents: 30363
diff changeset
  1879
interpretation INTEG: UP_pre_univ_prop INTEG INTEG id "UP INTEG"
28823
dcbef866c9e2 tuned unfold_locales invocation
haftmann
parents: 27933
diff changeset
  1880
  using INTEG_id_eval by simp_all
15763
b901a127ac73 Interpretation supports statically scoped attributes; documentation.
ballarin
parents: 15696
diff changeset
  1881
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1882
lemma INTEG_closed [intro, simp]:
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1883
  "z \<in> carrier INTEG"
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1884
  by (unfold INTEG_def) simp
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1885
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1886
lemma INTEG_mult [simp]:
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1887
  "mult INTEG z w = z * w"
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1888
  by (unfold INTEG_def) simp
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1889
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1890
lemma INTEG_pow [simp]:
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1891
  "pow INTEG z n = z ^ n"
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1892
  by (induct n) (simp_all add: INTEG_def nat_pow_def)
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1893
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1894
lemma "eval INTEG INTEG id 10 (monom (UP INTEG) 5 2) = 500"
15763
b901a127ac73 Interpretation supports statically scoped attributes; documentation.
ballarin
parents: 15696
diff changeset
  1895
  by (simp add: INTEG.eval_monom)
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1896
14590
276ef51cedbf simplified ML code for setsubgoaler;
wenzelm
parents: 14577
diff changeset
  1897
end