src/HOL/Algebra/UnivPoly.thy
author ballarin
Wed, 09 Mar 2005 18:44:52 +0100
changeset 15596 8665d08085df
parent 15481 fc075ae929e4
child 15696 1da4ce092c0b
permissions -rw-r--r--
First version of global registration command.
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
  Id:        $Id$
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
     4
  Author:    Clemens Ballarin, started 9 December 1996
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
     5
  Copyright: Clemens Ballarin
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
     6
*)
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
     7
14577
dbb95b825244 tuned document;
wenzelm
parents: 14553
diff changeset
     8
header {* Univariate Polynomials *}
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
     9
14577
dbb95b825244 tuned document;
wenzelm
parents: 14553
diff changeset
    10
theory UnivPoly = Module:
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
    11
14553
4740fc2da7bb Added brief intro text.
ballarin
parents: 14399
diff changeset
    12
text {*
14666
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
    13
  Polynomials are formalised as modules with additional operations for
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
    14
  extracting coefficients from polynomials and for obtaining monomials
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
    15
  from coefficients and exponents (record @{text "up_ring"}).  The
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
    16
  carrier set is a set of bounded functions from Nat to the
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
    17
  coefficient domain.  Bounded means that these functions return zero
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
    18
  above a certain bound (the degree).  There is a chapter on the
14706
71590b7733b7 tuned document;
wenzelm
parents: 14666
diff changeset
    19
  formalisation of polynomials in the PhD thesis \cite{Ballarin:1999},
71590b7733b7 tuned document;
wenzelm
parents: 14666
diff changeset
    20
  which was implemented with axiomatic type classes.  This was later
71590b7733b7 tuned document;
wenzelm
parents: 14666
diff changeset
    21
  ported to Locales.
14553
4740fc2da7bb Added brief intro text.
ballarin
parents: 14399
diff changeset
    22
*}
4740fc2da7bb Added brief intro text.
ballarin
parents: 14399
diff changeset
    23
14666
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
    24
13949
0ce528cd6f19 HOL-Algebra complete for release Isabelle2003 (modulo section headers).
ballarin
parents: 13940
diff changeset
    25
subsection {* The Constructor for Univariate Polynomials *}
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
    26
15095
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
    27
text {*
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
    28
  Functions with finite support.
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
    29
*}
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
    30
14666
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
    31
locale bound =
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
    32
  fixes z :: 'a
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
    33
    and n :: nat
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
    34
    and f :: "nat => 'a"
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
    35
  assumes bound: "!!m. n < m \<Longrightarrow> f m = z"
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
    36
14666
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
    37
declare bound.intro [intro!]
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
    38
  and bound.bound [dest]
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
    39
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
    40
lemma bound_below:
14666
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
    41
  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
    42
proof (rule classical)
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
    43
  assume "~ ?thesis"
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
    44
  then have "m < n" by arith
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
    45
  with bound have "f n = z" ..
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
    46
  with nonzero show ?thesis by contradiction
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
    47
qed
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
    48
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
    49
record ('a, 'p) up_ring = "('a, 'p) module" +
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
    50
  monom :: "['a, nat] => 'p"
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
    51
  coeff :: "['p, nat] => 'a"
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
    52
14651
02b8f3bcf7fe improved notation;
wenzelm
parents: 14590
diff changeset
    53
constdefs (structure R)
15095
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
    54
  up :: "('a, 'm) ring_scheme => (nat => 'a) set"
14651
02b8f3bcf7fe improved notation;
wenzelm
parents: 14590
diff changeset
    55
  "up R == {f. f \<in> UNIV -> carrier R & (EX n. bound \<zero> n f)}"
15095
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
    56
  UP :: "('a, 'm) ring_scheme => ('a, nat => 'a) up_ring"
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
    57
  "UP R == (|
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
    58
    carrier = up R,
14651
02b8f3bcf7fe improved notation;
wenzelm
parents: 14590
diff changeset
    59
    mult = (%p:up R. %q:up R. %n. \<Oplus>i \<in> {..n}. p i \<otimes> q (n-i)),
02b8f3bcf7fe improved notation;
wenzelm
parents: 14590
diff changeset
    60
    one = (%i. if i=0 then \<one> else \<zero>),
02b8f3bcf7fe improved notation;
wenzelm
parents: 14590
diff changeset
    61
    zero = (%i. \<zero>),
02b8f3bcf7fe improved notation;
wenzelm
parents: 14590
diff changeset
    62
    add = (%p:up R. %q:up R. %i. p i \<oplus> q i),
02b8f3bcf7fe improved notation;
wenzelm
parents: 14590
diff changeset
    63
    smult = (%a:carrier R. %p:up R. %i. a \<otimes> p i),
02b8f3bcf7fe improved notation;
wenzelm
parents: 14590
diff changeset
    64
    monom = (%a:carrier R. %n i. if i=n then a else \<zero>),
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
    65
    coeff = (%p:up R. %n. p n) |)"
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
    66
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
    67
text {*
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
    68
  Properties of the set of polynomials @{term up}.
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
    69
*}
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
    70
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
    71
lemma mem_upI [intro]:
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
    72
  "[| !!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
    73
  by (simp add: up_def Pi_def)
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_upD [dest]:
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
    76
  "f \<in> up R ==> f n \<in> carrier 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 (in cring) bound_upD [dest]:
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
    80
  "f \<in> up R ==> EX n. bound \<zero> n f"
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
    81
  by (simp add: up_def)
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
    82
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
    83
lemma (in cring) up_one_closed:
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
    84
   "(%n. if n = 0 then \<one> else \<zero>) \<in> up R"
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
    85
  using up_def by force
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
    86
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
    87
lemma (in cring) up_smult_closed:
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
    88
  "[| a \<in> carrier R; p \<in> up R |] ==> (%i. a \<otimes> p i) \<in> up R"
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
    89
  by force
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
    90
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
    91
lemma (in cring) up_add_closed:
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
    92
  "[| 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
    93
proof
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
    94
  fix n
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
    95
  assume "p \<in> up R" and "q \<in> up R"
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
    96
  then show "p n \<oplus> q n \<in> carrier R"
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
    97
    by auto
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
    98
next
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
    99
  assume UP: "p \<in> up R" "q \<in> up R"
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   100
  show "EX n. bound \<zero> n (%i. p i \<oplus> q i)"
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   101
  proof -
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   102
    from UP obtain n where boundn: "bound \<zero> n p" by fast
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   103
    from UP obtain m where boundm: "bound \<zero> m q" by fast
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   104
    have "bound \<zero> (max n m) (%i. p i \<oplus> q i)"
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   105
    proof
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   106
      fix i
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   107
      assume "max n m < i"
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   108
      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
   109
    qed
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   110
    then show ?thesis ..
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   111
  qed
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
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   114
lemma (in cring) up_a_inv_closed:
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   115
  "p \<in> up R ==> (%i. \<ominus> (p i)) \<in> up R"
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   116
proof
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   117
  assume R: "p \<in> up R"
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   118
  then obtain n where "bound \<zero> n p" by auto
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   119
  then have "bound \<zero> n (%i. \<ominus> p i)" by auto
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   120
  then show "EX n. bound \<zero> n (%i. \<ominus> p i)" by auto
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   121
qed auto
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   122
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   123
lemma (in cring) up_mult_closed:
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   124
  "[| p \<in> up R; q \<in> up R |] ==>
