src/HOL/Decision_Procs/Commutative_Ring.thy
author wenzelm
Sat, 07 Apr 2012 16:41:59 +0200
changeset 47389 e8552cba702d
parent 41807 ab5d2d81f9fb
child 47432 e1576d13e933
permissions -rw-r--r--
explicit checks stable_finished_theory/stable_command allow parallel asynchronous command transactions; tuned;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
28952
15a4b2cf8c34 made repository layout more coherent with logical distribution structure; stripped some $Id$s
haftmann
parents: 27487
diff changeset
     1
(*  Author:     Bernhard Haeupler
17516
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
     2
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
     3
Proving equalities in commutative rings done "right" in Isabelle/HOL.
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
     4
*)
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
     5
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
     6
header {* Proving equalities in commutative rings *}
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
     7
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
     8
theory Commutative_Ring
33356
9157d0f9f00e moved Commutative_Ring into session Decision_Procs
haftmann
parents: 32960
diff changeset
     9
imports Main Parity
9157d0f9f00e moved Commutative_Ring into session Decision_Procs
haftmann
parents: 32960
diff changeset
    10
uses ("commutative_ring_tac.ML")
17516
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
    11
begin
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
    12
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
    13
text {* Syntax of multivariate polynomials (pol) and polynomial expressions. *}
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
    14
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
    15
datatype 'a pol =
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
    16
    Pc 'a
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
    17
  | Pinj nat "'a pol"
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
    18
  | PX "'a pol" nat "'a pol"
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
    19
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
    20
datatype 'a polex =
20622
e1a573146be5 tuned proofs;
wenzelm
parents: 19736
diff changeset
    21
    Pol "'a pol"
17516
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
    22
  | Add "'a polex" "'a polex"
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
    23
  | Sub "'a polex" "'a polex"
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
    24
  | Mul "'a polex" "'a polex"
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
    25
  | Pow "'a polex" nat
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
    26
  | Neg "'a polex"
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
    27
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
    28
text {* Interpretation functions for the shadow syntax. *}
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
    29
31021
53642251a04f farewell to class recpower
haftmann
parents: 30663
diff changeset
    30
primrec
53642251a04f farewell to class recpower
haftmann
parents: 30663
diff changeset
    31
  Ipol :: "'a::{comm_ring_1} list \<Rightarrow> 'a pol \<Rightarrow> 'a"
22742
06165e40e7bd switched from recdef to function package; constants add, mul, pow now curried; infix syntax for algebraic operations.
haftmann
parents: 22665
diff changeset
    32
where
06165e40e7bd switched from recdef to function package; constants add, mul, pow now curried; infix syntax for algebraic operations.
haftmann
parents: 22665
diff changeset
    33
    "Ipol l (Pc c) = c"
06165e40e7bd switched from recdef to function package; constants add, mul, pow now curried; infix syntax for algebraic operations.
haftmann
parents: 22665
diff changeset
    34
  | "Ipol l (Pinj i P) = Ipol (drop i l) P"
06165e40e7bd switched from recdef to function package; constants add, mul, pow now curried; infix syntax for algebraic operations.
haftmann
parents: 22665
diff changeset
    35
  | "Ipol l (PX P x Q) = Ipol l P * (hd l)^x + Ipol (drop 1 l) Q"
17516
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
    36
31021
53642251a04f farewell to class recpower
haftmann
parents: 30663
diff changeset
    37
primrec
53642251a04f farewell to class recpower
haftmann
parents: 30663
diff changeset
    38
  Ipolex :: "'a::{comm_ring_1} list \<Rightarrow> 'a polex \<Rightarrow> 'a"
22742
06165e40e7bd switched from recdef to function package; constants add, mul, pow now curried; infix syntax for algebraic operations.
haftmann
parents: 22665
diff changeset
    39
where
06165e40e7bd switched from recdef to function package; constants add, mul, pow now curried; infix syntax for algebraic operations.
haftmann
parents: 22665
diff changeset
    40
    "Ipolex l (Pol P) = Ipol l P"
06165e40e7bd switched from recdef to function package; constants add, mul, pow now curried; infix syntax for algebraic operations.
haftmann
parents: 22665
diff changeset
    41
  | "Ipolex l (Add P Q) = Ipolex l P + Ipolex l Q"
06165e40e7bd switched from recdef to function package; constants add, mul, pow now curried; infix syntax for algebraic operations.
haftmann
parents: 22665
diff changeset
    42
  | "Ipolex l (Sub P Q) = Ipolex l P - Ipolex l Q"
06165e40e7bd switched from recdef to function package; constants add, mul, pow now curried; infix syntax for algebraic operations.
haftmann
parents: 22665
diff changeset
    43
  | "Ipolex l (Mul P Q) = Ipolex l P * Ipolex l Q"
06165e40e7bd switched from recdef to function package; constants add, mul, pow now curried; infix syntax for algebraic operations.
haftmann
parents: 22665
diff changeset
    44
  | "Ipolex l (Pow p n) = Ipolex l p ^ n"
06165e40e7bd switched from recdef to function package; constants add, mul, pow now curried; infix syntax for algebraic operations.
haftmann
parents: 22665
diff changeset
    45
  | "Ipolex l (Neg P) = - Ipolex l P"
17516
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
    46
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
    47
