src/HOL/Decision_Procs/Commutative_Ring.thy
author wenzelm
Thu, 09 Jul 2015 23:46:21 +0200
changeset 60708 f425e80a3eb0
parent 60534 b2add2b08412
child 64962 bf41e1109db3
permissions -rw-r--r--
tuned proofs;
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
60533
1e7ccd864b62 isabelle update_cartouches;
wenzelm
parents: 58889
diff changeset
     6
section \<open>Proving equalities in commutative rings\<close>
17516
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
58770
ae5e9b4f8daf downshift of theory Parity in the hierarchy
haftmann
parents: 58712
diff changeset
     9
imports Main
17516
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
    10
begin
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
    11
60533
1e7ccd864b62 isabelle update_cartouches;
wenzelm
parents: 58889
diff changeset
    12
text \<open>Syntax of multivariate polynomials (pol) and polynomial expressions.\<close>
17516
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
    13
58310
91ea607a34d8 updated news
blanchet
parents: 58259
diff changeset
    14
datatype 'a pol =
17516
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
    15
    Pc 'a
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
    16
  | Pinj nat "'a pol"
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
    17
  | PX "'a pol" nat "'a pol"
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
    18
58310
91ea607a34d8 updated news
blanchet
parents: 58259
diff changeset
    19
datatype 'a polex =
20622
e1a573146be5 tuned proofs;
wenzelm
parents: 19736
diff changeset
    20
    Pol "'a pol"
17516
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
    21
  | Add "'a polex" "'a polex"
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
    22
  | Sub "'a polex" "'a polex"
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
    23
  | Mul "'a polex" "'a polex"
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
    24
  | Pow "'a polex" nat
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
    25
  | Neg "'a polex"
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
    26
60533
1e7ccd864b62 isabelle update_cartouches;
wenzelm
parents: 58889
diff changeset
    27
text \<open>Interpretation functions for the shadow syntax.\<close>
17516
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
    28
60534
b2add2b08412 tuned proofs;
wenzelm
parents: 60533
diff changeset
    29
primrec 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
    30
where
06165e40e7bd switched from recdef to function package; constants add, mul, pow now curried; infix syntax for algebraic operations.
haftmann
parents: 22665
diff changeset
    31
    "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
    32
  | "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
    33
  | "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
    34
60534
b2add2b08412 tuned proofs;
wenzelm
parents: 60533
diff changeset
    35
primrec 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
    36
where
06165e40e7bd switched from recdef to function package; constants add, mul, pow now curried; infix syntax for algebraic operations.
haftmann
parents: 22665
diff changeset
    37
    "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
    38
  | "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
    39
  | "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
    40
  | "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
    41
  | "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
    42
  | "Ipolex l (Neg P) = - Ipolex l P"
17516
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
    43
60533
1e7ccd864b62 isabelle update_cartouches;
wenzelm
parents: 58889
diff changeset
    44
text \<open>Create polynomial normalized polynomials given normalized inputs.\<close>
17516
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
    45
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 53077
diff changeset
    46
definition mkPinj :: "nat \<Rightarrow> 'a pol \<Rightarrow> 'a pol"
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 53077
diff changeset
    47
where
60534
b2add2b08412 tuned proofs;
wenzelm
parents: 60533
diff changeset
    48
  "mkPinj x P =
b2add2b08412 tuned proofs;
wenzelm
parents: 60533
diff changeset
    49
    (case P of
b2add2b08412 tuned proofs;
wenzelm
parents: 60533
diff changeset
    50
      Pc c \<Rightarrow> Pc c
b2add2b08412 tuned proofs;
wenzelm
parents: 60533
diff changeset
    51
    | Pinj y P \<Rightarrow> Pinj (x + y) P
b2add2b08412 tuned proofs;
wenzelm
parents: 60533
diff changeset
    52
    | PX p1 y p2 \<Rightarrow> Pinj x P)"
17516
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
    53
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 53077
diff changeset
    54
definition mkPX :: "'a::comm_ring pol \<Rightarrow> nat \<Rightarrow> 'a pol \<Rightarrow> 'a pol"
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 53077
diff changeset
    55
