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