text {* Create polynomial normalized polynomials given normalized inputs. *}
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
    48
19736
wenzelm
parents: 18708
diff changeset
    49
definition
21404
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21256
diff changeset
    50
  mkPinj :: "nat \<Rightarrow> 'a pol \<Rightarrow> 'a pol" where
19736
wenzelm
parents: 18708
diff changeset
    51
  "mkPinj x P = (case P of
17516
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
    52
    Pc c \<Rightarrow> Pc c |
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
    53
    Pinj y P \<Rightarrow> Pinj (x + y) P |
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
    54
    PX p1 y p2 \<Rightarrow> Pinj x P)"
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
    55
19736
wenzelm
parents: 18708
diff changeset
    56
definition
31021
53642251a04f farewell to class recpower
haftmann
parents: 30663
diff changeset
    57
  mkPX :: "'a::{comm_ring} pol \<Rightarrow> nat \<Rightarrow> 'a pol \<Rightarrow> 'a pol" where
19736
wenzelm
parents: 18708
diff changeset
    58
  "mkPX P i Q = (case P of
17516
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
    59
    Pc c \<Rightarrow> (if (c = 0) then (mkPinj 1 Q) else (PX P i Q)) |
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
    60
    Pinj j R \<Rightarrow> PX P i Q |
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
    61
    PX P2 i2 Q2 \<Rightarrow> (if (Q2 = (Pc 0)) then (PX P2 (i+i2) Q) else (PX P i Q)) )"
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
    62
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
    63
text {* Defining the basic ring operations on normalized polynomials *}
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
    64
22742
06165e40e7bd switched from recdef to function package; constants add, mul, pow now curried; infix syntax for algebraic operations.
haftmann
parents: 22665
diff changeset
    65
function
31021
53642251a04f farewell to class recpower
haftmann
parents: 30663
diff changeset
    66
  add :: "'a::{comm_ring} pol \<Rightarrow> 'a pol \<Rightarrow> 'a pol" (infixl "\<oplus>" 65)
22742
06165e40e7bd switched from recdef to function package; constants add, mul, pow now curried; infix syntax for algebraic operations.
haftmann
parents: 22665
diff changeset
    67
where
06165e40e7bd switched from recdef to function package; constants add, mul, pow now curried; infix syntax for algebraic operations.
haftmann
parents: 22665
diff changeset
    68
    "Pc a \<oplus> Pc b = Pc (a + b)"
06165e40e7bd switched from recdef to function package; constants add, mul, pow now curried; infix syntax for algebraic operations.
haftmann
parents: 22665
diff changeset
    69
  | "Pc c \<oplus> Pinj i P = Pinj i (P \<oplus> Pc c)"
06165e40e7bd switched from recdef to function package; constants add, mul, pow now curried; infix syntax for algebraic operations.
haftmann
parents: 22665
diff changeset
    70
  | "Pinj i P \<oplus> Pc c = Pinj i (P \<oplus> Pc c)"
06165e40e7bd switched from recdef to function package; constants add, mul, pow now curried; infix syntax for algebraic operations.
haftmann
parents: 22665
diff changeset
    71
  | "Pc c \<oplus> PX P i Q = PX P i (Q \<oplus> Pc c)"
06165e40e7bd switched from recdef to function package; constants add, mul, pow now curried; infix syntax for algebraic operations.
haftmann
parents: 22665
diff changeset
    72
  | "PX P i Q \<oplus> Pc c = PX P i (Q \<oplus> Pc c)"
06165e40e7bd switched from recdef to function package; constants add, mul, pow now curried; infix syntax for algebraic operations.
haftmann
parents: 22665
diff changeset
    73
  | "Pinj x P \<oplus> Pinj y Q =
06165e40e7bd switched from recdef to function package; constants add, mul, pow now curried; infix syntax for algebraic operations.
haftmann
parents: 22665
diff changeset
    74
      (if x = y then mkPinj x (P \<oplus> Q)
06165e40e7bd switched from recdef to function package; constants add, mul, pow now curried; infix syntax for algebraic operations.
haftmann
parents: 22665
diff changeset
    75
       else (if x > y then mkPinj y (Pinj (x - y) P \<oplus> Q)
06165e40e7bd switched from recdef to function package; constants add, mul, pow now curried; infix syntax for algebraic operations.
haftmann
parents: 22665
diff changeset
    76
         else mkPinj x (Pinj (y - x) Q \<oplus> P)))"
06165e40e7bd switched from recdef to function package; constants add, mul, pow now curried; infix syntax for algebraic operations.
haftmann
parents: 22665
diff changeset
    77
  | "Pinj x P \<oplus> PX Q y R =
06165e40e7bd switched from recdef to function package; constants add, mul, pow now curried; infix syntax for algebraic operations.
haftmann
parents: 22665
diff changeset
    78
      (if x = 0 then P \<oplus> PX Q y R
06165e40e7bd switched from recdef to function package; constants add, mul, pow now curried; infix syntax for algebraic operations.
haftmann
parents: 22665
diff changeset
    79
       else (if x = 1 then PX Q y (R \<oplus> P)
06165e40e7bd switched from recdef to function package; constants add, mul, pow now curried; infix syntax for algebraic operations.
haftmann
parents: 22665
diff changeset
    80
         else PX Q y (R \<oplus> Pinj (x - 1) P)))"