where
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 53077
diff changeset
    56
  "mkPX P i Q =
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 53077
diff changeset
    57
    (case P of
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 53077
diff changeset
    58
      Pc c \<Rightarrow> if c = 0 then mkPinj 1 Q else PX P i Q
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 53077
diff changeset
    59
    | Pinj j R \<Rightarrow> PX P i Q
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 53077
diff changeset
    60
    | PX P2 i2 Q2 \<Rightarrow> if Q2 = Pc 0 then PX P2 (i + i2) Q else PX P i Q)"
17516
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
    61
60533
1e7ccd864b62 isabelle update_cartouches;
wenzelm
parents: 58889
diff changeset
    62
text \<open>Defining the basic ring operations on normalized polynomials\<close>
17516
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
    63
58259
52c35a59bbf5 ported Decision_Procs to new datatypes
blanchet
parents: 58249
diff changeset
    64
lemma pol_size_nz[simp]: "size (p :: 'a pol) \<noteq> 0"
52c35a59bbf5 ported Decision_Procs to new datatypes
blanchet
parents: 58249
diff changeset
    65
  by (cases p) simp_all
52c35a59bbf5 ported Decision_Procs to new datatypes
blanchet
parents: 58249
diff changeset
    66
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 53077
diff changeset
    67
function 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
    68
where
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 53077
diff changeset
    69
  "Pc a \<oplus> Pc b = Pc (a + b)"
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 53077
diff changeset
    70
| "Pc c \<oplus> Pinj i P = Pinj i (P \<oplus> Pc c)"
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 53077
diff changeset
    71
| "Pinj i P \<oplus> Pc c = Pinj i (P \<oplus> Pc c)"
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 53077
diff changeset
    72
| "Pc c \<oplus> PX P i Q = PX P i (Q \<oplus> Pc c)"
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 53077
diff changeset
    73
| "PX P i Q \<oplus> Pc c = PX P i (Q \<oplus> Pc c)"
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 53077
diff changeset
    74
| "Pinj x P \<oplus> Pinj y Q =
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 53077
diff changeset
    75
    (if x = y then mkPinj x (P \<oplus> Q)
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 53077
diff changeset
    76
     else (if x > y then mkPinj y (Pinj (x - y) P \<oplus> Q)
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 53077
diff changeset
    77
       else mkPinj x (Pinj (y - x) Q \<oplus> P)))"
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 53077
diff changeset
    78
| "Pinj x P \<oplus> PX Q y R =
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 53077
diff changeset
    79
    (if x = 0 then P \<oplus> PX Q y R
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 53077
diff changeset
    80
     else (if x = 1 then PX Q y (R \<oplus> P)
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 53077
diff changeset
    81
       else PX Q y (R \<oplus> Pinj (x - 1) P)))"
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 53077
diff changeset
    82
| "PX P x R \<oplus> Pinj y Q =
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 53077
diff changeset
    83
    (if y = 0 then PX P x R \<oplus> Q
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 53077
diff changeset
    84
     else (if y = 1 then PX P x (R \<oplus> Q)
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 53077
diff changeset
    85
       else PX P x (R \<oplus> Pinj (y - 1) Q)))"
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 53077
diff changeset
    86
| "PX P1 x P2 \<oplus> PX Q1 y Q2 =
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 53077
diff changeset
    87
    (if x = y then mkPX (P1 \<oplus> Q1) x (P2 \<oplus> Q2)
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 53077
diff changeset
    88
     else (if x > y then mkPX (PX P1 (x - y) (Pc 0) \<oplus> Q1) y (P2 \<oplus> Q2)
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 53077
diff changeset
    89
       else mkPX (PX Q1 (y-x) (Pc 0) \<oplus> P1) x (P2 \<oplus> Q2)))"
22742
06165e40e7bd switched from recdef to function package; constants add, mul, pow now curried; infix syntax for algebraic operations.
haftmann
parents: 22665
diff changeset
    90
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
    91
