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