06165e40e7bd switched from recdef to function package; constants add, mul, pow now curried; infix syntax for algebraic operations.
haftmann
parents: 22665
diff changeset
    81
  | "PX P x R \<oplus> Pinj y Q =
06165e40e7bd switched from recdef to function package; constants add, mul, pow now curried; infix syntax for algebraic operations.
haftmann
parents: 22665
diff changeset
    82
      (if y = 0 then PX P x R \<oplus> Q
06165e40e7bd switched from recdef to function package; constants add, mul, pow now curried; infix syntax for algebraic operations.
haftmann
parents: 22665
diff changeset
    83
       else (if y = 1 then PX P x (R \<oplus> Q)
06165e40e7bd switched from recdef to function package; constants add, mul, pow now curried; infix syntax for algebraic operations.
haftmann
parents: 22665
diff changeset
    84
         else PX P x (R \<oplus> Pinj (y - 1) Q)))"
06165e40e7bd switched from recdef to function package; constants add, mul, pow now curried; infix syntax for algebraic operations.
haftmann
parents: 22665
diff changeset
    85
  | "PX P1 x P2 \<oplus> PX Q1 y Q2 =
06165e40e7bd switched from recdef to function package; constants add, mul, pow now curried; infix syntax for algebraic operations.
haftmann
parents: 22665
diff changeset
    86
      (if x = y then mkPX (P1 \<oplus> Q1) x (P2 \<oplus> Q2)
06165e40e7bd switched from recdef to function package; constants add, mul, pow now curried; infix syntax for algebraic operations.
haftmann
parents: 22665
diff changeset
    87
       else (if x > y then mkPX (PX P1 (x - y) (Pc 0) \<oplus> Q1) y (P2 \<oplus> Q2)
06165e40e7bd switched from recdef to function package; constants add, mul, pow now curried; infix syntax for algebraic operations.
haftmann
parents: 22665
diff changeset
    88
         else mkPX (PX Q1 (y-x) (Pc 0) \<oplus> P1) x (P2 \<oplus> Q2)))"
06165e40e7bd switched from recdef to function package; constants add, mul, pow now curried; infix syntax for algebraic operations.
haftmann
parents: 22665
diff changeset
    89
by pat_completeness auto
06165e40e7bd switched from recdef to function package; constants add, mul, pow now curried; infix syntax for algebraic operations.
haftmann
parents: 22665
diff changeset
    90
termination by (relation "measure (\<lambda>(x, y). size x + size y)") auto
17516
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
    91
22742
06165e40e7bd switched from recdef to function package; constants add, mul, pow now curried; infix syntax for algebraic operations.
haftmann
parents: 22665
diff changeset
    92
function
31021
53642251a04f farewell to class recpower
haftmann
parents: 30663
diff changeset
    93
  mul :: "'a::{comm_ring} pol \<Rightarrow> 'a pol \<Rightarrow> 'a pol" (infixl "\<otimes>" 70)
22742
06165e40e7bd switched from recdef to function package; constants add, mul, pow now curried; infix syntax for algebraic operations.
haftmann
parents: 22665
diff changeset
    94
where
06165e40e7bd switched from recdef to function package; constants add, mul, pow now curried; infix syntax for algebraic operations.
haftmann
parents: 22665
diff changeset
    95
    "Pc a \<otimes> Pc b = Pc (a * b)"
06165e40e7bd switched from recdef to function package; constants add, mul, pow now curried; infix syntax for algebraic operations.
haftmann
parents: 22665
diff changeset
    96
  | "Pc c \<otimes> Pinj i P =
06165e40e7bd switched from recdef to function package; constants add, mul, pow now curried; infix syntax for algebraic operations.
haftmann
parents: 22665
diff changeset
    97
      (if c = 0 then Pc 0 else mkPinj i (P \<otimes> Pc c))"
06165e40e7bd switched from recdef to function package; constants add, mul, pow now curried; infix syntax for algebraic operations.
haftmann
parents: 22665
diff changeset
    98
  | "Pinj i P \<otimes> Pc c =
06165e40e7bd switched from recdef to function package; constants add, mul, pow now curried; infix syntax for algebraic operations.
haftmann
parents: 22665
diff changeset
    99
      (if c = 0 then Pc 0 else mkPinj i (P \<otimes> Pc c))"
06165e40e7bd switched from recdef to function package; constants add, mul, pow now curried; infix syntax for algebraic operations.
haftmann
parents: 22665
diff changeset
   100
  | "Pc c \<otimes> PX P i Q =
06165e40e7bd switched from recdef to function package; constants add, mul, pow now curried; infix syntax for algebraic operations.
haftmann
parents: 22665
diff changeset
   101
      (if c = 0 then Pc 0 else mkPX (P \<otimes> Pc c) i (Q \<otimes> Pc c))"
06165e40e7bd switched from recdef to function package; constants add, mul, pow now curried; infix syntax for algebraic operations.
haftmann
parents: 22665
diff changeset
   102
  | "PX P i Q \<otimes> Pc c =
06165e40e7bd switched from recdef to function package; constants add, mul, pow now curried; infix syntax for algebraic operations.
haftmann
parents: 22665
diff changeset
   103
      (if c = 0 then Pc 0 else mkPX (P \<otimes> Pc c) i (Q \<otimes> Pc c))"
