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