14666
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
   125
  (%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
   126
proof
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   127
  fix n
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   128
  assume "p \<in> up R" "q \<in> up R"
14666
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
   129
  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
   130
    by (simp add: mem_upD  funcsetI)
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   131
next
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   132
  assume UP: "p \<in> up R" "q \<in> up R"
14666
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
   133
  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
   134
  proof -
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   135
    from UP obtain n where boundn: "bound \<zero> n p" by fast
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   136
    from UP obtain m where boundm: "bound \<zero> m q" by fast
14666
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
   137
    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
   138
    proof
14666
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
   139
      fix k assume bound: "n + m < k"
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   140
      {
14666
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
   141
        fix i
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
   142
        have "p i \<otimes> q (k-i) = \<zero>"
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
   143
        proof (cases "n < i")
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
   144
          case True
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
   145
          with boundn have "p i = \<zero>" by auto
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   146
          moreover from UP have "q (k-i) \<in> carrier R" by auto
14666
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
   147
          ultimately show ?thesis by simp
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
   148
        next
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
   149
          case False
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
   150
          with bound have "m < k-i" by arith
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
   151
          with boundm have "q (k-i) = \<zero>" by auto
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
   152
          moreover from UP have "p i \<in> carrier R" by auto
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
   153
          ultimately show ?thesis by simp
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
   154
        qed
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   155
      }
14666
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
   156
      then show "(\<Oplus>i \<in> {..k}. p i \<otimes> q (k-i)) = \<zero>"
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
   157
        by (simp add: Pi_def)
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   158
    qed
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   159
    then show ?thesis by fast
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   160
  qed
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   161
qed
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   162
14666
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
   163
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   164
subsection {* Effect of operations on coefficients *}
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   165
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   166
locale UP = struct R + struct P +
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   167
  defines P_def: "P == UP R"
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   168
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   169
locale UP_cring = UP + cring R
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   170
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   171
locale UP_domain = UP_cring + "domain" R
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   172
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   173
text {*
15095
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
   174
  Temporarily declare @{thm [locale=UP] P_def} as simp rule.
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
   175
*}
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   176
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   177
declare (in UP) P_def [simp]
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   178
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   179
lemma (in UP_cring) coeff_monom [simp]:
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   180
  "a \<in> carrier R ==>
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   181
  coeff P (monom P a m) n = (if m=n then a else \<zero>)"
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   182
proof -
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   183
  assume R: "a \<in> carrier R"
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   184
  then have "(%n. if n = m then a else \<zero>) \<in> up R"
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   185
    using up_def by force
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   186
  with R show ?thesis by (simp add: UP_def)
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   187
qed
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   188
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   189
lemma (in UP_cring) coeff_zero [simp]:
15095
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
   190
  "coeff P \<zero>\<^bsub>P\<^esub> n = \<zero>"
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   191
  by (auto simp add: UP_def)
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   192
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   193
lemma (in UP_cring) coeff_one [simp]:
15095
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
   194
  "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
   195
  using up_one_closed by (simp add: UP_def)
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   196
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   197
lemma (in UP_cring) coeff_smult [simp]:
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   198
  "[| 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
   199
  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
   200
  by (simp add: UP_def up_smult_closed)
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   201
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   202
lemma (in UP_cring) coeff_add [simp]:
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   203
  "[| 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
   204
  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
   205
  by (simp add: UP_def up_add_closed)
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   206
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   207
lemma (in UP_cring) coeff_mult [simp]:
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   208
  "[| 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
   209
  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
   210
  by (simp add: UP_def up_mult_closed)
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   211
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   212
lemma (in UP) up_eqI:
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   213
  assumes prem: "!!n. coeff P p n = coeff P q n"
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   214
    and R: "p \<in> carrier P" "q \<in> carrier P"
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   215
  shows "p = q"
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   216
proof
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   217
  fix x
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   218
  from prem and R show "p x = q x" by (simp add: UP_def)
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   219
qed
14666
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
   220
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   221
subsection {* Polynomials form a commutative ring. *}
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   222
14666
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
   223
text {* Operations are closed over @{term P}. *}
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   224
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   225
lemma (in UP_cring) UP_mult_closed [simp]:
15095
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
   226
  "[| p \<in> carrier P; q \<in> carrier P |] ==> p \<otimes>\<^bsub>P\<^esub> q \<in> carrier P"
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   227
  by (simp add: UP_def up_mult_closed)
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   228
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   229
lemma (in UP_cring) UP_one_closed [simp]:
15095
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
   230
  "\<one>\<^bsub>P\<^esub> \<in> carrier P"
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   231
  by (simp add: UP_def up_one_closed)
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   232
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   233
lemma (in UP_cring) UP_zero_closed [intro, simp]:
15095
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
   234
  "\<zero>\<^bsub>P\<^esub> \<in> carrier P"
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   235
  by (auto simp add: UP_def)
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   236
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   237
lemma (in UP_cring) UP_a_closed [intro, simp]:
15095
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
   238
  "[| p \<in> carrier P; q \<in> carrier P |] ==> p \<oplus>\<^bsub>P\<^esub> q \<in> carrier P"
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   239
  by (simp add: UP_def up_add_closed)
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   240
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   241
lemma (in UP_cring) monom_closed [simp]:
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   242
  "a \<in> carrier R ==> monom P a n \<in> carrier P"
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   243
  by (auto simp add: UP_def up_def Pi_def)
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   244
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   245
lemma (in UP_cring) UP_smult_closed [simp]:
15095
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
   246
  "[| a \<in> carrier R; p \<in> carrier P |] ==> a \<odot>\<^bsub>P\<^esub> p \<in> carrier P"
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   247
  by (simp add: UP_def up_smult_closed)
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   248
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   249
lemma (in UP) coeff_closed [simp]:
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   250
  "p \<in> carrier P ==> coeff P p n \<in> carrier R"
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   251
  by (auto simp add: UP_def)
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   252
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   253
declare (in UP) P_def [simp del]
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   254
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   255
text {* Algebraic ring properties *}
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   256
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   257
lemma (in UP_cring) UP_a_assoc:
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   258
  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
   259
  shows "(p \<oplus>\<^bsub>P\<^esub> q) \<oplus>\<^bsub>P\<^esub> r = p \<oplus>\<^bsub>P\<^esub> (q \<oplus>\<^bsub>P\<^esub> r)"
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   260
  by (rule up_eqI, simp add: a_assoc R, simp_all add: R)
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   261
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   262
lemma (in UP_cring) UP_l_zero [simp]:
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   263
  assumes R: "p \<in> carrier P"
15095
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
   264
  shows "\<zero>\<^bsub>P\<^esub> \<oplus>\<^bsub>P\<^esub> p = p"
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   265
  by (rule up_eqI, simp_all add: R)
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   266
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   267
lemma (in UP_cring) UP_l_neg_ex:
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   268
  assumes R: "p \<in> carrier P"
15095
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
   269
  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
   270
proof -
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   271
  let ?q = "%i. \<ominus> (p i)"
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   272
  from R have closed: "?q \<in> carrier P"
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   273
    by (simp add: UP_def P_def up_a_inv_closed)
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   274
  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
   275
    by (simp add: UP_def P_def up_a_inv_closed)
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   276
  show ?thesis
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   277
  proof
15095
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
   278
    show "?q \<oplus>\<^bsub>P\<^esub> p = \<zero>\<^bsub>P\<^esub>"
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   279
      by (auto intro!: up_eqI simp add: R closed coeff R.l_neg)
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   280
  qed (rule closed)
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   281
qed
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   282
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   283
lemma (in UP_cring) UP_a_comm:
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   284
  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
   285
  shows "p \<oplus>\<^bsub>P\<^esub> q = q \<oplus>\<^bsub>P\<^esub> p"
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   286
  by (rule up_eqI, simp add: a_comm R, simp_all add: R)
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   287
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   288
ML_setup {*
14590
276ef51cedbf simplified ML code for setsubgoaler;
wenzelm
parents: 14577
diff changeset
   289
  simpset_ref() := simpset() setsubgoaler asm_full_simp_tac;
276ef51cedbf simplified ML code for setsubgoaler;
wenzelm
parents: 14577
diff changeset
   290
*}
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   291
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   292
lemma (in UP_cring) UP_m_assoc:
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   293
  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
   294
  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
   295
proof (rule up_eqI)
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   296
  fix n
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   297
  {
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   298
    fix k and a b c :: "nat=>'a"
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   299
    assume R: "a \<in> UNIV -> carrier R" "b \<in> UNIV -> carrier R"
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   300
      "c \<in> UNIV -> carrier R"
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   301
    then have "k <= n ==>
14666
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
   302
      (\<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
   303
      (\<Oplus>j \<in> {..k}. a j \<otimes> (\<Oplus>i \<in> {..k-j}. b i \<otimes> c (n-j-i)))"
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
   304
      (concl is "?eq k")
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   305
    proof (induct k)
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   306
      case 0 then show ?case by (simp add: Pi_def m_assoc)
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   307
    next
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   308
      case (Suc k)
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   309
      then have "k <= n" by arith
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   310
      then have "?eq k" by (rule Suc)
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   311
      with R show ?case
14666
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
   312
        by (simp cong: finsum_cong
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   313
             add: Suc_diff_le Pi_def l_distr r_distr m_assoc)
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   314
          (simp cong: finsum_cong add: Pi_def a_ac finsum_ldistr m_assoc)
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   315
    qed
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   316
  }
15095
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
   317
  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
   318
    by (simp add: Pi_def)
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   319
qed (simp_all add: R)
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   320
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   321
ML_setup {*
14590
276ef51cedbf simplified ML code for setsubgoaler;
wenzelm
parents: 14577
diff changeset
   322
  simpset_ref() := simpset() setsubgoaler asm_simp_tac;
276ef51cedbf simplified ML code for setsubgoaler;
wenzelm
parents: 14577
diff changeset
   323
*}
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   324
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   325
lemma (in UP_cring) UP_l_one [simp]:
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   326
  assumes R: "p \<in> carrier P"
15095
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
   327
  shows "\<one>\<^bsub>P\<^esub> \<otimes>\<^bsub>P\<^esub> p = p"
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   328
proof (rule up_eqI)
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   329
  fix n
15095
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
   330
  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
   331
  proof (cases n)
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   332
    case 0 with R show ?thesis by simp
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   333
  next
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   334
    case Suc with R show ?thesis
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   335
      by (simp del: finsum_Suc add: finsum_Suc2 Pi_def)
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   336
  qed
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   337
qed (simp_all add: R)
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   338
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   339
lemma (in UP_cring) UP_l_distr:
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   340
  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
   341
  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
   342
  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
   343
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   344
lemma (in UP_cring) UP_m_comm:
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   345
  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
   346
  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
   347
proof (rule up_eqI)
14666
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
   348
  fix n
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   349
  {
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   350
    fix k and a b :: "nat=>'a"
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   351
    assume R: "a \<in> UNIV -> carrier R" "b \<in> UNIV -> carrier R"
14666
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
   352
    then have "k <= n ==>
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
   353
      (\<Oplus>i \<in> {..k}. a i \<otimes> b (n-i)) =
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
   354
      (\<Oplus>i \<in> {..k}. a (k-i) \<otimes> b (i+n-k))"
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
   355
      (concl is "?eq k")
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   356
    proof (induct k)
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   357
      case 0 then show ?case by (simp add: Pi_def)
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   358
    next
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   359
      case (Suc k) then show ?case
15481
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents: 15095
diff changeset
   360
        by (simplesubst finsum_Suc2) (simp add: Pi_def a_comm)+
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   361
    qed
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   362
  }
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   363
  note l = this
15095
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
   364
  from R show "coeff P (p \<otimes>\<^bsub>P\<^esub> q) n =  coeff P (q \<otimes>\<^bsub>P\<^esub> p) n"
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   365
    apply (simp add: Pi_def)
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   366
    apply (subst l)
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   367
    apply (auto simp add: Pi_def)
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   368
    apply (simp add: m_comm)
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   369
    done
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   370
qed (simp_all add: R)
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   371
15596
8665d08085df First version of global registration command.
ballarin
parents: 15481
diff changeset
   372
(*
8665d08085df First version of global registration command.
ballarin
parents: 15481
diff changeset
   373
Strange phenomenon in Isar:
8665d08085df First version of global registration command.
ballarin
parents: 15481
diff changeset
   374
8665d08085df First version of global registration command.
ballarin
parents: 15481
diff changeset
   375
theorem (in UP_cring) UP_cring:
8665d08085df First version of global registration command.
ballarin
parents: 15481
diff changeset
   376
  "cring P"
8665d08085df First version of global registration command.
ballarin
parents: 15481
diff changeset
   377
proof (rule cringI)
8665d08085df First version of global registration command.
ballarin
parents: 15481
diff changeset
   378
  show "abelian_group P" proof (rule abelian_groupI)
8665d08085df First version of global registration command.
ballarin
parents: 15481
diff changeset
   379
  fix x y z
8665d08085df First version of global registration command.
ballarin
parents: 15481
diff changeset
   380
  assume "x \<in> carrier P" and "y \<in> carrier P" and "z \<in> carrier P"
8665d08085df First version of global registration command.
ballarin
parents: 15481
diff changeset
   381
  {
8665d08085df First version of global registration command.
ballarin
parents: 15481
diff changeset
   382
  show "x \<oplus>\<^bsub>P\<^esub> y \<in> carrier P" sorry
8665d08085df First version of global registration command.
ballarin
parents: 15481
diff changeset
   383
  next
8665d08085df First version of global registration command.
ballarin
parents: 15481
diff changeset
   384
  show "x \<oplus>\<^bsub>P\<^esub> y \<oplus>\<^bsub>P\<^esub> z = x \<oplus>\<^bsub>P\<^esub> (y \<oplus>\<^bsub>P\<^esub> z)" sorry
8665d08085df First version of global registration command.
ballarin
parents: 15481
diff changeset
   385
  next
8665d08085df First version of global registration command.
ballarin
parents: 15481
diff changeset
   386
  show "x \<oplus>\<^bsub>P\<^esub> y = y \<oplus>\<^bsub>P\<^esub> x" sorry
8665d08085df First version of global registration command.
ballarin
parents: 15481
diff changeset
   387
  next
8665d08085df First version of global registration command.
ballarin
parents: 15481
diff changeset
   388
  show "\<zero>\<^bsub>P\<^esub> \<oplus>\<^bsub>P\<^esub> x = x" sorry
8665d08085df First version of global registration command.
ballarin
parents: 15481
diff changeset
   389
  next
8665d08085df First version of global registration command.
ballarin
parents: 15481
diff changeset
   390
  show "\<exists>y\<in>carrier P. y \<oplus>\<^bsub>P\<^esub> x = \<zero>\<^bsub>P\<^esub>" sorry
8665d08085df First version of global registration command.
ballarin
parents: 15481
diff changeset
   391
  next
8665d08085df First version of global registration command.
ballarin
parents: 15481
diff changeset
   392
  show "\<zero>\<^bsub>P\<^esub> \<in> carrier P" sorry  last goal rejected!!!
8665d08085df First version of global registration command.
ballarin
parents: 15481
diff changeset
   393
*)
8665d08085df First version of global registration command.
ballarin
parents: 15481
diff changeset
   394
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   395
theorem (in UP_cring) UP_cring:
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   396
  "cring P"
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   397
  by (auto intro!: cringI abelian_groupI comm_monoidI UP_a_assoc UP_l_zero
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   398
    UP_l_neg_ex UP_a_comm UP_m_assoc UP_l_one UP_m_comm UP_l_distr)
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   399
14399
dc677b35e54f New lemmas about inversion of restricted functions.
ballarin
parents: 13975
diff changeset
   400
lemma (in UP_cring) UP_ring:  (* preliminary *)
dc677b35e54f New lemmas about inversion of restricted functions.
ballarin
parents: 13975
diff changeset
   401
  "ring P"
dc677b35e54f New lemmas about inversion of restricted functions.
ballarin
parents: 13975
diff changeset
   402
  by (auto intro: ring.intro cring.axioms UP_cring)
dc677b35e54f New lemmas about inversion of restricted functions.
ballarin
parents: 13975
diff changeset
   403
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   404
lemma (in UP_cring) UP_a_inv_closed [intro, simp]:
15095
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
   405
  "p \<in> carrier P ==> \<ominus>\<^bsub>P\<^esub> p \<in> carrier P"
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   406
  by (rule abelian_group.a_inv_closed
14399
dc677b35e54f New lemmas about inversion of restricted functions.
ballarin
parents: 13975
diff changeset
   407
    [OF ring.is_abelian_group [OF UP_ring]])
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   408
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   409
lemma (in UP_cring) coeff_a_inv [simp]:
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   410
  assumes R: "p \<in> carrier P"
15095
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
   411
  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
   412
proof -
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   413
  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
   414
    "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
   415
    by algebra
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   416
  also from R have "... =  \<ominus> (coeff P p n)"
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   417
    by (simp del: coeff_add add: coeff_add [THEN sym]
14399
dc677b35e54f New lemmas about inversion of restricted functions.
ballarin
parents: 13975
diff changeset
   418
      abelian_group.r_neg [OF ring.is_abelian_group [OF UP_ring]])
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   419
  finally show ?thesis .
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   420
qed
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   421
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   422
text {*
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   423
  Instantiation of lemmas from @{term cring}.
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   424
*}
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   425
15095
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
   426
(* TODO: this should be automated with an instantiation command. *)
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
   427
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   428
lemma (in UP_cring) UP_monoid:
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   429
  "monoid P"
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   430
  by (fast intro!: cring.is_comm_monoid comm_monoid.axioms monoid.intro
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   431
    UP_cring)
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   432
(* TODO: provide cring.is_monoid *)
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   433
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   434
lemma (in UP_cring) UP_comm_monoid:
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   435
  "comm_monoid P"
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   436
  by (fast intro!: cring.is_comm_monoid UP_cring)
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   437
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   438
lemma (in UP_cring) UP_abelian_monoid:
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   439
  "abelian_monoid P"
14399
dc677b35e54f New lemmas about inversion of restricted functions.
ballarin
parents: 13975
diff changeset
   440
  by (fast intro!: abelian_group.axioms ring.is_abelian_group UP_ring)
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   441
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   442
lemma (in UP_cring) UP_abelian_group:
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   443
  "abelian_group P"
14399
dc677b35e54f New lemmas about inversion of restricted functions.
ballarin
parents: 13975
diff changeset
   444
  by (fast intro!: ring.is_abelian_group UP_ring)
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   445
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   446
lemmas (in UP_cring) UP_r_one [simp] =
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   447
  monoid.r_one [OF UP_monoid]
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   448
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   449
lemmas (in UP_cring) UP_nat_pow_closed [intro, simp] =
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   450
  monoid.nat_pow_closed [OF UP_monoid]
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   451
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   452
lemmas (in UP_cring) UP_nat_pow_0 [simp] =
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   453
  monoid.nat_pow_0 [OF UP_monoid]
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   454
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   455
lemmas (in UP_cring) UP_nat_pow_Suc [simp] =
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   456
  monoid.nat_pow_Suc [OF UP_monoid]
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   457
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   458
lemmas (in UP_cring) UP_nat_pow_one [simp] =
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   459
  monoid.nat_pow_one [OF UP_monoid]
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   460
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   461
lemmas (in UP_cring) UP_nat_pow_mult =
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   462
  monoid.nat_pow_mult [OF UP_monoid]
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   463
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   464
lemmas (in UP_cring) UP_nat_pow_pow =
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   465
  monoid.nat_pow_pow [OF UP_monoid]
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   466
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   467
lemmas (in UP_cring) UP_m_lcomm =
14963
d584e32f7d46 removal of magmas and semigroups
paulson
parents: 14706
diff changeset
   468
  comm_monoid.m_lcomm [OF UP_comm_monoid]
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   469
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   470
lemmas (in UP_cring) UP_m_ac = UP_m_assoc UP_m_comm UP_m_lcomm
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   471
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   472
lemmas (in UP_cring) UP_nat_pow_distr =
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   473
  comm_monoid.nat_pow_distr [OF UP_comm_monoid]
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   474
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   475
lemmas (in UP_cring) UP_a_lcomm = abelian_monoid.a_lcomm [OF UP_abelian_monoid]
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   476
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   477
lemmas (in UP_cring) UP_r_zero [simp] =
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   478
  abelian_monoid.r_zero [OF UP_abelian_monoid]
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   479
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   480
lemmas (in UP_cring) UP_a_ac = UP_a_assoc UP_a_comm UP_a_lcomm
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   481
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   482
lemmas (in UP_cring) UP_finsum_empty [simp] =
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   483
  abelian_monoid.finsum_empty [OF UP_abelian_monoid]
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   484
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   485
lemmas (in UP_cring) UP_finsum_insert [simp] =
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   486
  abelian_monoid.finsum_insert [OF UP_abelian_monoid]
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   487
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   488
lemmas (in UP_cring) UP_finsum_zero [simp] =
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   489
  abelian_monoid.finsum_zero [OF UP_abelian_monoid]
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   490
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   491
lemmas (in UP_cring) UP_finsum_closed [simp] =
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   492
  abelian_monoid.finsum_closed [OF UP_abelian_monoid]
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   493
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   494
lemmas (in UP_cring) UP_finsum_Un_Int =
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   495
  abelian_monoid.finsum_Un_Int [OF UP_abelian_monoid]
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   496
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   497
lemmas (in UP_cring) UP_finsum_Un_disjoint =
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   498
  abelian_monoid.finsum_Un_disjoint [OF UP_abelian_monoid]
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   499
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   500
lemmas (in UP_cring) UP_finsum_addf =
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   501
  abelian_monoid.finsum_addf [OF UP_abelian_monoid]
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   502
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   503
lemmas (in UP_cring) UP_finsum_cong' =
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   504
  abelian_monoid.finsum_cong' [OF UP_abelian_monoid]
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   505
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   506
lemmas (in UP_cring) UP_finsum_0 [simp] =
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   507
  abelian_monoid.finsum_0 [OF UP_abelian_monoid]
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   508
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   509
lemmas (in UP_cring) UP_finsum_Suc [simp] =
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   510
  abelian_monoid.finsum_Suc [OF UP_abelian_monoid]
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   511
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   512
lemmas (in UP_cring) UP_finsum_Suc2 =
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   513
  abelian_monoid.finsum_Suc2 [OF UP_abelian_monoid]
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   514
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   515
lemmas (in UP_cring) UP_finsum_add [simp] =
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   516
  abelian_monoid.finsum_add [OF UP_abelian_monoid]
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   517
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   518
lemmas (in UP_cring) UP_finsum_cong =
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   519
  abelian_monoid.finsum_cong [OF UP_abelian_monoid]
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   520
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   521
lemmas (in UP_cring) UP_minus_closed [intro, simp] =
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   522
  abelian_group.minus_closed [OF UP_abelian_group]
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   523
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   524
lemmas (in UP_cring) UP_a_l_cancel [simp] =
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   525
  abelian_group.a_l_cancel [OF UP_abelian_group]
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   526
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   527
lemmas (in UP_cring) UP_a_r_cancel [simp] =
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   528
  abelian_group.a_r_cancel [OF UP_abelian_group]
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   529
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   530
lemmas (in UP_cring) UP_l_neg =
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   531
  abelian_group.l_neg [OF UP_abelian_group]
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   532
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   533
lemmas (in UP_cring) UP_r_neg =
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   534
  abelian_group.r_neg [OF UP_abelian_group]
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   535
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   536
lemmas (in UP_cring) UP_minus_zero [simp] =
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   537
  abelian_group.minus_zero [OF UP_abelian_group]
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   538
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   539
lemmas (in UP_cring) UP_minus_minus [simp] =
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   540
  abelian_group.minus_minus [OF UP_abelian_group]
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   541
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   542
lemmas (in UP_cring) UP_minus_add =
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   543
  abelian_group.minus_add [OF UP_abelian_group]
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   544
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   545
lemmas (in UP_cring) UP_r_neg2 =
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   546
  abelian_group.r_neg2 [OF UP_abelian_group]
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   547
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   548
lemmas (in UP_cring) UP_r_neg1 =
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   549
  abelian_group.r_neg1 [OF UP_abelian_group]
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   550
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   551
lemmas (in UP_cring) UP_r_distr =
14399
dc677b35e54f New lemmas about inversion of restricted functions.
ballarin
parents: 13975
diff changeset
   552
  ring.r_distr [OF UP_ring]
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   553
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   554
lemmas (in UP_cring) UP_l_null [simp] =
14399
dc677b35e54f New lemmas about inversion of restricted functions.
ballarin
parents: 13975
diff changeset
   555
  ring.l_null [OF UP_ring]
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   556
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   557
lemmas (in UP_cring) UP_r_null [simp] =
14399
dc677b35e54f New lemmas about inversion of restricted functions.
ballarin
parents: 13975
diff changeset
   558
  ring.r_null [OF UP_ring]
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   559
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   560
lemmas (in UP_cring) UP_l_minus =
14399
dc677b35e54f New lemmas about inversion of restricted functions.
ballarin
parents: 13975
diff changeset
   561
  ring.l_minus [OF UP_ring]
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   562
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   563
lemmas (in UP_cring) UP_r_minus =
14399
dc677b35e54f New lemmas about inversion of restricted functions.
ballarin
parents: 13975
diff changeset
   564
  ring.r_minus [OF UP_ring]
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   565
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   566
lemmas (in UP_cring) UP_finsum_ldistr =
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   567
  cring.finsum_ldistr [OF UP_cring]
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   568
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   569
lemmas (in UP_cring) UP_finsum_rdistr =
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   570
  cring.finsum_rdistr [OF UP_cring]
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   571
14666
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
   572
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   573
subsection {* Polynomials form an Algebra *}
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   574
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   575
lemma (in UP_cring) UP_smult_l_distr:
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   576
  "[| 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
   577
  (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
   578
  by (rule up_eqI) (simp_all add: R.l_distr)
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   579
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   580
lemma (in UP_cring) UP_smult_r_distr:
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   581
  "[| 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
   582
  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
   583
  by (rule up_eqI) (simp_all add: R.r_distr)
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   584
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   585
lemma (in UP_cring) UP_smult_assoc1:
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   586
      "[| 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
   587
      (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
   588
  by (rule up_eqI) (simp_all add: R.m_assoc)
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   589
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   590
lemma (in UP_cring) UP_smult_one [simp]:
15095
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
   591
      "p \<in> carrier P ==> \<one> \<odot>\<^bsub>P\<^esub> p = p"
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   592
  by (rule up_eqI) simp_all
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   593
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   594
lemma (in UP_cring) UP_smult_assoc2:
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   595
  "[| 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
   596
  (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
   597
  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
   598
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   599
text {*
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   600
  Instantiation of lemmas from @{term algebra}.
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   601
*}
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   602
15095
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
   603
(* TODO: this should be automated with an instantiation command. *)
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
   604
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   605
(* TODO: move to CRing.thy, really a fact missing from the locales package *)
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   606
lemma (in cring) cring:
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   607
  "cring R"
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   608
  by (fast intro: cring.intro prems)
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   609
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   610
lemma (in UP_cring) UP_algebra:
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   611
  "algebra R P"
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   612
  by (auto intro: algebraI cring UP_cring UP_smult_l_distr UP_smult_r_distr
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   613
    UP_smult_assoc1 UP_smult_assoc2)
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   614
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   615
lemmas (in UP_cring) UP_smult_l_null [simp] =
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   616
  algebra.smult_l_null [OF UP_algebra]
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   617
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   618
lemmas (in UP_cring) UP_smult_r_null [simp] =
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   619
  algebra.smult_r_null [OF UP_algebra]
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   620
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   621
lemmas (in UP_cring) UP_smult_l_minus =
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   622
  algebra.smult_l_minus [OF UP_algebra]
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   623
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   624
lemmas (in UP_cring) UP_smult_r_minus =
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   625
  algebra.smult_r_minus [OF UP_algebra]
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   626
13949
0ce528cd6f19 HOL-Algebra complete for release Isabelle2003 (modulo section headers).
ballarin
parents: 13940
diff changeset
   627
subsection {* Further lemmas involving monomials *}
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   628
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   629
lemma (in UP_cring) monom_zero [simp]:
15095
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
   630
  "monom P \<zero> n = \<zero>\<^bsub>P\<^esub>"
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   631
  by (simp add: UP_def P_def)
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   632
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   633
ML_setup {*
14590
276ef51cedbf simplified ML code for setsubgoaler;
wenzelm
parents: 14577
diff changeset
   634
  simpset_ref() := simpset() setsubgoaler asm_full_simp_tac;
276ef51cedbf simplified ML code for setsubgoaler;
wenzelm
parents: 14577
diff changeset
   635
*}
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   636
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   637
lemma (in UP_cring) monom_mult_is_smult:
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   638
  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
   639
  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
   640
proof (rule up_eqI)
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   641
  fix n
15095
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
   642
  have "coeff P (p \<otimes>\<^bsub>P\<^esub> monom P a 0) n = coeff P (a \<odot>\<^bsub>P\<^esub> p) n"
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   643
  proof (cases n)
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   644
    case 0 with R show ?thesis by (simp add: R.m_comm)
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   645
  next
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   646
    case Suc with R show ?thesis
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   647
      by (simp cong: finsum_cong add: R.r_null Pi_def)
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   648
        (simp add: m_comm)
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   649
  qed
15095
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
   650
  with R 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
   651
    by (simp add: UP_m_comm)
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   652
qed (simp_all add: R)
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   653
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   654
ML_setup {*
14590
276ef51cedbf simplified ML code for setsubgoaler;
wenzelm
parents: 14577
diff changeset
   655
  simpset_ref() := simpset() setsubgoaler asm_simp_tac;
276ef51cedbf simplified ML code for setsubgoaler;
wenzelm
parents: 14577
diff changeset
   656
*}
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   657
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   658
lemma (in UP_cring) monom_add [simp]:
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   659
  "[| 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
   660
  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
   661
  by (rule up_eqI) simp_all
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   662
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   663
ML_setup {*
14590
276ef51cedbf simplified ML code for setsubgoaler;
wenzelm
parents: 14577
diff changeset
   664
  simpset_ref() := simpset() setsubgoaler asm_full_simp_tac;
276ef51cedbf simplified ML code for setsubgoaler;
wenzelm
parents: 14577
diff changeset
   665
*}
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   666
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   667
lemma (in UP_cring) monom_one_Suc:
15095
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
   668
  "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
   669
proof (rule up_eqI)
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   670
  fix k
15095
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
   671
  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
   672
  proof (cases "k = Suc n")
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   673
    case True show ?thesis
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   674
    proof -
14666
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
   675
      from True have less_add_diff:
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
   676
        "!!i. [| n < i; i <= n + m |] ==> n + m - i < m" by arith
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   677
      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
   678
      also from True
15045
d59f7e2e18d3 Moved to new m<..<n syntax for set intervals.
nipkow
parents: 14963
diff changeset
   679
      have "... = (\<Oplus>i \<in> {..<n} \<union> {n}. coeff P (monom P \<one> n) i \<otimes>
14666
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
   680
        coeff P (monom P \<one> 1) (k - i))"
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
   681
        by (simp cong: finsum_cong add: finsum_Un_disjoint Pi_def)
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
   682
      also have "... = (\<Oplus>i \<in>  {..n}. coeff P (monom P \<one> n) i \<otimes>
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
   683
        coeff P (monom P \<one> 1) (k - i))"
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
   684
        by (simp only: ivl_disj_un_singleton)
15095
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
   685
      also from True
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
   686
      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
   687
        coeff P (monom P \<one> 1) (k - i))"
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
   688
        by (simp cong: finsum_cong add: finsum_Un_disjoint ivl_disj_int_one
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
   689
          order_less_imp_not_eq Pi_def)
15095
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
   690
      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
   691
        by (simp add: ivl_disj_un_one)
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   692
      finally show ?thesis .
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   693
    qed
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   694
  next
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   695
    case False
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   696
    note neq = False
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   697
    let ?s =
14666
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
   698
      "\<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
   699
    from neq have "coeff P (monom P \<one> (Suc n)) k = \<zero>" by simp
14666
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
   700
    also have "... = (\<Oplus>i \<in> {..k}. ?s i)"
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   701
    proof -
15095
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
   702
      have f1: "(\<Oplus>i \<in> {..<n}. ?s i) = \<zero>"
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
   703
        by (simp cong: finsum_cong add: Pi_def)
14666
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
   704
      from neq have f2: "(\<Oplus>i \<in> {n}. ?s i) = \<zero>"
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
   705
        by (simp cong: finsum_cong add: Pi_def) arith
15045
d59f7e2e18d3 Moved to new m<..<n syntax for set intervals.
nipkow
parents: 14963
diff changeset
   706
      have f3: "n < k ==> (\<Oplus>i \<in> {n<..k}. ?s i) = \<zero>"
14666
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
   707
        by (simp cong: finsum_cong add: order_less_imp_not_eq Pi_def)
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   708
      show ?thesis
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   709
      proof (cases "k < n")
14666
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
   710
        case True then show ?thesis by (simp cong: finsum_cong add: Pi_def)
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   711
      next
14666
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
   712
        case False then have n_le_k: "n <= k" by arith
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
   713
        show ?thesis
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
   714
        proof (cases "n = k")
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
   715
          case True
15045
d59f7e2e18d3 Moved to new m<..<n syntax for set intervals.
nipkow
parents: 14963
diff changeset
   716
          then have "\<zero> = (\<Oplus>i \<in> {..<n} \<union> {n}. ?s i)"
14666
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
   717
            by (simp cong: finsum_cong add: finsum_Un_disjoint
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
   718
              ivl_disj_int_singleton Pi_def)
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
   719
          also from True have "... = (\<Oplus>i \<in> {..k}. ?s i)"
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
   720
            by (simp only: ivl_disj_un_singleton)
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
   721
          finally show ?thesis .
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
   722
        next
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
   723
          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
   724
          with neq have "\<zero> = (\<Oplus>i \<in> {..<n} \<union> {n}. ?s i)"
14666
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
   725
            by (simp add: finsum_Un_disjoint f1 f2
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
   726
              ivl_disj_int_singleton Pi_def del: Un_insert_right)
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
   727
          also have "... = (\<Oplus>i \<in> {..n}. ?s i)"
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
   728
            by (simp only: ivl_disj_un_singleton)
15045
d59f7e2e18d3 Moved to new m<..<n syntax for set intervals.
nipkow
parents: 14963
diff changeset
   729
          also from n_less_k neq have "... = (\<Oplus>i \<in> {..n} \<union> {n<..k}. ?s i)"
14666
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
   730
            by (simp add: finsum_Un_disjoint f3 ivl_disj_int_one Pi_def)
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
   731
          also from n_less_k have "... = (\<Oplus>i \<in> {..k}. ?s i)"
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
   732
            by (simp only: ivl_disj_un_one)
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
   733
          finally show ?thesis .
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
   734
        qed
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   735
      qed
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   736
    qed
15095
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
   737
    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
   738
    finally show ?thesis .
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   739
  qed
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   740
qed (simp_all)
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   741
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   742
ML_setup {*
14590
276ef51cedbf simplified ML code for setsubgoaler;
wenzelm
parents: 14577
diff changeset
   743
  simpset_ref() := simpset() setsubgoaler asm_simp_tac;
276ef51cedbf simplified ML code for setsubgoaler;
wenzelm
parents: 14577
diff changeset
   744
*}
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   745
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   746
lemma (in UP_cring) monom_mult_smult:
15095
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
   747
  "[| 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
   748
  by (rule up_eqI) simp_all
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   749
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   750
lemma (in UP_cring) monom_one [simp]:
15095
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
   751
  "monom P \<one> 0 = \<one>\<^bsub>P\<^esub>"
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   752
  by (rule up_eqI) simp_all
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   753
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   754
lemma (in UP_cring) monom_one_mult:
15095
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
   755
  "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
   756
proof (induct n)
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   757
  case 0 show ?case by simp
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   758
next
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   759
  case Suc then show ?case
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   760
    by (simp only: add_Suc monom_one_Suc) (simp add: UP_m_ac)
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   761
qed
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   762
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   763
lemma (in UP_cring) monom_mult [simp]:
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   764
  assumes R: "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
   765
  shows "monom P (a \<otimes> b) (n + m) = monom P a n \<otimes>\<^bsub>P\<^esub> monom P b m"
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   766
proof -
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   767
  from R have "monom P (a \<otimes> b) (n + m) = monom P (a \<otimes> b \<otimes> \<one>) (n + m)" by simp
15095
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
   768
  also from R have "... = a \<otimes> b \<odot>\<^bsub>P\<^esub> monom P \<one> (n + m)"
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   769
    by (simp add: monom_mult_smult del: r_one)
15095
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
   770
  also have "... = a \<otimes> b \<odot>\<^bsub>P\<^esub> (monom P \<one> n \<otimes>\<^bsub>P\<^esub> monom P \<one> m)"
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   771
    by (simp only: monom_one_mult)
15095
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
   772
  also from R have "... = a \<odot>\<^bsub>P\<^esub> (b \<odot>\<^bsub>P\<^esub> (monom P \<one> n \<otimes>\<^bsub>P\<^esub> monom P \<one> m))"
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   773
    by (simp add: UP_smult_assoc1)
15095
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
   774
  also from R have "... = a \<odot>\<^bsub>P\<^esub> (b \<odot>\<^bsub>P\<^esub> (monom P \<one> m \<otimes>\<^bsub>P\<^esub> monom P \<one> n))"
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   775
    by (simp add: UP_m_comm)
15095
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
   776
  also from R have "... = a \<odot>\<^bsub>P\<^esub> ((b \<odot>\<^bsub>P\<^esub> monom P \<one> m) \<otimes>\<^bsub>P\<^esub> monom P \<one> n)"
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   777
    by (simp add: UP_smult_assoc2)
15095
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
   778
  also from R have "... = a \<odot>\<^bsub>P\<^esub> (monom P \<one> n \<otimes>\<^bsub>P\<^esub> (b \<odot>\<^bsub>P\<^esub> monom P \<one> m))"
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   779
    by (simp add: UP_m_comm)
15095
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
   780
  also from R have "... = (a \<odot>\<^bsub>P\<^esub> monom P \<one> n) \<otimes>\<^bsub>P\<^esub> (b \<odot>\<^bsub>P\<^esub> monom P \<one> m)"
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   781
    by (simp add: UP_smult_assoc2)
15095
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
   782
  also from R have "... = monom P (a \<otimes> \<one>) n \<otimes>\<^bsub>P\<^esub> monom P (b \<otimes> \<one>) m"
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   783
    by (simp add: monom_mult_smult del: r_one)
15095
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
   784
  also from R have "... = monom P a n \<otimes>\<^bsub>P\<^esub> monom P b m" by simp
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   785
  finally show ?thesis .
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   786
qed
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   787
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   788
lemma (in UP_cring) monom_a_inv [simp]:
15095
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
   789
  "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
   790
  by (rule up_eqI) simp_all
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   791
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   792
lemma (in UP_cring) monom_inj:
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   793
  "inj_on (%a. monom P a n) (carrier R)"
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   794
proof (rule inj_onI)
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   795
  fix x y
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   796
  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
   797
  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
   798
  with R show "x = y" by simp
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   799
qed
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   800
13949
0ce528cd6f19 HOL-Algebra complete for release Isabelle2003 (modulo section headers).
ballarin
parents: 13940
diff changeset
   801
subsection {* The degree function *}
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   802
14651
02b8f3bcf7fe improved notation;
wenzelm
parents: 14590
diff changeset
   803
constdefs (structure R)
15095
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
   804
  deg :: "[('a, 'm) ring_scheme, nat => 'a] => nat"
14651
02b8f3bcf7fe improved notation;
wenzelm
parents: 14590
diff changeset
   805
  "deg R p == LEAST n. bound \<zero> n (coeff (UP R) p)"
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   806
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   807
lemma (in UP_cring) deg_aboveI:
14666
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
   808
  "[| (!!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
   809
  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
   810
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   811
(*
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   812
lemma coeff_bound_ex: "EX n. bound n (coeff p)"
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   813
proof -
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   814
  have "(%n. coeff p n) : UP" by (simp add: coeff_def Rep_UP)
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   815
  then obtain n where "bound n (coeff p)" by (unfold UP_def) fast
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   816
  then show ?thesis ..
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   817
qed
14666
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
   818
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   819
lemma bound_coeff_obtain:
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   820
  assumes prem: "(!!n. bound n (coeff p) ==> P)" shows "P"
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   821
proof -
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   822
  have "(%n. coeff p n) : UP" by (simp add: coeff_def Rep_UP)
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   823
  then obtain n where "bound n (coeff p)" by (unfold UP_def) fast
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   824
  with prem show P .
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   825
qed
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   826
*)
15095
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
   827
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   828
lemma (in UP_cring) deg_aboveD:
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   829
  "[| deg R p < m; p \<in> carrier P |] ==> coeff P p m = \<zero>"
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   830
proof -
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   831
  assume R: "p \<in> carrier P" and "deg R p < m"
14666
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
   832
  from R obtain n where "bound \<zero> n (coeff P p)"
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   833
    by (auto simp add: UP_def P_def)
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   834
  then have "bound \<zero> (deg R p) (coeff P p)"
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   835
    by (auto simp: deg_def P_def dest: LeastI)
14666
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
   836
  then show ?thesis ..
13940
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
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   839
lemma (in UP_cring) deg_belowI:
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   840
  assumes non_zero: "n ~= 0 ==> coeff P p n ~= \<zero>"
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   841
    and R: "p \<in> carrier P"
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   842
  shows "n <= deg R p"
14666
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
   843
-- {* 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
   844
   @{thm [source] deg_aboveD} *}
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   845
proof (cases "n=0")
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   846
  case True then show ?thesis by simp
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   847
next
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   848
  case False then have "coeff P p n ~= \<zero>" by (rule non_zero)
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   849
  then have "~ deg R p < n" by (fast dest: deg_aboveD intro: R)
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   850
  then show ?thesis by arith
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   851
qed
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   852
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   853
lemma (in UP_cring) lcoeff_nonzero_deg:
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   854
  assumes deg: "deg R p ~= 0" and R: "p \<in> carrier P"
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   855
  shows "coeff P p (deg R p) ~= \<zero>"
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   856
proof -
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   857
  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
   858
  proof -
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   859
    have minus: "!!(n::nat) m. n ~= 0 ==> (n - Suc 0 < m) = (n <= m)"
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   860
      by arith
15095
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
   861
(* TODO: why does simplification below not work with "1" *)
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   862
    from deg have "deg R p - 1 < (LEAST n. bound \<zero> n (coeff P p))"
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   863
      by (unfold deg_def P_def) arith
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   864
    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
   865
    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
   866
      by (unfold bound_def) fast
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   867
    then have "EX m. deg R p <= m & coeff P p m ~= \<zero>" by (simp add: deg minus)
14666
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
   868
    then show ?thesis by auto
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   869
  qed
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   870
  with deg_belowI R have "deg R p = m" by fastsimp
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   871
  with m_coeff show ?thesis by simp
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   872
qed
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   873
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   874
lemma (in UP_cring) lcoeff_nonzero_nonzero:
15095
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
   875
  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
   876
  shows "coeff P p 0 ~= \<zero>"
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   877
proof -
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   878
  have "EX m. coeff P p m ~= \<zero>"
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   879
  proof (rule classical)
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   880
    assume "~ ?thesis"
15095
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
   881
    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
   882
    with nonzero show ?thesis by contradiction
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   883
  qed
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   884
  then obtain m where coeff: "coeff P p m ~= \<zero>" ..
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   885
  then have "m <= deg R p" by (rule deg_belowI)
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   886
  then have "m = 0" by (simp add: deg)
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   887
  with coeff show ?thesis by simp
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
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   890
lemma (in UP_cring) lcoeff_nonzero:
15095
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
   891
  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
   892
  shows "coeff P p (deg R p) ~= \<zero>"
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   893
proof (cases "deg R p = 0")
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   894
  case True with neq R show ?thesis by (simp add: lcoeff_nonzero_nonzero)
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   895
next
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   896
  case False with neq R show ?thesis by (simp add: lcoeff_nonzero_deg)
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   897
qed
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   898
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   899
lemma (in UP_cring) deg_eqI:
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   900
  "[| !!m. n < m ==> coeff P p m = \<zero>;
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   901
      !!n. n ~= 0 ==> coeff P p n ~= \<zero>; p \<in> carrier P |] ==> deg R p = n"
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   902
by (fast intro: le_anti_sym deg_aboveI deg_belowI)
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   903
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   904
(* Degree and polynomial operations *)
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   905
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   906
lemma (in UP_cring) deg_add [simp]:
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   907
  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
   908
  shows "deg R (p \<oplus>\<^bsub>P\<^esub> q) <= max (deg R p) (deg R q)"
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   909
proof (cases "deg R p <= deg R q")
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   910
  case True show ?thesis
14666
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
   911
    by (rule deg_aboveI) (simp_all add: True R deg_aboveD)
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   912
next
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   913
  case False show ?thesis
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   914
    by (rule deg_aboveI) (simp_all add: False R deg_aboveD)
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   915
qed
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   916
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   917
lemma (in UP_cring) deg_monom_le:
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   918
  "a \<in> carrier R ==> deg R (monom P a n) <= n"
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   919
  by (intro deg_aboveI) simp_all
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   920
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   921
lemma (in UP_cring) deg_monom [simp]:
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   922
  "[| a ~= \<zero>; a \<in> carrier R |] ==> deg R (monom P a n) = n"
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   923
  by (fastsimp intro: le_anti_sym deg_aboveI deg_belowI)
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   924
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   925
lemma (in UP_cring) deg_const [simp]:
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   926
  assumes R: "a \<in> carrier R" shows "deg R (monom P a 0) = 0"
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   927
proof (rule le_anti_sym)
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   928
  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
   929
next
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   930
  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
   931
qed
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   932
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   933
lemma (in UP_cring) deg_zero [simp]:
15095
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
   934
  "deg R \<zero>\<^bsub>P\<^esub> = 0"
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   935
proof (rule le_anti_sym)
15095
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
   936
  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
   937
next
15095
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
   938
  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
   939
qed
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   941
lemma (in UP_cring) deg_one [simp]:
15095
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
   942
  "deg R \<one>\<^bsub>P\<^esub> = 0"
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   943
proof (rule le_anti_sym)
15095
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
   944
  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
   945
next
15095
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
   946
  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
   947
qed
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   948
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   949
lemma (in UP_cring) deg_uminus [simp]:
15095
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
   950
  assumes R: "p \<in> carrier P" shows "deg R (\<ominus>\<^bsub>P\<^esub> p) = deg R p"
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   951
proof (rule le_anti_sym)
15095
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
   952
  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
   953
next
15095
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
   954
  show "deg R p <= deg R (\<ominus>\<^bsub>P\<^esub> p)"
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   955
    by (simp add: deg_belowI lcoeff_nonzero_deg
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   956
      inj_on_iff [OF a_inv_inj, of _ "\<zero>", simplified] R)
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   957
qed
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   958
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   959
lemma (in UP_domain) deg_smult_ring:
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   960
  "[| 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
   961
  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
   962
  by (cases "a = \<zero>") (simp add: deg_aboveI deg_aboveD)+
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   963
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   964
lemma (in UP_domain) deg_smult [simp]:
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   965
  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
   966
  shows "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
   967
proof (rule le_anti_sym)
15095
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
   968
  show "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
   969
    by (rule deg_smult_ring)
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   970
next
15095
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
   971
  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
   972
  proof (cases "a = \<zero>")
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   973
  qed (simp, simp add: deg_belowI lcoeff_nonzero_deg integral_iff R)
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   974
qed
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   975
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   976
lemma (in UP_cring) deg_mult_cring:
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   977
  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
   978
  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
   979
proof (rule deg_aboveI)
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   980
  fix m
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   981
  assume boundm: "deg R p + deg R q < m"
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   982
  {
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   983
    fix k i
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   984
    assume boundk: "deg R p + deg R q < k"
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   985
    then have "coeff P p i \<otimes> coeff P q (k - i) = \<zero>"
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   986
    proof (cases "deg R p < i")
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   987
      case True then show ?thesis by (simp add: deg_aboveD R)
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   988
    next
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   989
      case False with boundk have "deg R q < k - i" by arith
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   990
      then show ?thesis by (simp add: deg_aboveD R)
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   991
    qed
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   992
  }
15095
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
   993
  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
   994
qed (simp add: R)
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   995
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   996
ML_setup {*
14590
276ef51cedbf simplified ML code for setsubgoaler;
wenzelm
parents: 14577
diff changeset
   997
  simpset_ref() := simpset() setsubgoaler asm_full_simp_tac;
276ef51cedbf simplified ML code for setsubgoaler;
wenzelm
parents: 14577
diff changeset
   998
*}
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
   999
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1000
lemma (in UP_domain) deg_mult [simp]:
15095
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
  1001
  "[| 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
  1002
  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
  1003
proof (rule le_anti_sym)
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1004
  assume "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
  1005
  show "deg R (p \<otimes>\<^bsub>P\<^esub> q) <= deg R p + deg R q" by (rule deg_mult_cring)
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1006
next
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1007
  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
  1008
  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
  1009
  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
  1010
  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
  1011
  proof (rule deg_belowI, simp add: R)
15095
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
  1012
    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
  1013
      = (\<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
  1014
      by (simp only: ivl_disj_un_one)
15095
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
  1015
    also have "... = (\<Oplus>i \<in> {deg R p .. deg R p + deg R q}. ?s i)"
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1016
      by (simp cong: finsum_cong add: finsum_Un_disjoint ivl_disj_int_one
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1017
        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
  1018
    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
  1019
      by (simp only: ivl_disj_un_singleton)
14666
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
  1020
    also have "... = coeff P p (deg R p) \<otimes> coeff P q (deg R q)"
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1021
      by (simp cong: finsum_cong add: finsum_Un_disjoint
14666
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
  1022
        ivl_disj_int_singleton deg_aboveD R Pi_def)
15095
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
  1023
    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
  1024
      = 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
  1025
    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
  1026
      by (simp add: integral_iff lcoeff_nonzero R)
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1027
    qed (simp add: R)
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1028
  qed
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1029
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1030
lemma (in UP_cring) coeff_finsum:
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1031
  assumes fin: "finite A"
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1032
  shows "p \<in> A -> carrier P ==>
15095
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
  1033
    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
  1034
  using fin by induct (auto simp: Pi_def)
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1035
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1036
ML_setup {*
14590
276ef51cedbf simplified ML code for setsubgoaler;
wenzelm
parents: 14577
diff changeset
  1037
  simpset_ref() := simpset() setsubgoaler asm_full_simp_tac;
276ef51cedbf simplified ML code for setsubgoaler;
wenzelm
parents: 14577
diff changeset
  1038
*}
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1039
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1040
lemma (in UP_cring) up_repr:
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1041
  assumes R: "p \<in> carrier P"
15095
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
  1042
  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
  1043
proof (rule up_eqI)
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1044
  let ?s = "(%i. monom P (coeff P p i) i)"
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1045
  fix k
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1046
  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
  1047
    by simp
15095
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
  1048
  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
  1049
  proof (cases "k <= deg R p")
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1050
    case True
15095
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
  1051
    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
  1052
          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
  1053
      by (simp only: ivl_disj_un_one)
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1054
    also from True
15095
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
  1055
    have "... = coeff P (\<Oplus>\<^bsub>P\<^esub> i \<in> {..k}. ?s i) k"
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1056
      by (simp cong: finsum_cong add: finsum_Un_disjoint
14666
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
  1057
        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
  1058
    also
15095
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
  1059
    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
  1060
      by (simp only: ivl_disj_un_singleton)
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1061
    also have "... = coeff P p k"
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1062
      by (simp cong: finsum_cong add: setsum_Un_disjoint
14666
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
  1063
        ivl_disj_int_singleton coeff_finsum deg_aboveD R RR Pi_def)
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1064
    finally show ?thesis .
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1065
  next
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1066
    case False
15095
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
  1067
    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
  1068
          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
  1069
      by (simp only: ivl_disj_un_singleton)
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1070
    also from False have "... = coeff P p k"
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1071
      by (simp cong: finsum_cong add: setsum_Un_disjoint ivl_disj_int_singleton
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1072
        coeff_finsum deg_aboveD R Pi_def)
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1073
    finally show ?thesis .
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1074
  qed
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1075
qed (simp_all add: R Pi_def)
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1076
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1077
lemma (in UP_cring) up_repr_le:
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1078
  "[| 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
  1079
  (\<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
  1080
proof -
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1081
  let ?s = "(%i. monom P (coeff P p i) i)"
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1082
  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
  1083
  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
  1084
    by (simp only: ivl_disj_un_one)
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1085
  also have "... = finsum P ?s {..deg R p}"
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1086
    by (simp cong: UP_finsum_cong add: UP_finsum_Un_disjoint ivl_disj_int_one
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1087
      deg_aboveD R Pi_def)
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1088
  also have "... = p" by (rule up_repr)
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1089
  finally show ?thesis .
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1090
qed
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1091
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1092
ML_setup {*
14590
276ef51cedbf simplified ML code for setsubgoaler;
wenzelm
parents: 14577
diff changeset
  1093
  simpset_ref() := simpset() setsubgoaler asm_simp_tac;
276ef51cedbf simplified ML code for setsubgoaler;
wenzelm
parents: 14577
diff changeset
  1094
*}
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1095
13949
0ce528cd6f19 HOL-Algebra complete for release Isabelle2003 (modulo section headers).
ballarin
parents: 13940
diff changeset
  1096
subsection {* Polynomials over an integral domain form an integral domain *}
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1097
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1098
lemma domainI:
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1099
  assumes cring: "cring R"
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1100
    and one_not_zero: "one R ~= zero R"
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1101
    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
  1102
      b \<in> carrier R |] ==> a = zero R | b = zero R"
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1103
  shows "domain R"
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1104
  by (auto intro!: domain.intro domain_axioms.intro cring.axioms prems
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1105
    del: disjCI)
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1106
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1107
lemma (in UP_domain) UP_one_not_zero:
15095
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
  1108
  "\<one>\<^bsub>P\<^esub> ~= \<zero>\<^bsub>P\<^esub>"
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1109
proof
15095
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
  1110
  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
  1111
  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
  1112
  hence "\<one> = \<zero>" by simp
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1113
  with one_not_zero show "False" by contradiction
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1114
qed
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1115
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1116
lemma (in UP_domain) UP_integral:
15095
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
  1117
  "[| 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
  1118
proof -
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1119
  fix p q
15095
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
  1120
  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
  1121
  show "p = \<zero>\<^bsub>P\<^esub> | q = \<zero>\<^bsub>P\<^esub>"
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1122
  proof (rule classical)
15095
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
  1123
    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
  1124
    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
  1125
    also from pq have "... = 0" by simp
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1126
    finally have "deg R p + deg R q = 0" .
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1127
    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
  1128
    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
  1129
      by (simp only: up_repr_le)
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1130
    also from R have "... = monom P (coeff P p 0) 0" by simp
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1131
    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
  1132
    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
  1133
      by (simp only: up_repr_le)
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1134
    also from R have "... = monom P (coeff P q 0) 0" by simp
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1135
    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
  1136
    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
  1137
    also from pq have "... = \<zero>" by simp
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1138
    finally have "coeff P p 0 \<otimes> coeff P q 0 = \<zero>" .
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1139
    with R have "coeff P p 0 = \<zero> | coeff P q 0 = \<zero>"
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1140
      by (simp add: R.integral_iff)
15095
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
  1141
    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
  1142
  qed
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1143
qed
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1144
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1145
theorem (in UP_domain) UP_domain:
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1146
  "domain P"
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1147
  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
  1148
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1149
text {*
15095
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
  1150
  Instantiation of theorems from @{term domain}.
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1151
*}
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1152
15095
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
  1153
(* TODO: this should be automated with an instantiation command. *)
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
  1154
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1155
lemmas (in UP_domain) UP_zero_not_one [simp] =
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1156
  domain.zero_not_one [OF UP_domain]
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1157
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1158
lemmas (in UP_domain) UP_integral_iff =
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1159
  domain.integral_iff [OF UP_domain]
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1160
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1161
lemmas (in UP_domain) UP_m_lcancel =
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1162
  domain.m_lcancel [OF UP_domain]
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1163
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1164
lemmas (in UP_domain) UP_m_rcancel =
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1165
  domain.m_rcancel [OF UP_domain]
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1166
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1167
lemma (in UP_domain) smult_integral:
15095
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
  1168
  "[| a \<odot>\<^bsub>P\<^esub> p = \<zero>\<^bsub>P\<^esub>; a \<in> carrier R; p \<in> carrier P |] ==> a = \<zero> | p = \<zero>\<^bsub>P\<^esub>"
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1169
  by (simp add: monom_mult_is_smult [THEN sym] UP_integral_iff
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1170
    inj_on_iff [OF monom_inj, of _ "\<zero>", simplified])
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1171
14666
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
  1172
13949
0ce528cd6f19 HOL-Algebra complete for release Isabelle2003 (modulo section headers).
ballarin
parents: 13940
diff changeset
  1173
subsection {* Evaluation Homomorphism and Universal Property*}
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1174
14666
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
  1175
(* alternative congruence rule (possibly more efficient)
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
  1176
lemma (in abelian_monoid) finsum_cong2:
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
  1177
  "[| !!i. i \<in> A ==> f i \<in> carrier G = True; A = B;
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
  1178
  !!i. i \<in> B ==> f i = g i |] ==> finsum G f A = finsum G g B"
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
  1179
  sorry*)
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
  1180
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1181
ML_setup {*
14590
276ef51cedbf simplified ML code for setsubgoaler;
wenzelm
parents: 14577
diff changeset
  1182
  simpset_ref() := simpset() setsubgoaler asm_full_simp_tac;
276ef51cedbf simplified ML code for setsubgoaler;
wenzelm
parents: 14577
diff changeset
  1183
*}
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1184
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1185
theorem (in cring) diagonal_sum:
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1186
  "[| f \<in> {..n + m::nat} -> carrier R; g \<in> {..n + m} -> carrier R |] ==>
14666
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
  1187
  (\<Oplus>k \<in> {..n + m}. \<Oplus>i \<in> {..k}. f i \<otimes> g (k - i)) =
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
  1188
  (\<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
  1189
proof -
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1190
  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
  1191
  {
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1192
    fix j
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1193
    have "j <= n + m ==>
14666
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
  1194
      (\<Oplus>k \<in> {..j}. \<Oplus>i \<in> {..k}. f i \<otimes> g (k - i)) =
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
  1195
      (\<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
  1196
    proof (induct j)
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1197
      case 0 from Rf Rg show ?case by (simp add: Pi_def)
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1198
    next
14666
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
  1199
      case (Suc j)
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1200
      have R6: "!!i k. [| k <= j; i <= Suc j - k |] ==> g i \<in> carrier R"
14666
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
  1201
        using Suc by (auto intro!: funcset_mem [OF Rg]) arith
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1202
      have R8: "!!i k. [| k <= Suc j; i <= k |] ==> g (k - i) \<in> carrier R"
14666
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
  1203
        using Suc by (auto intro!: funcset_mem [OF Rg]) arith
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1204
      have R9: "!!i k. [| k <= Suc j |] ==> f k \<in> carrier R"
14666
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
  1205
        using Suc by (auto intro!: funcset_mem [OF Rf])
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1206
      have R10: "!!i k. [| k <= Suc j; i <= Suc j - k |] ==> g i \<in> carrier R"
14666
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
  1207
        using Suc by (auto intro!: funcset_mem [OF Rg]) arith
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1208
      have R11: "g 0 \<in> carrier R"
14666
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
  1209
        using Suc by (auto intro!: funcset_mem [OF Rg])
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1210
      from Suc show ?case
14666
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
  1211
        by (simp cong: finsum_cong add: Suc_diff_le a_ac
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
  1212
          Pi_def R6 R8 R9 R10 R11)
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1213
    qed
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1214
  }
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1215
  then show ?thesis by fast
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1216
qed
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1217
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1218
lemma (in abelian_monoid) boundD_carrier:
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1219
  "[| bound \<zero> n f; n < m |] ==> f m \<in> carrier G"
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1220
  by auto
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1221
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1222
theorem (in cring) cauchy_product:
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1223
  assumes bf: "bound \<zero> n f" and bg: "bound \<zero> m g"
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1224
    and Rf: "f \<in> {..n} -> carrier R" and Rg: "g \<in> {..m} -> carrier R"
14666
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
  1225
  shows "(\<Oplus>k \<in> {..n + m}. \<Oplus>i \<in> {..k}. f i \<otimes> g (k - i)) =
15095
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
  1226
    (\<Oplus>i \<in> {..n}. f i) \<otimes> (\<Oplus>i \<in> {..m}. g i)"       (* State revese direction? *)
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1227
proof -
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1228
  have f: "!!x. f x \<in> carrier R"
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1229
  proof -
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1230
    fix x
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1231
    show "f x \<in> carrier R"
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1232
      using Rf bf boundD_carrier by (cases "x <= n") (auto simp: Pi_def)
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1233
  qed
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1234
  have g: "!!x. g x \<in> carrier R"
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1235
  proof -
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1236
    fix x
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1237
    show "g x \<in> carrier R"
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1238
      using Rg bg boundD_carrier by (cases "x <= m") (auto simp: Pi_def)
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1239
  qed
14666
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
  1240
  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
  1241
      (\<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
  1242
    by (simp add: diagonal_sum Pi_def)
15045
d59f7e2e18d3 Moved to new m<..<n syntax for set intervals.
nipkow
parents: 14963
diff changeset
  1243
  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
  1244
    by (simp only: ivl_disj_un_one)
14666
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
  1245
  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
  1246
    by (simp cong: finsum_cong
14666
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
  1247
      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
  1248
  also from f g
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
  1249
  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
  1250
    by (simp cong: finsum_cong add: ivl_disj_un_one le_add_diff Pi_def)
14666
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
  1251
  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
  1252
    by (simp cong: finsum_cong
14666
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
  1253
      add: bound.bound [OF bg] finsum_Un_disjoint ivl_disj_int_one Pi_def)
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
  1254
  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
  1255
    by (simp add: finsum_ldistr diagonal_sum Pi_def,
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1256
      simp cong: finsum_cong add: finsum_rdistr Pi_def)
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1257
  finally show ?thesis .
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1258
qed
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1259
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1260
lemma (in UP_cring) const_ring_hom:
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1261
  "(%a. monom P a 0) \<in> ring_hom R P"
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1262
  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
  1263
14651
02b8f3bcf7fe improved notation;
wenzelm
parents: 14590
diff changeset
  1264
constdefs (structure S)
15095
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
  1265
  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
  1266
           'a => 'b, 'b, nat => 'a] => 'b"
14651
02b8f3bcf7fe improved notation;
wenzelm
parents: 14590
diff changeset
  1267
  "eval R S phi s == \<lambda>p \<in> carrier (UP R).
15095
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
  1268
    \<Oplus>i \<in> {..deg R p}. phi (coeff (UP R) p i) \<otimes> s (^) i"
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
  1269
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
  1270
locale UP_univ_prop = ring_hom_cring R S + UP_cring R
14666
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
  1271
15095
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
  1272
lemma (in UP) eval_on_carrier:
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
  1273
  includes struct S
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
  1274
  shows  "p \<in> carrier P ==>
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1275
    eval R S phi s p =
15095
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
  1276
    (\<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
  1277
  by (unfold eval_def, fold P_def) simp
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1278
15095
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
  1279
lemma (in UP) eval_extensional:
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1280
  "eval R S phi s \<in> extensional (carrier P)"
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1281
  by (unfold eval_def, fold P_def) simp
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1282
15095
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
  1283
theorem (in UP_univ_prop) eval_ring_hom:
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1284
  "s \<in> carrier S ==> eval R S h s \<in> ring_hom P S"
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1285
proof (rule ring_hom_memI)
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1286
  fix p
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1287
  assume RS: "p \<in> carrier P" "s \<in> carrier S"
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1288
  then show "eval R S h s p \<in> carrier S"
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1289
    by (simp only: eval_on_carrier) (simp add: Pi_def)
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1290
next
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1291
  fix p q
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1292
  assume RS: "p \<in> carrier P" "q \<in> carrier P" "s \<in> carrier S"
15095
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
  1293
  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"
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1294
  proof (simp only: eval_on_carrier UP_mult_closed)
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1295
    from RS have
15095
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
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) =
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
  1297
      (\<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}.
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
  1298
        h (coeff P (p \<otimes>\<^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
  1299
      by (simp cong: finsum_cong
14666
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
  1300
        add: deg_aboveD finsum_Un_disjoint ivl_disj_int_one Pi_def
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
  1301
        del: coeff_mult)
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1302
    also from RS have "... =
15095
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
  1303
      (\<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)"
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1304
      by (simp only: ivl_disj_un_one deg_mult_cring)
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1305
    also from RS have "... =
15095
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
  1306
      (\<Oplus>\<^bsub>S\<^esub> i \<in> {..deg R p + deg R q}.
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
  1307
         \<Oplus>\<^bsub>S\<^esub> k \<in> {..i}.
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
  1308
           h (coeff P p k) \<otimes>\<^bsub>S\<^esub> h (coeff P q (i - k)) \<otimes>\<^bsub>S\<^esub>
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
  1309
           (s (^)\<^bsub>S\<^esub> k \<otimes>\<^bsub>S\<^esub> s (^)\<^bsub>S\<^esub> (i - k)))"
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1310
      by (simp cong: finsum_cong add: nat_pow_mult Pi_def
14666
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
  1311
        S.m_ac S.finsum_rdistr)
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1312
    also from RS have "... =
15095
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
  1313
      (\<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>
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
  1314
      (\<Oplus>\<^bsub>S\<^esub> i\<in>{..deg R q}. h (coeff P q i) \<otimes>\<^bsub>S\<^esub> s (^)\<^bsub>S\<^esub> i)"
14666
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
  1315
      by (simp add: S.cauchy_product [THEN sym] bound.intro deg_aboveD S.m_ac
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
  1316
        Pi_def)
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1317
    finally show
15095
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
  1318
      "(\<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) =
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
  1319
      (\<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>
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
  1320
      (\<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
  1321
  qed
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1322
next
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1323
  fix p q
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1324
  assume RS: "p \<in> carrier P" "q \<in> carrier P" "s \<in> carrier S"
15095
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
  1325
  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"
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1326
  proof (simp only: eval_on_carrier UP_a_closed)
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1327
    from RS 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>{..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
  1329
      (\<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
  1330
        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
  1331
      by (simp cong: finsum_cong
14666
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
  1332
        add: deg_aboveD finsum_Un_disjoint ivl_disj_int_one Pi_def
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
  1333
        del: coeff_add)
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1334
    also from RS have "... =
15095
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
  1335
        (\<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
  1336
          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
  1337
      by (simp add: ivl_disj_un_one)
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1338
    also from RS have "... =
15095
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
  1339
      (\<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
  1340
      (\<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)"
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1341
      by (simp cong: finsum_cong
14666
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
  1342
        add: l_distr deg_aboveD finsum_Un_disjoint ivl_disj_int_one Pi_def)
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1343
    also have "... =
15095
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
  1344
        (\<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
  1345
          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
  1346
        (\<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
  1347
          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
  1348
      by (simp only: ivl_disj_un_one le_maxI1 le_maxI2)
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1349
    also from RS have "... =
15095
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
  1350
      (\<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
  1351
      (\<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
  1352
      by (simp cong: finsum_cong
14666
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
  1353
        add: deg_aboveD finsum_Un_disjoint ivl_disj_int_one Pi_def)
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1354
    finally show
15095
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
  1355
      "(\<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
  1356
      (\<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
  1357
      (\<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
  1358
  qed
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
  assume S: "s \<in> carrier S"
15095
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
  1361
  then 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
  1362
    by (simp only: eval_on_carrier UP_one_closed) simp
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1363
qed
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1364
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1365
text {* Instantiation of ring homomorphism lemmas. *}
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1366
15095
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
  1367
(* TODO: again, automate with instantiation command *)
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
  1368
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
  1369
lemma (in UP_univ_prop) ring_hom_cring_P_S:
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1370
  "s \<in> carrier S ==> ring_hom_cring P S (eval R S h s)"
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1371
  by (fast intro!: ring_hom_cring.intro UP_cring cring.axioms prems
15095
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
  1372
    intro: ring_hom_cring_axioms.intro eval_ring_hom)
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1373
15095
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
  1374
(*
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
  1375
lemma (in UP_univ_prop) UP_hom_closed [intro, simp]:
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1376
  "[| s \<in> carrier S; p \<in> carrier P |] ==> eval R S h s p \<in> carrier S"
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1377
  by (rule ring_hom_cring.hom_closed [OF ring_hom_cring_P_S])
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1378
15095
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
  1379
lemma (in UP_univ_prop) UP_hom_mult [simp]:
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1380
  "[| s \<in> carrier S; 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
  1381
  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"
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1382
  by (rule ring_hom_cring.hom_mult [OF ring_hom_cring_P_S])
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1383
15095
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
  1384
lemma (in UP_univ_prop) UP_hom_add [simp]:
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1385
  "[| s \<in> carrier S; 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
  1386
  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"
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1387
  by (rule ring_hom_cring.hom_add [OF ring_hom_cring_P_S])
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1388
15095
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
  1389
lemma (in UP_univ_prop) UP_hom_one [simp]:
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
  1390
  "s \<in> carrier S ==> 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
  1391
  by (rule ring_hom_cring.hom_one [OF ring_hom_cring_P_S])
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1392
15095
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
  1393
lemma (in UP_univ_prop) UP_hom_zero [simp]:
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
  1394
  "s \<in> carrier S ==> eval R S h s \<zero>\<^bsub>P\<^esub> = \<zero>\<^bsub>S\<^esub>"
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1395
  by (rule ring_hom_cring.hom_zero [OF ring_hom_cring_P_S])
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1396
15095
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
  1397
lemma (in UP_univ_prop) UP_hom_a_inv [simp]:
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1398
  "[| s \<in> carrier S; p \<in> carrier P |] ==>
15095
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
  1399
  (eval R S h s) (\<ominus>\<^bsub>P\<^esub> p) = \<ominus>\<^bsub>S\<^esub> (eval R S h s) p"
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1400
  by (rule ring_hom_cring.hom_a_inv [OF ring_hom_cring_P_S])
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1401
15095
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
  1402
lemma (in UP_univ_prop) UP_hom_finsum [simp]:
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1403
  "[| s \<in> carrier S; finite A; f \<in> A -> carrier P |] ==>
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1404
  (eval R S h s) (finsum P f A) = finsum S (eval R S h s o f) A"
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1405
  by (rule ring_hom_cring.hom_finsum [OF ring_hom_cring_P_S])
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1406
15095
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
  1407
lemma (in UP_univ_prop) UP_hom_finprod [simp]:
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1408
  "[| s \<in> carrier S; finite A; f \<in> A -> carrier P |] ==>
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1409
  (eval R S h s) (finprod P f A) = finprod S (eval R S h s o f) A"
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1410
  by (rule ring_hom_cring.hom_finprod [OF ring_hom_cring_P_S])
15095
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
  1411
*)
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1412
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1413
text {* Further properties of the evaluation homomorphism. *}
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1414
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1415
(* The following lemma could be proved in UP\_cring with the additional
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1416
   assumption that h is closed. *)
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1417
15095
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
  1418
lemma (in UP_univ_prop) eval_const:
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1419
  "[| 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
  1420
  by (simp only: eval_on_carrier monom_closed) simp
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1421
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1422
text {* The following proof is complicated by the fact that in arbitrary
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1423
  rings one might have @{term "one R = zero R"}. *}
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1424
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1425
(* TODO: simplify by cases "one R = zero R" *)
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1426
15095
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
  1427
lemma (in UP_univ_prop) eval_monom1:
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1428
  "s \<in> carrier S ==> eval R S h s (monom P \<one> 1) = s"
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1429
proof (simp only: eval_on_carrier monom_closed R.one_closed)
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1430
  assume S: "s \<in> carrier S"
14666
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
  1431
  then have
15095
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
  1432
    "(\<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
  1433
    (\<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
  1434
      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
  1435
    by (simp cong: finsum_cong del: coeff_monom
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1436
      add: deg_aboveD finsum_Un_disjoint ivl_disj_int_one Pi_def)
14666
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
  1437
  also have "... =
15095
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
  1438
    (\<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
  1439
    by (simp only: ivl_disj_un_one deg_monom_le R.one_closed)
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1440
  also have "... = s"
15095
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
  1441
  proof (cases "s = \<zero>\<^bsub>S\<^esub>")
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1442
    case True then show ?thesis by (simp add: Pi_def)
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1443
  next
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1444
    case False with S show ?thesis by (simp add: Pi_def)
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1445
  qed
15095
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
  1446
  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
  1447
    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
  1448
qed
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1449
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1450
lemma (in UP_cring) monom_pow:
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1451
  assumes R: "a \<in> carrier R"
15095
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
  1452
  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
  1453
proof (induct m)
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1454
  case 0 from R show ?case by simp
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1455
next
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1456
  case Suc with R show ?case
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1457
    by (simp del: monom_mult add: monom_mult [THEN sym] add_commute)
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1458
qed
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1459
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1460
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
  1461
  "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
  1462
  by (induct n) simp_all
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1463
15095
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
  1464
lemma (in UP_univ_prop) eval_monom:
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1465
  "[| s \<in> carrier S; r \<in> carrier R |] ==>
15095
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
  1466
  eval R S h s (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
  1467
proof -
15095
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
  1468
  assume S: "s \<in> carrier S" and R: "r \<in> carrier R"
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
  1469
  from R S have "eval R S h s (monom P r n) =
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
  1470
    eval R S h s (monom P r 0 \<otimes>\<^bsub>P\<^esub> (monom P \<one> 1) (^)\<^bsub>P\<^esub> n)"
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
  1471
    by (simp del: monom_mult (* eval.hom_mult eval.hom_pow, delayed inst! *)
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1472
      add: monom_mult [THEN sym] monom_pow)
15095
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
  1473
  also
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
  1474
  from ring_hom_cring_P_S [OF S] instantiate eval: ring_hom_cring
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
  1475
  from R S eval_monom1 have "... = h r \<otimes>\<^bsub>S\<^esub> s (^)\<^bsub>S\<^esub> n"
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1476
    by (simp add: eval_const)
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1477
  finally show ?thesis .
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1478
qed
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1479
15095
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
  1480
lemma (in UP_univ_prop) eval_smult:
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1481
  "[| s \<in> carrier S; r \<in> carrier R; p \<in> carrier P |] ==>
15095
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
  1482
  eval R S h s (r \<odot>\<^bsub>P\<^esub> p) = h r \<otimes>\<^bsub>S\<^esub> eval R S h s p"
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
  1483
proof -
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
  1484
  assume S: "s \<in> carrier S" and R: "r \<in> carrier R" and P: "p \<in> carrier P"
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
  1485
  from ring_hom_cring_P_S [OF S] instantiate eval: ring_hom_cring
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
  1486
  from S R P show ?thesis
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
  1487
    by (simp add: monom_mult_is_smult [THEN sym] eval_const)
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
  1488
qed
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1489
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1490
lemma ring_hom_cringI:
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1491
  assumes "cring R"
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1492
    and "cring S"
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1493
    and "h \<in> ring_hom R S"
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1494
  shows "ring_hom_cring R S h"
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1495
  by (fast intro: ring_hom_cring.intro ring_hom_cring_axioms.intro
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1496
    cring.axioms prems)
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1497
15095
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
  1498
lemma (in UP_univ_prop) UP_hom_unique:
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1499
  assumes Phi: "Phi \<in> ring_hom P S" "Phi (monom P \<one> (Suc 0)) = s"
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1500
      "!!r. r \<in> carrier R ==> Phi (monom P r 0) = h r"
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1501
    and Psi: "Psi \<in> ring_hom P S" "Psi (monom P \<one> (Suc 0)) = s"
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1502
      "!!r. r \<in> carrier R ==> Psi (monom P r 0) = h r"
15095
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
  1503
    and S: "s \<in> carrier S" and P: "p \<in> carrier P"
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1504
  shows "Phi p = Psi p"
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1505
proof -
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1506
  have Phi_hom: "ring_hom_cring P S Phi"
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1507
    by (auto intro: ring_hom_cringI UP_cring S.cring Phi)
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1508
  have Psi_hom: "ring_hom_cring P S Psi"
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1509
    by (auto intro: ring_hom_cringI UP_cring S.cring Psi)
15095
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
  1510
  have "Phi p =
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
  1511
      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)"
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
  1512
    by (simp add: up_repr P S monom_mult [THEN sym] monom_pow del: monom_mult)
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
  1513
  also 
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
  1514
    from Phi_hom instantiate Phi: ring_hom_cring
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
  1515
    from Psi_hom instantiate Psi: ring_hom_cring
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
  1516
    have "... =
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
  1517
      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)"
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
  1518
    by (simp add: Phi Psi P S Pi_def comp_def)
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
  1519
(* Without instantiate, the following command would have been necessary.
14666
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
  1520
    by (simp add: ring_hom_cring.hom_finsum [OF Phi_hom]
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1521
      ring_hom_cring.hom_mult [OF Phi_hom]
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1522
      ring_hom_cring.hom_pow [OF Phi_hom] Phi
14666
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
  1523
      ring_hom_cring.hom_finsum [OF Psi_hom]
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1524
      ring_hom_cring.hom_mult [OF Psi_hom]
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1525
      ring_hom_cring.hom_pow [OF Psi_hom] Psi RS Pi_def comp_def)
15095
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
  1526
*)
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1527
  also have "... = Psi p"
15095
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
  1528
    by (simp add: up_repr P S monom_mult [THEN sym] monom_pow del: monom_mult)
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1529
  finally show ?thesis .
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1530
qed
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1531
15095
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
  1532
theorem (in UP_univ_prop) UP_universal_property:
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1533
  "s \<in> carrier S ==>
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1534
  EX! Phi. Phi \<in> ring_hom P S \<inter> extensional (carrier P) &
14666
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
  1535
    Phi (monom P \<one> 1) = s &
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1536
    (ALL r : carrier R. Phi (monom P r 0) = h r)"
14666
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
  1537
  using eval_monom1
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1538
  apply (auto intro: eval_ring_hom eval_const eval_extensional)
14666
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
  1539
  apply (rule extensionalityI)
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
  1540
  apply (auto intro: UP_hom_unique)
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
  1541
  done
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1542
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1543
subsection {* Sample application of evaluation homomorphism *}
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1544
15095
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
  1545
lemma UP_univ_propI:
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1546
  assumes "cring R"
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1547
    and "cring S"
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1548
    and "h \<in> ring_hom R S"
15095
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
  1549
  shows "UP_univ_prop R S h"
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
  1550
  by (fast intro: UP_univ_prop.intro ring_hom_cring_axioms.intro
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1551
    cring.axioms prems)
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1552
13975
c8e9a89883ce Small changes for release Isabelle 2003.
ballarin
parents: 13949
diff changeset
  1553
constdefs
c8e9a89883ce Small changes for release Isabelle 2003.
ballarin
parents: 13949
diff changeset
  1554
  INTEG :: "int ring"
c8e9a89883ce Small changes for release Isabelle 2003.
ballarin
parents: 13949
diff changeset
  1555
  "INTEG == (| carrier = UNIV, mult = op *, one = 1, zero = 0, add = op + |)"
c8e9a89883ce Small changes for release Isabelle 2003.
ballarin
parents: 13949
diff changeset
  1556
15095
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
  1557
lemma INTEG_cring:
13975
c8e9a89883ce Small changes for release Isabelle 2003.
ballarin
parents: 13949
diff changeset
  1558
  "cring INTEG"
c8e9a89883ce Small changes for release Isabelle 2003.
ballarin
parents: 13949
diff changeset
  1559
  by (unfold INTEG_def) (auto intro!: cringI abelian_groupI comm_monoidI
c8e9a89883ce Small changes for release Isabelle 2003.
ballarin
parents: 13949
diff changeset
  1560
    zadd_zminus_inverse2 zadd_zmult_distrib)
c8e9a89883ce Small changes for release Isabelle 2003.
ballarin
parents: 13949
diff changeset
  1561
15095
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
  1562
lemma INTEG_id_eval:
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
  1563
  "UP_univ_prop INTEG INTEG id"
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
  1564
  by (fast intro: UP_univ_propI INTEG_cring id_ring_hom)
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1565
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1566
text {*
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1567
  An instantiation mechanism would now import all theorems and lemmas
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1568
  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
  1569
  "UP INTEG"} globally.
14666
65f8680c3f16 improved notation;
wenzelm
parents: 14651
diff changeset
  1570
*}
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1571
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1572
lemma INTEG_closed [intro, simp]:
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1573
  "z \<in> carrier INTEG"
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1574
  by (unfold INTEG_def) simp
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1575
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1576
lemma INTEG_mult [simp]:
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1577
  "mult INTEG z w = z * w"
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1578
  by (unfold INTEG_def) simp
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1579
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1580
lemma INTEG_pow [simp]:
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1581
  "pow INTEG z n = z ^ n"
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1582
  by (induct n) (simp_all add: INTEG_def nat_pow_def)
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1583
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1584
lemma "eval INTEG INTEG id 10 (monom (UP INTEG) 5 2) = 500"
15095
63f5f4c265dd Theories now take advantage of recent syntax improvements with (structure).
ballarin
parents: 15076
diff changeset
  1585
  by (simp add: UP_univ_prop.eval_monom [OF INTEG_id_eval])
13940
c67798653056 HOL-Algebra: New polynomial development added.
ballarin
parents:
diff changeset
  1586
14590
276ef51cedbf simplified ML code for setsubgoaler;
wenzelm
parents: 14577
diff changeset
  1587
end