06165e40e7bd switched from recdef to function package; constants add, mul, pow now curried; infix syntax for algebraic operations.
haftmann
parents: 22665
diff changeset
   104
  | "Pinj x P \<otimes> Pinj y Q =
06165e40e7bd switched from recdef to function package; constants add, mul, pow now curried; infix syntax for algebraic operations.
haftmann
parents: 22665
diff changeset
   105
      (if x = y then mkPinj x (P \<otimes> Q) else
06165e40e7bd switched from recdef to function package; constants add, mul, pow now curried; infix syntax for algebraic operations.
haftmann
parents: 22665
diff changeset
   106
         (if x > y then mkPinj y (Pinj (x-y) P \<otimes> Q)
06165e40e7bd switched from recdef to function package; constants add, mul, pow now curried; infix syntax for algebraic operations.
haftmann
parents: 22665
diff changeset
   107
           else mkPinj x (Pinj (y - x) Q \<otimes> P)))"
06165e40e7bd switched from recdef to function package; constants add, mul, pow now curried; infix syntax for algebraic operations.
haftmann
parents: 22665
diff changeset
   108
  | "Pinj x P \<otimes> PX Q y R =
06165e40e7bd switched from recdef to function package; constants add, mul, pow now curried; infix syntax for algebraic operations.
haftmann
parents: 22665
diff changeset
   109
      (if x = 0 then P \<otimes> PX Q y R else
06165e40e7bd switched from recdef to function package; constants add, mul, pow now curried; infix syntax for algebraic operations.
haftmann
parents: 22665
diff changeset
   110
         (if x = 1 then mkPX (Pinj x P \<otimes> Q) y (R \<otimes> P)
06165e40e7bd switched from recdef to function package; constants add, mul, pow now curried; infix syntax for algebraic operations.
haftmann
parents: 22665
diff changeset
   111
           else mkPX (Pinj x P \<otimes> Q) y (R \<otimes> Pinj (x - 1) P)))"
06165e40e7bd switched from recdef to function package; constants add, mul, pow now curried; infix syntax for algebraic operations.
haftmann
parents: 22665
diff changeset
   112
  | "PX P x R \<otimes> Pinj y Q =
06165e40e7bd switched from recdef to function package; constants add, mul, pow now curried; infix syntax for algebraic operations.
haftmann
parents: 22665
diff changeset
   113
      (if y = 0 then PX P x R \<otimes> Q else
06165e40e7bd switched from recdef to function package; constants add, mul, pow now curried; infix syntax for algebraic operations.
haftmann
parents: 22665
diff changeset
   114
         (if y = 1 then mkPX (Pinj y Q \<otimes> P) x (R \<otimes> Q)
06165e40e7bd switched from recdef to function package; constants add, mul, pow now curried; infix syntax for algebraic operations.
haftmann
parents: 22665
diff changeset
   115
           else mkPX (Pinj y Q \<otimes> P) x (R \<otimes> Pinj (y - 1) Q)))"
06165e40e7bd switched from recdef to function package; constants add, mul, pow now curried; infix syntax for algebraic operations.
haftmann
parents: 22665
diff changeset
   116
  | "PX P1 x P2 \<otimes> PX Q1 y Q2 =
06165e40e7bd switched from recdef to function package; constants add, mul, pow now curried; infix syntax for algebraic operations.
haftmann
parents: 22665
diff changeset
   117
      mkPX (P1 \<otimes> Q1) (x + y) (P2 \<otimes> Q2) \<oplus>
06165e40e7bd switched from recdef to function package; constants add, mul, pow now curried; infix syntax for algebraic operations.
haftmann
parents: 22665
diff changeset
   118
        (mkPX (P1 \<otimes> mkPinj 1 Q2) x (Pc 0) \<oplus>
06165e40e7bd switched from recdef to function package; constants add, mul, pow now curried; infix syntax for algebraic operations.
haftmann
parents: 22665
diff changeset
   119
          (mkPX (Q1 \<otimes> mkPinj 1 P2) y (Pc 0)))"
06165e40e7bd switched from recdef to function package; constants add, mul, pow now curried; infix syntax for algebraic operations.
haftmann
parents: 22665
diff changeset
   120
by pat_completeness auto
06165e40e7bd switched from recdef to function package; constants add, mul, pow now curried; infix syntax for algebraic operations.
haftmann
parents: 22665
diff changeset
   121
termination by (relation "measure (\<lambda>(x, y). size x + size y)")
06165e40e7bd switched from recdef to function package; constants add, mul, pow now curried; infix syntax for algebraic operations.
haftmann
parents: 22665
diff changeset
   122
  (auto simp add: mkPinj_def split: pol.split)
17516
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   123
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   124
text {* Negation*}
31021
53642251a04f farewell to class recpower
haftmann
parents: 30663
diff changeset
   125
primrec
53642251a04f farewell to class recpower
haftmann
parents: 30663
diff changeset
   126
  neg :: "'a::{comm_ring} pol \<Rightarrow> 'a pol"
22742
06165e40e7bd switched from recdef to function package; constants add, mul, pow now curried; infix syntax for algebraic operations.
haftmann
parents: 22665
diff changeset
   127
where
06165e40e7bd switched from recdef to function package; constants add, mul, pow now curried; infix syntax for algebraic operations.
haftmann
parents: 22665
diff changeset
   128
    "neg (Pc c) = Pc (-c)"