termination by (relation "measure (\<lambda>(x, y). size x + size y)") auto
17516
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
    92
60534
b2add2b08412 tuned proofs;
wenzelm
parents: 60533
diff changeset
    93
function 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
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 53077
diff changeset
    95
  "Pc a \<otimes> Pc b = Pc (a * b)"
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 53077
diff changeset
    96
| "Pc c \<otimes> Pinj i P =
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 53077
diff changeset
    97
    (if c = 0 then Pc 0 else mkPinj i (P \<otimes> Pc c))"
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 53077
diff changeset
    98
| "Pinj i P \<otimes> Pc c =
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 53077
diff changeset
    99
    (if c = 0 then Pc 0 else mkPinj i (P \<otimes> Pc c))"
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 53077
diff changeset
   100
| "Pc c \<otimes> PX P i Q =
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 53077
diff changeset
   101
    (if c = 0 then Pc 0 else mkPX (P \<otimes> Pc c) i (Q \<otimes> Pc c))"
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 53077
diff changeset
   102
| "PX P i Q \<otimes> Pc c =
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 53077
diff changeset
   103
    (if c = 0 then Pc 0 else mkPX (P \<otimes> Pc c) i (Q \<otimes> Pc c))"
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 53077
diff changeset
   104
| "Pinj x P \<otimes> Pinj y Q =
60534
b2add2b08412 tuned proofs;
wenzelm
parents: 60533
diff changeset
   105
    (if x = y then mkPinj x (P \<otimes> Q)
b2add2b08412 tuned proofs;
wenzelm
parents: 60533
diff changeset
   106
     else
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 53077
diff changeset
   107
       (if x > y then mkPinj y (Pinj (x-y) P \<otimes> Q)
60534
b2add2b08412 tuned proofs;
wenzelm
parents: 60533
diff changeset
   108
        else mkPinj x (Pinj (y - x) Q \<otimes> P)))"
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 53077
diff changeset
   109
| "Pinj x P \<otimes> PX Q y R =
60534
b2add2b08412 tuned proofs;
wenzelm
parents: 60533
diff changeset
   110
    (if x = 0 then P \<otimes> PX Q y R
b2add2b08412 tuned proofs;
wenzelm
parents: 60533
diff changeset
   111
     else
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 53077
diff changeset
   112
       (if x = 1 then mkPX (Pinj x P \<otimes> Q) y (R \<otimes> P)
60534
b2add2b08412 tuned proofs;
wenzelm
parents: 60533
diff changeset
   113
        else mkPX (Pinj x P \<otimes> Q) y (R \<otimes> Pinj (x - 1) P)))"
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 53077
diff changeset
   114
| "PX P x R \<otimes> Pinj y Q =
60534
b2add2b08412 tuned proofs;
wenzelm
parents: 60533
diff changeset
   115
    (if y = 0 then PX P x R \<otimes> Q
b2add2b08412 tuned proofs;
wenzelm
parents: 60533
diff changeset
   116
     else
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 53077
diff changeset
   117
       (if y = 1 then mkPX (Pinj y Q \<otimes> P) x (R \<otimes> Q)
60534
b2add2b08412 tuned proofs;
wenzelm
parents: 60533
diff changeset
   118
        else mkPX (Pinj y Q \<otimes> P) x (R \<otimes> Pinj (y - 1) Q)))"
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 53077
diff changeset
   119
| "PX P1 x P2 \<otimes> PX Q1 y Q2 =
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 53077
diff changeset
   120
    mkPX (P1 \<otimes> Q1) (x + y) (P2 \<otimes> Q2) \<oplus>
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 53077
diff changeset
   121
      (mkPX (P1 \<otimes> mkPinj 1 Q2) x (Pc 0) \<oplus>
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 53077
diff changeset
   122
        (mkPX (Q1 \<otimes> mkPinj 1 P2) y (Pc 0)))"
22742
06165e40e7bd switched from recdef to function package; constants add, mul, pow now curried; infix syntax for algebraic operations.
haftmann
parents: 22665
diff changeset
   123
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
   124
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
   125
  (auto simp add: mkPinj_def split: pol.split)