06165e40e7bd switched from recdef to function package; constants add, mul, pow now curried; infix syntax for algebraic operations.
haftmann
parents: 22665
diff changeset
   129
  | "neg (Pinj i P) = Pinj i (neg P)"
06165e40e7bd switched from recdef to function package; constants add, mul, pow now curried; infix syntax for algebraic operations.
haftmann
parents: 22665
diff changeset
   130
  | "neg (PX P x Q) = PX (neg P) x (neg Q)"
17516
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   131
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   132
text {* Substraction *}
19736
wenzelm
parents: 18708
diff changeset
   133
definition
31021
53642251a04f farewell to class recpower
haftmann
parents: 30663
diff changeset
   134
  sub :: "'a::{comm_ring} pol \<Rightarrow> 'a pol \<Rightarrow> 'a pol" (infixl "\<ominus>" 65)
22742
06165e40e7bd switched from recdef to function package; constants add, mul, pow now curried; infix syntax for algebraic operations.
haftmann
parents: 22665
diff changeset
   135
where
06165e40e7bd switched from recdef to function package; constants add, mul, pow now curried; infix syntax for algebraic operations.
haftmann
parents: 22665
diff changeset
   136
  "sub P Q = P \<oplus> neg Q"
17516
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   137
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   138
text {* Square for Fast Exponentation *}
31021
53642251a04f farewell to class recpower
haftmann
parents: 30663
diff changeset
   139
primrec
53642251a04f farewell to class recpower
haftmann
parents: 30663
diff changeset
   140
  sqr :: "'a::{comm_ring_1} pol \<Rightarrow> 'a pol"
22742
06165e40e7bd switched from recdef to function package; constants add, mul, pow now curried; infix syntax for algebraic operations.
haftmann
parents: 22665
diff changeset
   141
where
06165e40e7bd switched from recdef to function package; constants add, mul, pow now curried; infix syntax for algebraic operations.
haftmann
parents: 22665
diff changeset
   142
    "sqr (Pc c) = Pc (c * c)"
06165e40e7bd switched from recdef to function package; constants add, mul, pow now curried; infix syntax for algebraic operations.
haftmann
parents: 22665
diff changeset
   143
  | "sqr (Pinj i P) = mkPinj i (sqr P)"
06165e40e7bd switched from recdef to function package; constants add, mul, pow now curried; infix syntax for algebraic operations.
haftmann
parents: 22665
diff changeset
   144
  | "sqr (PX A x B) = mkPX (sqr A) (x + x) (sqr B) \<oplus>
06165e40e7bd switched from recdef to function package; constants add, mul, pow now curried; infix syntax for algebraic operations.
haftmann
parents: 22665
diff changeset
   145
      mkPX (Pc (1 + 1) \<otimes> A \<otimes> mkPinj 1 B) x (Pc 0)"
17516
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   146
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   147
text {* Fast Exponentation *}
22742
06165e40e7bd switched from recdef to function package; constants add, mul, pow now curried; infix syntax for algebraic operations.
haftmann
parents: 22665
diff changeset
   148
fun
31021
53642251a04f farewell to class recpower
haftmann
parents: 30663
diff changeset
   149
  pow :: "nat \<Rightarrow> 'a::{comm_ring_1} pol \<Rightarrow> 'a pol"
22742
06165e40e7bd switched from recdef to function package; constants add, mul, pow now curried; infix syntax for algebraic operations.
haftmann
parents: 22665
diff changeset
   150
where
06165e40e7bd switched from recdef to function package; constants add, mul, pow now curried; infix syntax for algebraic operations.
haftmann
parents: 22665
diff changeset
   151
    "pow 0 P = Pc 1"
06165e40e7bd switched from recdef to function package; constants add, mul, pow now curried; infix syntax for algebraic operations.
haftmann
parents: 22665
diff changeset
   152
  | "pow n P = (if even n then pow (n div 2) (sqr P)
06165e40e7bd switched from recdef to function package; constants add, mul, pow now curried; infix syntax for algebraic operations.
haftmann
parents: 22665
diff changeset
   153
       else P \<otimes> pow (n div 2) (sqr P))"
06165e40e7bd switched from recdef to function package; constants add, mul, pow now curried; infix syntax for algebraic operations.
haftmann
parents: 22665
diff changeset
   154
  
17516
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   155
lemma pow_if:
22742
06165e40e7bd switched from recdef to function package; constants add, mul, pow now curried; infix syntax for algebraic operations.
haftmann
parents: 22665
diff changeset
   156
  "pow n P =
06165e40e7bd switched from recdef to function package; constants add, mul, pow now curried; infix syntax for algebraic operations.
haftmann
parents: 22665
diff changeset
   157
   (if n = 0 then Pc 1 else if even n then pow (n div 2) (sqr P)
06165e40e7bd switched from recdef to function package; constants add, mul, pow now curried; infix syntax for algebraic operations.
haftmann
parents: 22665
diff changeset
   158
    else P \<otimes> pow (n div 2) (sqr P))"
17516
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   159
  by (cases n) simp_all
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   160
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   161
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   162
text {* Normalization of polynomial expressions *}
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   163
31021
53642251a04f farewell to class recpower
haftmann
parents: 30663
diff changeset
   164
primrec
53642251a04f farewell to class recpower
haftmann
parents: 30663
diff changeset
   165
  norm :: "'a::{comm_ring_1} polex \<Rightarrow> 'a pol"
22742
06165e40e7bd switched from recdef to function package; constants add, mul, pow now curried; infix syntax for algebraic operations.
haftmann
parents: 22665
diff changeset
   166
where
06165e40e7bd switched from recdef to function package; constants add, mul, pow now curried; infix syntax for algebraic operations.
haftmann
parents: 22665
diff changeset
   167
    "norm (Pol P) = P"
06165e40e7bd switched from recdef to function package; constants add, mul, pow now curried; infix syntax for algebraic operations.
haftmann
parents: 22665
diff changeset
   168
  | "norm (Add P Q) = norm P \<oplus> norm Q"
06165e40e7bd switched from recdef to function package; constants add, mul, pow now curried; infix syntax for algebraic operations.
haftmann
parents: 22665
diff changeset
   169
  | "norm (Sub P Q) = norm P \<ominus> norm Q"
06165e40e7bd switched from recdef to function package; constants add, mul, pow now curried; infix syntax for algebraic operations.
haftmann
parents: 22665
diff changeset
   170
  | "norm (Mul P Q) = norm P \<otimes> norm Q"
06165e40e7bd switched from recdef to function package; constants add, mul, pow now curried; infix syntax for algebraic operations.
haftmann
parents: 22665
diff changeset
   171
  | "norm (Pow P n) = pow n (norm P)"
06165e40e7bd switched from recdef to function package; constants add, mul, pow now curried; infix syntax for algebraic operations.
haftmann
parents: 22665
diff changeset
   172
  | "norm (Neg P) = neg (norm P)"
17516
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   173
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   174
text {* mkPinj preserve semantics *}
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   175
lemma mkPinj_ci: "Ipol l (mkPinj a B) = Ipol l (Pinj a B)"
29667
53103fc8ffa3 Replaced group_ and ring_simps by algebra_simps;
nipkow
parents: 28952
diff changeset
   176
  by (induct B) (auto simp add: mkPinj_def algebra_simps)
17516
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   177
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   178
text {* mkPX preserves semantics *}
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   179
lemma mkPX_ci: "Ipol l (mkPX A b C) = Ipol l (PX A b C)"
29667
53103fc8ffa3 Replaced group_ and ring_simps by algebra_simps;
nipkow
parents: 28952
diff changeset
   180
  by (cases A) (auto simp add: mkPX_def mkPinj_ci power_add algebra_simps)
17516
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   181
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   182
text {* Correctness theorems for the implemented operations *}
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   183
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   184
text {* Negation *}
20622
e1a573146be5 tuned proofs;
wenzelm
parents: 19736
diff changeset
   185
lemma neg_ci: "Ipol l (neg P) = -(Ipol l P)"
e1a573146be5 tuned proofs;
wenzelm
parents: 19736
diff changeset
   186
  by (induct P arbitrary: l) auto
17516
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   187
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   188
text {* Addition *}
22742
06165e40e7bd switched from recdef to function package; constants add, mul, pow now curried; infix syntax for algebraic operations.
haftmann
parents: 22665
diff changeset
   189
lemma add_ci: "Ipol l (P \<oplus> Q) = Ipol l P + Ipol l Q"
20622
e1a573146be5 tuned proofs;
wenzelm
parents: 19736
diff changeset
   190
proof (induct P Q arbitrary: l rule: add.induct)
17516
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   191
  case (6 x P y Q)
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   192
  show ?case
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   193
  proof (rule linorder_cases)
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   194
    assume "x < y"
29667
53103fc8ffa3 Replaced group_ and ring_simps by algebra_simps;
nipkow
parents: 28952
diff changeset
   195
    with 6 show ?case by (simp add: mkPinj_ci algebra_simps)
17516
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   196
  next
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   197
    assume "x = y"
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   198
    with 6 show ?case by (simp add: mkPinj_ci)
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   199
  next
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   200
    assume "x > y"
29667
53103fc8ffa3 Replaced group_ and ring_simps by algebra_simps;
nipkow
parents: 28952
diff changeset
   201
    with 6 show ?case by (simp add: mkPinj_ci algebra_simps)
17516
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   202
  qed
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   203
next
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   204
  case (7 x P Q y R)
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   205
  have "x = 0 \<or> x = 1 \<or> x > 1" by arith
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   206
  moreover
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   207
  { assume "x = 0" with 7 have ?case by simp }
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   208
  moreover
29667
53103fc8ffa3 Replaced group_ and ring_simps by algebra_simps;
nipkow
parents: 28952
diff changeset
   209
  { assume "x = 1" with 7 have ?case by (simp add: algebra_simps) }
17516
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   210
  moreover
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   211
  { assume "x > 1" from 7 have ?case by (cases x) simp_all }
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   212
  ultimately show ?case by blast
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   213
next
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   214
  case (8 P x R y Q)
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   215
  have "y = 0 \<or> y = 1 \<or> y > 1" by arith
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   216
  moreover
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   217
  { assume "y = 0" with 8 have ?case by simp }
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   218
  moreover
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   219
  { assume "y = 1" with 8 have ?case by simp }
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   220
  moreover
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   221
  { assume "y > 1" with 8 have ?case by simp }
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   222
  ultimately show ?case by blast
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   223
next
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   224
  case (9 P1 x P2 Q1 y Q2)
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   225
  show ?case
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   226
  proof (rule linorder_cases)
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   227
    assume a: "x < y" hence "EX d. d + x = y" by arith