17516
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   126
60533
1e7ccd864b62 isabelle update_cartouches;
wenzelm
parents: 58889
diff changeset
   127
text \<open>Negation\<close>
60534
b2add2b08412 tuned proofs;
wenzelm
parents: 60533
diff changeset
   128
primrec 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
   129
where
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 53077
diff changeset
   130
  "neg (Pc c) = Pc (-c)"
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 53077
diff changeset
   131
| "neg (Pinj i P) = Pinj i (neg P)"
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 53077
diff changeset
   132
| "neg (PX P x Q) = PX (neg P) x (neg Q)"
17516
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   133
60533
1e7ccd864b62 isabelle update_cartouches;
wenzelm
parents: 58889
diff changeset
   134
text \<open>Substraction\<close>
60534
b2add2b08412 tuned proofs;
wenzelm
parents: 60533
diff changeset
   135
definition sub :: "'a::comm_ring pol \<Rightarrow> 'a pol \<Rightarrow> 'a pol"  (infixl "\<ominus>" 65)
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 53077
diff changeset
   136
  where "sub P Q = P \<oplus> neg Q"
17516
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   137
60533
1e7ccd864b62 isabelle update_cartouches;
wenzelm
parents: 58889
diff changeset
   138
text \<open>Square for Fast Exponentation\<close>
60534
b2add2b08412 tuned proofs;
wenzelm
parents: 60533
diff changeset
   139
primrec 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
   140
where
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 53077
diff changeset
   141
  "sqr (Pc c) = Pc (c * c)"
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 53077
diff changeset
   142
| "sqr (Pinj i P) = mkPinj i (sqr P)"
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 53077
diff changeset
   143
| "sqr (PX A x B) =
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 53077
diff changeset
   144
    mkPX (sqr A) (x + x) (sqr B) \<oplus> 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
   145
60533
1e7ccd864b62 isabelle update_cartouches;
wenzelm
parents: 58889
diff changeset
   146
text \<open>Fast Exponentation\<close>
58710
7216a10d69ba augmented and tuned facts on even/odd and division
haftmann
parents: 58310
diff changeset
   147
60534
b2add2b08412 tuned proofs;
wenzelm
parents: 60533
diff changeset
   148
fun 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
   149
where
58710
7216a10d69ba augmented and tuned facts on even/odd and division
haftmann
parents: 58310
diff changeset
   150
  pow_if [simp del]: "pow n P =