29667
53103fc8ffa3 Replaced group_ and ring_simps by algebra_simps;
nipkow
parents: 28952
diff changeset
   228
    with 9 a show ?case by (auto simp add: mkPX_ci power_add algebra_simps)
17516
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   229
  next
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   230
    assume a: "y < x" hence "EX d. d + y = x" by arith
29667
53103fc8ffa3 Replaced group_ and ring_simps by algebra_simps;
nipkow
parents: 28952
diff changeset
   231
    with 9 a show ?case by (auto simp add: power_add mkPX_ci algebra_simps)
17516
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   232
  next
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   233
    assume "x = y"
29667
53103fc8ffa3 Replaced group_ and ring_simps by algebra_simps;
nipkow
parents: 28952
diff changeset
   234
    with 9 show ?case by (simp add: mkPX_ci algebra_simps)
17516
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   235
  qed
29667
53103fc8ffa3 Replaced group_ and ring_simps by algebra_simps;
nipkow
parents: 28952
diff changeset
   236
qed (auto simp add: algebra_simps)
17516
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   237
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   238
text {* Multiplication *}
22742
06165e40e7bd switched from recdef to function package; constants add, mul, pow now curried; infix syntax for algebraic operations.
haftmann
parents: 22665
diff changeset
   239
lemma mul_ci: "Ipol l (P \<otimes> Q) = Ipol l P * Ipol l Q"
20622
e1a573146be5 tuned proofs;
wenzelm
parents: 19736
diff changeset
   240
  by (induct P Q arbitrary: l rule: mul.induct)
29667
53103fc8ffa3 Replaced group_ and ring_simps by algebra_simps;
nipkow
parents: 28952
diff changeset
   241
    (simp_all add: mkPX_ci mkPinj_ci algebra_simps add_ci power_add)
17516
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   242
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   243
text {* Substraction *}
22742
06165e40e7bd switched from recdef to function package; constants add, mul, pow now curried; infix syntax for algebraic operations.
haftmann
parents: 22665
diff changeset
   244
lemma sub_ci: "Ipol l (P \<ominus> Q) = Ipol l P - Ipol l Q"
17516
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   245
  by (simp add: add_ci neg_ci sub_def)
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   246
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   247
text {* Square *}
22742
06165e40e7bd switched from recdef to function package; constants add, mul, pow now curried; infix syntax for algebraic operations.
haftmann
parents: 22665
diff changeset
   248
lemma sqr_ci: "Ipol ls (sqr P) = Ipol ls P * Ipol ls P"
06165e40e7bd switched from recdef to function package; constants add, mul, pow now curried; infix syntax for algebraic operations.
haftmann
parents: 22665
diff changeset
   249
  by (induct P arbitrary: ls)
29667
53103fc8ffa3 Replaced group_ and ring_simps by algebra_simps;
nipkow
parents: 28952
diff changeset
   250
    (simp_all add: add_ci mkPinj_ci mkPX_ci mul_ci algebra_simps power_add)
17516
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   251
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   252
text {* Power *}
22742
06165e40e7bd switched from recdef to function package; constants add, mul, pow now curried; infix syntax for algebraic operations.
haftmann
parents: 22665
diff changeset
   253
lemma even_pow:"even n \<Longrightarrow> pow n P = pow (n div 2) (sqr P)"
20622
e1a573146be5 tuned proofs;
wenzelm
parents: 19736
diff changeset
   254
  by (induct n) simp_all
17516
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   255
22742
06165e40e7bd switched from recdef to function package; constants add, mul, pow now curried; infix syntax for algebraic operations.
haftmann
parents: 22665
diff changeset
   256
lemma pow_ci: "Ipol ls (pow n P) = Ipol ls P ^ n"
06165e40e7bd switched from recdef to function package; constants add, mul, pow now curried; infix syntax for algebraic operations.
haftmann
parents: 22665
diff changeset
   257
proof (induct n arbitrary: P rule: nat_less_induct)
17516
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   258
  case (1 k)
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   259
  show ?case
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   260
  proof (cases k)
20622
e1a573146be5 tuned proofs;
wenzelm
parents: 19736
diff changeset
   261
    case 0
e1a573146be5 tuned proofs;
wenzelm
parents: 19736
diff changeset
   262
    then show ?thesis by simp
e1a573146be5 tuned proofs;
wenzelm
parents: 19736
diff changeset
   263
  next
17516
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   264
    case (Suc l)
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   265
    show ?thesis
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   266
    proof cases
20622
e1a573146be5 tuned proofs;
wenzelm
parents: 19736
diff changeset
   267
      assume "even l"
e1a573146be5 tuned proofs;
wenzelm
parents: 19736
diff changeset
   268
      then have "Suc l div 2 = l div 2"
40077
c8a9eaaa2f59 nat_number -> eval_nat_numeral
nipkow
parents: 33356
diff changeset
   269
        by (simp add: eval_nat_numeral even_nat_plus_one_div_two)
17516
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   270
      moreover
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   271
      from Suc have "l < k" by simp
22742
06165e40e7bd switched from recdef to function package; constants add, mul, pow now curried; infix syntax for algebraic operations.
haftmann
parents: 22665
diff changeset
   272
      with 1 have "\<And>P. Ipol ls (pow l P) = Ipol ls P ^ l" by simp
17516
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   273
      moreover
20622
e1a573146be5 tuned proofs;
wenzelm
parents: 19736
diff changeset
   274
      note Suc `even l` even_nat_plus_one_div_two
17516
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   275
      ultimately show ?thesis by (auto simp add: mul_ci power_Suc even_pow)
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   276
    next
20622
e1a573146be5 tuned proofs;
wenzelm
parents: 19736
diff changeset
   277
      assume "odd l"
e1a573146be5 tuned proofs;
wenzelm
parents: 19736
diff changeset
   278
      {
e1a573146be5 tuned proofs;
wenzelm
parents: 19736
diff changeset
   279
        fix p
22742
06165e40e7bd switched from recdef to function package; constants add, mul, pow now curried; infix syntax for algebraic operations.
haftmann
parents: 22665
diff changeset
   280
        have "Ipol ls (sqr P) ^ (Suc l div 2) = Ipol ls P ^ Suc l"
20622
e1a573146be5 tuned proofs;
wenzelm
parents: 19736
diff changeset
   281
        proof (cases l)
e1a573146be5 tuned proofs;
wenzelm
parents: 19736
diff changeset
   282
          case 0
e1a573146be5 tuned proofs;
wenzelm
parents: 19736
diff changeset
   283
          with `odd l` show ?thesis by simp
e1a573146be5 tuned proofs;
wenzelm
parents: 19736
diff changeset
   284
        next
e1a573146be5 tuned proofs;
wenzelm
parents: 19736
diff changeset
   285
          case (Suc w)
e1a573146be5 tuned proofs;
wenzelm
parents: 19736
diff changeset
   286
          with `odd l` have "even w" by simp
20678
f6d602833557 tuned proofs;
wenzelm
parents: 20622
diff changeset
   287
          have two_times: "2 * (w div 2) = w"
f6d602833557 tuned proofs;
wenzelm
parents: 20622
diff changeset
   288
            by (simp only: numerals even_nat_div_two_times_two [OF `even w`])
22742
06165e40e7bd switched from recdef to function package; constants add, mul, pow now curried; infix syntax for algebraic operations.
haftmann
parents: 22665
diff changeset
   289
          have "Ipol ls P * Ipol ls P = Ipol ls P ^ Suc (Suc 0)"
20622
e1a573146be5 tuned proofs;
wenzelm
parents: 19736
diff changeset
   290
            by (simp add: power_Suc)
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 31021
diff changeset
   291
          then have "Ipol ls P * Ipol ls P = Ipol ls P ^ 2"
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 31021
diff changeset
   292
            by (simp add: numerals)
20622
e1a573146be5 tuned proofs;
wenzelm
parents: 19736
diff changeset
   293
          with Suc show ?thesis
30273
ecd6f0ca62ea declare power_Suc [simp]; remove redundant type-specific versions of power_Suc
huffman
parents: 29667
diff changeset
   294
            by (auto simp add: power_mult [symmetric, of _ 2 _] two_times mul_ci sqr_ci
ecd6f0ca62ea declare power_Suc [simp]; remove redundant type-specific versions of power_Suc
huffman
parents: 29667
diff changeset
   295
                     simp del: power_Suc)
20622
e1a573146be5 tuned proofs;
wenzelm
parents: 19736
diff changeset
   296
        qed
e1a573146be5 tuned proofs;
wenzelm
parents: 19736
diff changeset
   297
      } with 1 Suc `odd l` show ?thesis by simp
17516
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   298
    qed
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   299
  qed
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   300
qed
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   301
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   302
text {* Normalization preserves semantics  *}
20622
e1a573146be5 tuned proofs;
wenzelm
parents: 19736
diff changeset
   303
lemma norm_ci: "Ipolex l Pe = Ipol l (norm Pe)"
17516
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   304
  by (induct Pe) (simp_all add: add_ci sub_ci mul_ci neg_ci pow_ci)
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   305
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   306
text {* Reflection lemma: Key to the (incomplete) decision procedure *}
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   307
lemma norm_eq:
20622
e1a573146be5 tuned proofs;
wenzelm
parents: 19736
diff changeset
   308
  assumes "norm P1 = norm P2"
17516
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   309
  shows "Ipolex l P1 = Ipolex l P2"
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   310
proof -
41807
ab5d2d81f9fb tuned proofs -- eliminated prems;
wenzelm
parents: 40077
diff changeset
   311
  from assms have "Ipol l (norm P1) = Ipol l (norm P2)" by simp
20622
e1a573146be5 tuned proofs;
wenzelm
parents: 19736
diff changeset
   312
  then show ?thesis by (simp only: norm_ci)
17516
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   313
qed
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   314
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   315
33356
9157d0f9f00e moved Commutative_Ring into session Decision_Procs
haftmann
parents: 32960
diff changeset
   316
use "commutative_ring_tac.ML"
9157d0f9f00e moved Commutative_Ring into session Decision_Procs
haftmann
parents: 32960
diff changeset
   317
setup Commutative_Ring_Tac.setup
17516
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   318
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   319
end