60534
b2add2b08412 tuned proofs;
wenzelm
parents: 60533
diff changeset
   151
   (if n = 0 then Pc 1
b2add2b08412 tuned proofs;
wenzelm
parents: 60533
diff changeset
   152
    else if even n then pow (n div 2) (sqr 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
   153
    else P \<otimes> pow (n div 2) (sqr P))"
17516
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   154
60534
b2add2b08412 tuned proofs;
wenzelm
parents: 60533
diff changeset
   155
lemma pow_simps [simp]:
58710
7216a10d69ba augmented and tuned facts on even/odd and division
haftmann
parents: 58310
diff changeset
   156
  "pow 0 P = Pc 1"
58712
709d8f68ec29 avoid unsafe simp rules
haftmann
parents: 58710
diff changeset
   157
  "pow (2 * n) P = pow n (sqr P)"
709d8f68ec29 avoid unsafe simp rules
haftmann
parents: 58710
diff changeset
   158
  "pow (Suc (2 * n)) P = P \<otimes> pow n (sqr P)"
58710
7216a10d69ba augmented and tuned facts on even/odd and division
haftmann
parents: 58310
diff changeset
   159
  by (simp_all add: pow_if)
17516
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   160
60534
b2add2b08412 tuned proofs;
wenzelm
parents: 60533
diff changeset
   161
lemma even_pow: "even n \<Longrightarrow> pow n P = pow (n div 2) (sqr P)"
58712
709d8f68ec29 avoid unsafe simp rules
haftmann
parents: 58710
diff changeset
   162
  by (erule evenE) simp
709d8f68ec29 avoid unsafe simp rules
haftmann
parents: 58710
diff changeset
   163
60534
b2add2b08412 tuned proofs;
wenzelm
parents: 60533
diff changeset
   164
lemma odd_pow: "odd n \<Longrightarrow> pow n P = P \<otimes> pow (n div 2) (sqr P)"
58712
709d8f68ec29 avoid unsafe simp rules
haftmann
parents: 58710
diff changeset
   165
  by (erule oddE) simp
709d8f68ec29 avoid unsafe simp rules
haftmann
parents: 58710
diff changeset
   166
60534
b2add2b08412 tuned proofs;
wenzelm
parents: 60533
diff changeset
   167
60533
1e7ccd864b62 isabelle update_cartouches;
wenzelm
parents: 58889
diff changeset
   168
text \<open>Normalization of polynomial expressions\<close>
17516
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   169
60534
b2add2b08412 tuned proofs;
wenzelm
parents: 60533
diff changeset
   170
primrec 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
   171
where
55754
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 53077
diff changeset
   172
  "norm (Pol P) = P"
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 53077
diff changeset
   173
| "norm (Add P Q) = norm P \<oplus> norm Q"
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 53077
diff changeset
   174
| "norm (Sub P Q) = norm P \<ominus> norm Q"
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 53077
diff changeset
   175
| "norm (Mul P Q) = norm P \<otimes> norm Q"
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 53077
diff changeset
   176
| "norm (Pow P n) = pow n (norm P)"
d14072d53c1e tuned specifications and proofs;
wenzelm
parents: 53077
diff changeset
   177
| "norm (Neg P) = neg (norm P)"
17516
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   178
60533
1e7ccd864b62 isabelle update_cartouches;
wenzelm
parents: 58889
diff changeset
   179
text \<open>mkPinj preserve semantics\<close>
17516
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   180
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
   181
  by (induct B) (auto simp add: mkPinj_def algebra_simps)
17516
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   182
60533
1e7ccd864b62 isabelle update_cartouches;
wenzelm
parents: 58889
diff changeset
   183
text \<open>mkPX preserves semantics\<close>
17516
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   184
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
   185
  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
   186
60533
1e7ccd864b62 isabelle update_cartouches;
wenzelm
parents: 58889
diff changeset
   187
text \<open>Correctness theorems for the implemented operations\<close>
17516
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   188
60533
1e7ccd864b62 isabelle update_cartouches;
wenzelm
parents: 58889
diff changeset
   189
text \<open>Negation\<close>
20622
e1a573146be5 tuned proofs;
wenzelm
parents: 19736
diff changeset
   190
lemma neg_ci: "Ipol l (neg P) = -(Ipol l P)"
e1a573146be5 tuned proofs;
wenzelm
parents: 19736
diff changeset
   191
  by (induct P arbitrary: l) auto
17516
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   192
60533
1e7ccd864b62 isabelle update_cartouches;
wenzelm
parents: 58889
diff changeset
   193
text \<open>Addition\<close>
22742
06165e40e7bd switched from recdef to function package; constants add, mul, pow now curried; infix syntax for algebraic operations.
haftmann
parents: 22665
diff changeset
   194
lemma add_ci: "Ipol l (P \<oplus> Q) = Ipol l P + Ipol l Q"
20622
e1a573146be5 tuned proofs;
wenzelm
parents: 19736
diff changeset
   195
proof (induct P Q arbitrary: l rule: add.induct)
17516
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   196
  case (6 x P y Q)
60708
f425e80a3eb0 tuned proofs;
wenzelm
parents: 60534
diff changeset
   197
  consider "x < y" | "x = y" | "x > y" by arith
f425e80a3eb0 tuned proofs;
wenzelm
parents: 60534
diff changeset
   198
  then
17516
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   199
  show ?case
60708
f425e80a3eb0 tuned proofs;
wenzelm
parents: 60534
diff changeset
   200
  proof cases
f425e80a3eb0 tuned proofs;
wenzelm
parents: 60534
diff changeset
   201
    case 1
f425e80a3eb0 tuned proofs;
wenzelm
parents: 60534
diff changeset
   202
    with 6 show ?thesis by (simp add: mkPinj_ci algebra_simps)
17516
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   203
  next
60708
f425e80a3eb0 tuned proofs;
wenzelm
parents: 60534
diff changeset
   204
    case 2
f425e80a3eb0 tuned proofs;
wenzelm
parents: 60534
diff changeset
   205
    with 6 show ?thesis by (simp add: mkPinj_ci)
17516
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   206
  next
60708
f425e80a3eb0 tuned proofs;
wenzelm
parents: 60534
diff changeset
   207
    case 3
f425e80a3eb0 tuned proofs;
wenzelm
parents: 60534
diff changeset
   208
    with 6 show ?thesis by (simp add: mkPinj_ci algebra_simps)
17516
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   209
  qed
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   210
next
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   211
  case (7 x P Q y R)
60534
b2add2b08412 tuned proofs;
wenzelm
parents: 60533
diff changeset
   212
  consider "x = 0" | "x = 1" | "x > 1" by arith
b2add2b08412 tuned proofs;
wenzelm
parents: 60533
diff changeset
   213
  then show ?case
b2add2b08412 tuned proofs;
wenzelm
parents: 60533
diff changeset
   214
  proof cases
b2add2b08412 tuned proofs;
wenzelm
parents: 60533
diff changeset
   215
    case 1
b2add2b08412 tuned proofs;
wenzelm
parents: 60533
diff changeset
   216
    with 7 show ?thesis by simp
b2add2b08412 tuned proofs;
wenzelm
parents: 60533
diff changeset
   217
  next
b2add2b08412 tuned proofs;
wenzelm
parents: 60533
diff changeset
   218
    case 2
b2add2b08412 tuned proofs;
wenzelm
parents: 60533
diff changeset
   219
    with 7 show ?thesis by (simp add: algebra_simps)
b2add2b08412 tuned proofs;
wenzelm
parents: 60533
diff changeset
   220
  next
b2add2b08412 tuned proofs;
wenzelm
parents: 60533
diff changeset
   221
    case 3
b2add2b08412 tuned proofs;
wenzelm
parents: 60533
diff changeset
   222
    from 7 show ?thesis by (cases x) simp_all
b2add2b08412 tuned proofs;
wenzelm
parents: 60533
diff changeset
   223
  qed
17516
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   224
next
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   225
  case (8 P x R y Q)
60534
b2add2b08412 tuned proofs;
wenzelm
parents: 60533
diff changeset
   226
  then show ?case by simp
17516
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   227
next
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   228
  case (9 P1 x P2 Q1 y Q2)
60534
b2add2b08412 tuned proofs;
wenzelm
parents: 60533
diff changeset
   229
  consider "x = y" | d where "d + x = y" | d where "d + y = x"
b2add2b08412 tuned proofs;
wenzelm
parents: 60533
diff changeset
   230
    by atomize_elim arith
b2add2b08412 tuned proofs;
wenzelm
parents: 60533
diff changeset
   231
  then show ?case
b2add2b08412 tuned proofs;
wenzelm
parents: 60533
diff changeset
   232
  proof cases
b2add2b08412 tuned proofs;
wenzelm
parents: 60533
diff changeset
   233
    case 1
b2add2b08412 tuned proofs;
wenzelm
parents: 60533
diff changeset
   234
    with 9 show ?thesis by (simp add: mkPX_ci algebra_simps)
17516
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   235
  next
60534
b2add2b08412 tuned proofs;
wenzelm
parents: 60533
diff changeset
   236
    case 2
b2add2b08412 tuned proofs;
wenzelm
parents: 60533
diff changeset
   237
    with 9 show ?thesis by (auto simp add: mkPX_ci power_add algebra_simps)
17516
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   238
  next
60534
b2add2b08412 tuned proofs;
wenzelm
parents: 60533
diff changeset
   239
    case 3
b2add2b08412 tuned proofs;
wenzelm
parents: 60533
diff changeset
   240
    with 9 show ?thesis by (auto simp add: power_add mkPX_ci algebra_simps)
17516
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   241
  qed
29667
53103fc8ffa3 Replaced group_ and ring_simps by algebra_simps;
nipkow
parents: 28952
diff changeset
   242
qed (auto simp add: algebra_simps)
17516
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   243
60533
1e7ccd864b62 isabelle update_cartouches;
wenzelm
parents: 58889
diff changeset
   244
text \<open>Multiplication\<close>
22742
06165e40e7bd switched from recdef to function package; constants add, mul, pow now curried; infix syntax for algebraic operations.
haftmann
parents: 22665
diff changeset
   245
lemma mul_ci: "Ipol l (P \<otimes> Q) = Ipol l P * Ipol l Q"
20622
e1a573146be5 tuned proofs;
wenzelm
parents: 19736
diff changeset
   246
  by (induct P Q arbitrary: l rule: mul.induct)
29667
53103fc8ffa3 Replaced group_ and ring_simps by algebra_simps;
nipkow
parents: 28952
diff changeset
   247
    (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
   248
60533
1e7ccd864b62 isabelle update_cartouches;
wenzelm
parents: 58889
diff changeset
   249
text \<open>Substraction\<close>
22742
06165e40e7bd switched from recdef to function package; constants add, mul, pow now curried; infix syntax for algebraic operations.
haftmann
parents: 22665
diff changeset
   250
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
   251
  by (simp add: add_ci neg_ci sub_def)
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   252
60533
1e7ccd864b62 isabelle update_cartouches;
wenzelm
parents: 58889
diff changeset
   253
text \<open>Square\<close>
22742
06165e40e7bd switched from recdef to function package; constants add, mul, pow now curried; infix syntax for algebraic operations.
haftmann
parents: 22665
diff changeset
   254
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
   255
  by (induct P arbitrary: ls)
29667
53103fc8ffa3 Replaced group_ and ring_simps by algebra_simps;
nipkow
parents: 28952
diff changeset
   256
    (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
   257
60533
1e7ccd864b62 isabelle update_cartouches;
wenzelm
parents: 58889
diff changeset
   258
text \<open>Power\<close>
22742
06165e40e7bd switched from recdef to function package; constants add, mul, pow now curried; infix syntax for algebraic operations.
haftmann
parents: 22665
diff changeset
   259
lemma pow_ci: "Ipol ls (pow n P) = Ipol ls P ^ n"
58712
709d8f68ec29 avoid unsafe simp rules
haftmann
parents: 58710
diff changeset
   260
proof (induct n arbitrary: P rule: less_induct)
709d8f68ec29 avoid unsafe simp rules
haftmann
parents: 58710
diff changeset
   261
  case (less k)
60708
f425e80a3eb0 tuned proofs;
wenzelm
parents: 60534
diff changeset
   262
  consider "k = 0" | "k > 0" by arith
f425e80a3eb0 tuned proofs;
wenzelm
parents: 60534
diff changeset
   263
  then
17516
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   264
  show ?case
60708
f425e80a3eb0 tuned proofs;
wenzelm
parents: 60534
diff changeset
   265
  proof cases
f425e80a3eb0 tuned proofs;
wenzelm
parents: 60534
diff changeset
   266
    case 1
60534
b2add2b08412 tuned proofs;
wenzelm
parents: 60533
diff changeset
   267
    then show ?thesis by simp
20622
e1a573146be5 tuned proofs;
wenzelm
parents: 19736
diff changeset
   268
  next
60708
f425e80a3eb0 tuned proofs;
wenzelm
parents: 60534
diff changeset
   269
    case 2
58712
709d8f68ec29 avoid unsafe simp rules
haftmann
parents: 58710
diff changeset
   270
    then have "k div 2 < k" by arith
709d8f68ec29 avoid unsafe simp rules
haftmann
parents: 58710
diff changeset
   271
    with less have *: "Ipol ls (pow (k div 2) (sqr P)) = Ipol ls (sqr P) ^ (k div 2)"
709d8f68ec29 avoid unsafe simp rules
haftmann
parents: 58710
diff changeset
   272
      by simp
17516
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   273
    show ?thesis
58712
709d8f68ec29 avoid unsafe simp rules
haftmann
parents: 58710
diff changeset
   274
    proof (cases "even k")
60534
b2add2b08412 tuned proofs;
wenzelm
parents: 60533
diff changeset
   275
      case True
b2add2b08412 tuned proofs;
wenzelm
parents: 60533
diff changeset
   276
      with * show ?thesis
b2add2b08412 tuned proofs;
wenzelm
parents: 60533
diff changeset
   277
        by (simp add: even_pow sqr_ci power_mult_distrib power_add [symmetric]
b2add2b08412 tuned proofs;
wenzelm
parents: 60533
diff changeset
   278
          mult_2 [symmetric] even_two_times_div_two)
17516
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   279
    next
60534
b2add2b08412 tuned proofs;
wenzelm
parents: 60533
diff changeset
   280
      case False
b2add2b08412 tuned proofs;
wenzelm
parents: 60533
diff changeset
   281
      with * show ?thesis
b2add2b08412 tuned proofs;
wenzelm
parents: 60533
diff changeset
   282
        by (simp add: odd_pow mul_ci sqr_ci power_mult_distrib power_add [symmetric]
b2add2b08412 tuned proofs;
wenzelm
parents: 60533
diff changeset
   283
          mult_2 [symmetric] power_Suc [symmetric])
17516
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   284
    qed
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   285
  qed
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   286
qed
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   287
60533
1e7ccd864b62 isabelle update_cartouches;
wenzelm
parents: 58889
diff changeset
   288
text \<open>Normalization preserves semantics\<close>
20622
e1a573146be5 tuned proofs;
wenzelm
parents: 19736
diff changeset
   289
lemma norm_ci: "Ipolex l Pe = Ipol l (norm Pe)"
17516
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   290
  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
   291
60533
1e7ccd864b62 isabelle update_cartouches;
wenzelm
parents: 58889
diff changeset
   292
text \<open>Reflection lemma: Key to the (incomplete) decision procedure\<close>
17516
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   293
lemma norm_eq:
20622
e1a573146be5 tuned proofs;
wenzelm
parents: 19736
diff changeset
   294
  assumes "norm P1 = norm P2"
17516
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   295
  shows "Ipolex l P1 = Ipolex l P2"
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   296
proof -
60534
b2add2b08412 tuned proofs;
wenzelm
parents: 60533
diff changeset
   297
  from assms have "Ipol l (norm P1) = Ipol l (norm P2)"
b2add2b08412 tuned proofs;
wenzelm
parents: 60533
diff changeset
   298
    by simp
b2add2b08412 tuned proofs;
wenzelm
parents: 60533
diff changeset
   299
  then show ?thesis
b2add2b08412 tuned proofs;
wenzelm
parents: 60533
diff changeset
   300
    by (simp only: norm_ci)
17516
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   301
qed
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   302
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   303
48891
c0eafbd55de3 prefer ML_file over old uses;
wenzelm
parents: 47432
diff changeset
   304
ML_file "commutative_ring_tac.ML"
47432
e1576d13e933 more standard method setup;
wenzelm
parents: 41807
diff changeset
   305
60533
1e7ccd864b62 isabelle update_cartouches;
wenzelm
parents: 58889
diff changeset
   306
method_setup comm_ring = \<open>
47432
e1576d13e933 more standard method setup;
wenzelm
parents: 41807
diff changeset
   307
  Scan.succeed (SIMPLE_METHOD' o Commutative_Ring_Tac.tac)
60533
1e7ccd864b62 isabelle update_cartouches;
wenzelm
parents: 58889
diff changeset
   308
\<close> "reflective decision procedure for equalities over commutative rings"
17516
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   309
45164074dad4 added Commutative_Ring (from Main HOL);
wenzelm
parents:
diff changeset
   310
end