author | haftmann |
Fri, 06 Feb 2009 09:05:19 +0100 | |
changeset 29814 | 15344c0899e1 |
parent 29667 | 53103fc8ffa3 |
child 30273 | ecd6f0ca62ea |
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 |
|
28952
15a4b2cf8c34
made repository layout more coherent with logical distribution structure; stripped some $Id$s
haftmann
parents:
27487
diff
changeset
|
9 |
imports Plain "~~/src/HOL/List" "~~/src/HOL/Parity" |
17516 | 10 |
uses ("comm_ring.ML") |
11 |
begin |
|
12 |
||
13 |
text {* Syntax of multivariate polynomials (pol) and polynomial expressions. *} |
|
14 |
||
15 |
datatype 'a pol = |
|
16 |
Pc 'a |
|
17 |
| Pinj nat "'a pol" |
|
18 |
| PX "'a pol" nat "'a pol" |
|
19 |
||
20 |
datatype 'a polex = |
|
20622 | 21 |
Pol "'a pol" |
17516 | 22 |
| Add "'a polex" "'a polex" |
23 |
| Sub "'a polex" "'a polex" |
|
24 |
| Mul "'a polex" "'a polex" |
|
25 |
| Pow "'a polex" nat |
|
26 |
| Neg "'a polex" |
|
27 |
||
28 |
text {* Interpretation functions for the shadow syntax. *} |
|
29 |
||
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 |
fun |
17516 | 31 |
Ipol :: "'a::{comm_ring,recpower} list \<Rightarrow> 'a pol \<Rightarrow> 'a" |
22742
06165e40e7bd
switched from recdef to function package; constants add, mul, pow now curried; infix syntax for algebraic operations.
haftmann
parents:
22665
diff
changeset
|
32 |
where |
06165e40e7bd
switched from recdef to function package; constants add, mul, pow now curried; infix syntax for algebraic operations.
haftmann
parents:
22665
diff
changeset
|
33 |
"Ipol l (Pc c) = c" |
06165e40e7bd
switched from recdef to function package; constants add, mul, pow now curried; infix syntax for algebraic operations.
haftmann
parents:
22665
diff
changeset
|
34 |
| "Ipol l (Pinj i P) = Ipol (drop i l) P" |
06165e40e7bd
switched from recdef to function package; constants add, mul, pow now curried; infix syntax for algebraic operations.
haftmann
parents:
22665
diff
changeset
|
35 |
| "Ipol l (PX P x Q) = Ipol l P * (hd l)^x + Ipol (drop 1 l) Q" |
17516 | 36 |
|
22742
06165e40e7bd
switched from recdef to function package; constants add, mul, pow now curried; infix syntax for algebraic operations.
haftmann
parents:
22665
diff
changeset
|
37 |
fun |
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 :: "'a::{comm_ring,recpower} list \<Rightarrow> 'a polex \<Rightarrow> 'a" |
06165e40e7bd
switched from recdef to function package; constants add, mul, pow now curried; infix syntax for algebraic operations.
haftmann
parents:
22665
diff
changeset
|
39 |
where |
06165e40e7bd
switched from recdef to function package; constants add, mul, pow now curried; infix syntax for algebraic operations.
haftmann
parents:
22665
diff
changeset
|
40 |
"Ipolex l (Pol P) = Ipol l P" |
06165e40e7bd
switched from recdef to function package; constants add, mul, pow now curried; infix syntax for algebraic operations.
haftmann
parents:
22665
diff
changeset
|
41 |
| "Ipolex l (Add P Q) = Ipolex l P + Ipolex l Q" |
06165e40e7bd
switched from recdef to function package; constants add, mul, pow now curried; infix syntax for algebraic operations.
haftmann
parents:
22665
diff
changeset
|
42 |
| "Ipolex l (Sub P Q) = Ipolex l P - Ipolex l Q" |
06165e40e7bd
switched from recdef to function package; constants add, mul, pow now curried; infix syntax for algebraic operations.
haftmann
parents:
22665
diff
changeset
|
43 |
| "Ipolex l (Mul P Q) = Ipolex l P * Ipolex l Q" |
06165e40e7bd
switched from recdef to function package; constants add, mul, pow now curried; infix syntax for algebraic operations.
haftmann
parents:
22665
diff
changeset
|
44 |
| "Ipolex l (Pow p n) = Ipolex l p ^ n" |
06165e40e7bd
switched from recdef to function package; constants add, mul, pow now curried; infix syntax for algebraic operations.
haftmann
parents:
22665
diff
changeset
|
45 |
| "Ipolex l (Neg P) = - Ipolex l P" |
17516 | 46 |
|
47 |
text {* Create polynomial normalized polynomials given normalized inputs. *} |
|
48 |
||
19736 | 49 |
definition |
21404
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
wenzelm
parents:
21256
diff
changeset
|
50 |
mkPinj :: "nat \<Rightarrow> 'a pol \<Rightarrow> 'a pol" where |
19736 | 51 |
"mkPinj x P = (case P of |
17516 | 52 |
Pc c \<Rightarrow> Pc c | |
53 |
Pinj y P \<Rightarrow> Pinj (x + y) P | |
|
54 |
PX p1 y p2 \<Rightarrow> Pinj x P)" |
|
55 |
||
19736 | 56 |
definition |
21404
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
wenzelm
parents:
21256
diff
changeset
|
57 |
mkPX :: "'a::{comm_ring,recpower} pol \<Rightarrow> nat \<Rightarrow> 'a pol \<Rightarrow> 'a pol" where |
19736 | 58 |
"mkPX P i Q = (case P of |
17516 | 59 |
Pc c \<Rightarrow> (if (c = 0) then (mkPinj 1 Q) else (PX P i Q)) | |
60 |
Pinj j R \<Rightarrow> PX P i Q | |
|
61 |
PX P2 i2 Q2 \<Rightarrow> (if (Q2 = (Pc 0)) then (PX P2 (i+i2) Q) else (PX P i Q)) )" |
|
62 |
||
63 |
text {* Defining the basic ring operations on normalized polynomials *} |
|
64 |
||
22742
06165e40e7bd
switched from recdef to function package; constants add, mul, pow now curried; infix syntax for algebraic operations.
haftmann
parents:
22665
diff
changeset
|
65 |
function |
06165e40e7bd
switched from recdef to function package; constants add, mul, pow now curried; infix syntax for algebraic operations.
haftmann
parents:
22665
diff
changeset
|
66 |
add :: "'a::{comm_ring,recpower} pol \<Rightarrow> 'a pol \<Rightarrow> 'a pol" (infixl "\<oplus>" 65) |
06165e40e7bd
switched from recdef to function package; constants add, mul, pow now curried; infix syntax for algebraic operations.
haftmann
parents:
22665
diff
changeset
|
67 |
where |
06165e40e7bd
switched from recdef to function package; constants add, mul, pow now curried; infix syntax for algebraic operations.
haftmann
parents:
22665
diff
changeset
|
68 |
"Pc a \<oplus> Pc b = Pc (a + b)" |
06165e40e7bd
switched from recdef to function package; constants add, mul, pow now curried; infix syntax for algebraic operations.
haftmann
parents:
22665
diff
changeset
|
69 |
| "Pc c \<oplus> Pinj i P = Pinj i (P \<oplus> Pc c)" |
06165e40e7bd
switched from recdef to function package; constants add, mul, pow now curried; infix syntax for algebraic operations.
haftmann
parents:
22665
diff
changeset
|
70 |
| "Pinj i P \<oplus> Pc c = Pinj i (P \<oplus> Pc c)" |
06165e40e7bd
switched from recdef to function package; constants add, mul, pow now curried; infix syntax for algebraic operations.
haftmann
parents:
22665
diff
changeset
|
71 |
| "Pc c \<oplus> PX P i Q = PX P i (Q \<oplus> Pc c)" |
06165e40e7bd
switched from recdef to function package; constants add, mul, pow now curried; infix syntax for algebraic operations.
haftmann
parents:
22665
diff
changeset
|
72 |
| "PX P i Q \<oplus> Pc c = PX P i (Q \<oplus> Pc c)" |
06165e40e7bd
switched from recdef to function package; constants add, mul, pow now curried; infix syntax for algebraic operations.
haftmann
parents:
22665
diff
changeset
|
73 |
| "Pinj x P \<oplus> Pinj y Q = |
06165e40e7bd
switched from recdef to function package; constants add, mul, pow now curried; infix syntax for algebraic operations.
haftmann
parents:
22665
diff
changeset
|
74 |
(if x = y then mkPinj x (P \<oplus> Q) |
06165e40e7bd
switched from recdef to function package; constants add, mul, pow now curried; infix syntax for algebraic operations.
haftmann
parents:
22665
diff
changeset
|
75 |
else (if x > y then mkPinj y (Pinj (x - y) P \<oplus> Q) |
06165e40e7bd
switched from recdef to function package; constants add, mul, pow now curried; infix syntax for algebraic operations.
haftmann
parents:
22665
diff
changeset
|
76 |
else mkPinj x (Pinj (y - x) Q \<oplus> P)))" |
06165e40e7bd
switched from recdef to function package; constants add, mul, pow now curried; infix syntax for algebraic operations.
haftmann
parents:
22665
diff
changeset
|
77 |
| "Pinj x P \<oplus> PX Q y R = |
06165e40e7bd
switched from recdef to function package; constants add, mul, pow now curried; infix syntax for algebraic operations.
haftmann
parents:
22665
diff
changeset
|
78 |
(if x = 0 then P \<oplus> PX Q y R |
06165e40e7bd
switched from recdef to function package; constants add, mul, pow now curried; infix syntax for algebraic operations.
haftmann
parents:
22665
diff
changeset
|
79 |
else (if x = 1 then PX Q y (R \<oplus> P) |
06165e40e7bd
switched from recdef to function package; constants add, mul, pow now curried; infix syntax for algebraic operations.
haftmann
parents:
22665
diff
changeset
|
80 |
else PX Q y (R \<oplus> Pinj (x - 1) P)))" |
06165e40e7bd
switched from recdef to function package; constants add, mul, pow now curried; infix syntax for algebraic operations.
haftmann
parents:
22665
diff
changeset
|
81 |
| "PX P x R \<oplus> Pinj y Q = |
06165e40e7bd
switched from recdef to function package; constants add, mul, pow now curried; infix syntax for algebraic operations.
haftmann
parents:
22665
diff
changeset
|
82 |
(if y = 0 then PX P x R \<oplus> Q |
06165e40e7bd
switched from recdef to function package; constants add, mul, pow now curried; infix syntax for algebraic operations.
haftmann
parents:
22665
diff
changeset
|
83 |
else (if y = 1 then PX P x (R \<oplus> Q) |
06165e40e7bd
switched from recdef to function package; constants add, mul, pow now curried; infix syntax for algebraic operations.
haftmann
parents:
22665
diff
changeset
|
84 |
else PX P x (R \<oplus> Pinj (y - 1) Q)))" |
06165e40e7bd
switched from recdef to function package; constants add, mul, pow now curried; infix syntax for algebraic operations.
haftmann
parents:
22665
diff
changeset
|
85 |
| "PX P1 x P2 \<oplus> PX Q1 y Q2 = |
06165e40e7bd
switched from recdef to function package; constants add, mul, pow now curried; infix syntax for algebraic operations.
haftmann
parents:
22665
diff
changeset
|
86 |
(if x = y then mkPX (P1 \<oplus> Q1) x (P2 \<oplus> Q2) |
06165e40e7bd
switched from recdef to function package; constants add, mul, pow now curried; infix syntax for algebraic operations.
haftmann
parents:
22665
diff
changeset
|
87 |
else (if x > y then mkPX (PX P1 (x - y) (Pc 0) \<oplus> Q1) y (P2 \<oplus> Q2) |
06165e40e7bd
switched from recdef to function package; constants add, mul, pow now curried; infix syntax for algebraic operations.
haftmann
parents:
22665
diff
changeset
|
88 |
else mkPX (PX Q1 (y-x) (Pc 0) \<oplus> P1) x (P2 \<oplus> Q2)))" |
06165e40e7bd
switched from recdef to function package; constants add, mul, pow now curried; infix syntax for algebraic operations.
haftmann
parents:
22665
diff
changeset
|
89 |
by pat_completeness auto |
06165e40e7bd
switched from recdef to function package; constants add, mul, pow now curried; infix syntax for algebraic operations.
haftmann
parents:
22665
diff
changeset
|
90 |
termination by (relation "measure (\<lambda>(x, y). size x + size y)") auto |
17516 | 91 |
|
22742
06165e40e7bd
switched from recdef to function package; constants add, mul, pow now curried; infix syntax for algebraic operations.
haftmann
parents:
22665
diff
changeset
|
92 |
function |
06165e40e7bd
switched from recdef to function package; constants add, mul, pow now curried; infix syntax for algebraic operations.
haftmann
parents:
22665
diff
changeset
|
93 |
mul :: "'a::{comm_ring,recpower} pol \<Rightarrow> 'a pol \<Rightarrow> 'a pol" (infixl "\<otimes>" 70) |
06165e40e7bd
switched from recdef to function package; constants add, mul, pow now curried; infix syntax for algebraic operations.
haftmann
parents:
22665
diff
changeset
|
94 |
where |
06165e40e7bd
switched from recdef to function package; constants add, mul, pow now curried; infix syntax for algebraic operations.
haftmann
parents:
22665
diff
changeset
|
95 |
"Pc a \<otimes> Pc b = Pc (a * b)" |
06165e40e7bd
switched from recdef to function package; constants add, mul, pow now curried; infix syntax for algebraic operations.
haftmann
parents:
22665
diff
changeset
|
96 |
| "Pc c \<otimes> Pinj i P = |
06165e40e7bd
switched from recdef to function package; constants add, mul, pow now curried; infix syntax for algebraic operations.
haftmann
parents:
22665
diff
changeset
|
97 |
(if c = 0 then Pc 0 else mkPinj i (P \<otimes> Pc c))" |
06165e40e7bd
switched from recdef to function package; constants add, mul, pow now curried; infix syntax for algebraic operations.
haftmann
parents:
22665
diff
changeset
|
98 |
| "Pinj i P \<otimes> Pc c = |
06165e40e7bd
switched from recdef to function package; constants add, mul, pow now curried; infix syntax for algebraic operations.
haftmann
parents:
22665
diff
changeset
|
99 |
(if c = 0 then Pc 0 else mkPinj i (P \<otimes> Pc c))" |
06165e40e7bd
switched from recdef to function package; constants add, mul, pow now curried; infix syntax for algebraic operations.
haftmann
parents:
22665
diff
changeset
|
100 |
| "Pc c \<otimes> PX P i Q = |
06165e40e7bd
switched from recdef to function package; constants add, mul, pow now curried; infix syntax for algebraic operations.
haftmann
parents:
22665
diff
changeset
|
101 |
(if c = 0 then Pc 0 else mkPX (P \<otimes> Pc c) i (Q \<otimes> Pc c))" |
06165e40e7bd
switched from recdef to function package; constants add, mul, pow now curried; infix syntax for algebraic operations.
haftmann
parents:
22665
diff
changeset
|
102 |
| "PX P i Q \<otimes> Pc c = |
06165e40e7bd
switched from recdef to function package; constants add, mul, pow now curried; infix syntax for algebraic operations.
haftmann
parents:
22665
diff
changeset
|
103 |
(if c = 0 then Pc 0 else mkPX (P \<otimes> Pc c) i (Q \<otimes> Pc c))" |
06165e40e7bd
switched from recdef to function package; constants add, mul, pow now curried; infix syntax for algebraic operations.
haftmann
parents:
22665
diff
changeset
|
104 |
| "Pinj x P \<otimes> Pinj y Q = |
06165e40e7bd
switched from recdef to function package; constants add, mul, pow now curried; infix syntax for algebraic operations.
haftmann
parents:
22665
diff
changeset
|
105 |
(if x = y then mkPinj x (P \<otimes> Q) else |
06165e40e7bd
switched from recdef to function package; constants add, mul, pow now curried; infix syntax for algebraic operations.
haftmann
parents:
22665
diff
changeset
|
106 |
(if x > y then mkPinj y (Pinj (x-y) P \<otimes> Q) |
06165e40e7bd
switched from recdef to function package; constants add, mul, pow now curried; infix syntax for algebraic operations.
haftmann
parents:
22665
diff
changeset
|
107 |
else mkPinj x (Pinj (y - x) Q \<otimes> P)))" |
06165e40e7bd
switched from recdef to function package; constants add, mul, pow now curried; infix syntax for algebraic operations.
haftmann
parents:
22665
diff
changeset
|
108 |
| "Pinj x P \<otimes> PX Q y R = |
06165e40e7bd
switched from recdef to function package; constants add, mul, pow now curried; infix syntax for algebraic operations.
haftmann
parents:
22665
diff
changeset
|
109 |
(if x = 0 then P \<otimes> PX Q y R else |
06165e40e7bd
switched from recdef to function package; constants add, mul, pow now curried; infix syntax for algebraic operations.
haftmann
parents:
22665
diff
changeset
|
110 |
(if x = 1 then mkPX (Pinj x P \<otimes> Q) y (R \<otimes> P) |
06165e40e7bd
switched from recdef to function package; constants add, mul, pow now curried; infix syntax for algebraic operations.
haftmann
parents:
22665
diff
changeset
|
111 |
else mkPX (Pinj x P \<otimes> Q) y (R \<otimes> Pinj (x - 1) P)))" |
06165e40e7bd
switched from recdef to function package; constants add, mul, pow now curried; infix syntax for algebraic operations.
haftmann
parents:
22665
diff
changeset
|
112 |
| "PX P x R \<otimes> Pinj y Q = |
06165e40e7bd
switched from recdef to function package; constants add, mul, pow now curried; infix syntax for algebraic operations.
haftmann
parents:
22665
diff
changeset
|
113 |
(if y = 0 then PX P x R \<otimes> Q else |
06165e40e7bd
switched from recdef to function package; constants add, mul, pow now curried; infix syntax for algebraic operations.
haftmann
parents:
22665
diff
changeset
|
114 |
(if y = 1 then mkPX (Pinj y Q \<otimes> P) x (R \<otimes> Q) |
06165e40e7bd
switched from recdef to function package; constants add, mul, pow now curried; infix syntax for algebraic operations.
haftmann
parents:
22665
diff
changeset
|
115 |
else mkPX (Pinj y Q \<otimes> P) x (R \<otimes> Pinj (y - 1) Q)))" |
06165e40e7bd
switched from recdef to function package; constants add, mul, pow now curried; infix syntax for algebraic operations.
haftmann
parents:
22665
diff
changeset
|
116 |
| "PX P1 x P2 \<otimes> PX Q1 y Q2 = |
06165e40e7bd
switched from recdef to function package; constants add, mul, pow now curried; infix syntax for algebraic operations.
haftmann
parents:
22665
diff
changeset
|
117 |
mkPX (P1 \<otimes> Q1) (x + y) (P2 \<otimes> Q2) \<oplus> |
06165e40e7bd
switched from recdef to function package; constants add, mul, pow now curried; infix syntax for algebraic operations.
haftmann
parents:
22665
diff
changeset
|
118 |
(mkPX (P1 \<otimes> mkPinj 1 Q2) x (Pc 0) \<oplus> |
06165e40e7bd
switched from recdef to function package; constants add, mul, pow now curried; infix syntax for algebraic operations.
haftmann
parents:
22665
diff
changeset
|
119 |
(mkPX (Q1 \<otimes> mkPinj 1 P2) y (Pc 0)))" |
06165e40e7bd
switched from recdef to function package; constants add, mul, pow now curried; infix syntax for algebraic operations.
haftmann
parents:
22665
diff
changeset
|
120 |
by pat_completeness auto |
06165e40e7bd
switched from recdef to function package; constants add, mul, pow now curried; infix syntax for algebraic operations.
haftmann
parents:
22665
diff
changeset
|
121 |
termination by (relation "measure (\<lambda>(x, y). size x + size y)") |
06165e40e7bd
switched from recdef to function package; constants add, mul, pow now curried; infix syntax for algebraic operations.
haftmann
parents:
22665
diff
changeset
|
122 |
(auto simp add: mkPinj_def split: pol.split) |
17516 | 123 |
|
124 |
text {* Negation*} |
|
22742
06165e40e7bd
switched from recdef to function package; constants add, mul, pow now curried; infix syntax for algebraic operations.
haftmann
parents:
22665
diff
changeset
|
125 |
fun |
06165e40e7bd
switched from recdef to function package; constants add, mul, pow now curried; infix syntax for algebraic operations.
haftmann
parents:
22665
diff
changeset
|
126 |
neg :: "'a::{comm_ring,recpower} pol \<Rightarrow> 'a pol" |
06165e40e7bd
switched from recdef to function package; constants add, mul, pow now curried; infix syntax for algebraic operations.
haftmann
parents:
22665
diff
changeset
|
127 |
where |
06165e40e7bd
switched from recdef to function package; constants add, mul, pow now curried; infix syntax for algebraic operations.
haftmann
parents:
22665
diff
changeset
|
128 |
"neg (Pc c) = Pc (-c)" |
06165e40e7bd
switched from recdef to function package; constants add, mul, pow now curried; infix syntax for algebraic operations.
haftmann
parents:
22665
diff
changeset
|
129 |
| "neg (Pinj i P) = Pinj i (neg P)" |
06165e40e7bd
switched from recdef to function package; constants add, mul, pow now curried; infix syntax for algebraic operations.
haftmann
parents:
22665
diff
changeset
|
130 |
| "neg (PX P x Q) = PX (neg P) x (neg Q)" |
17516 | 131 |
|
132 |
text {* Substraction *} |
|
19736 | 133 |
definition |
22742
06165e40e7bd
switched from recdef to function package; constants add, mul, pow now curried; infix syntax for algebraic operations.
haftmann
parents:
22665
diff
changeset
|
134 |
sub :: "'a::{comm_ring,recpower} pol \<Rightarrow> 'a pol \<Rightarrow> 'a pol" (infixl "\<ominus>" 65) |
06165e40e7bd
switched from recdef to function package; constants add, mul, pow now curried; infix syntax for algebraic operations.
haftmann
parents:
22665
diff
changeset
|
135 |
where |
06165e40e7bd
switched from recdef to function package; constants add, mul, pow now curried; infix syntax for algebraic operations.
haftmann
parents:
22665
diff
changeset
|
136 |
"sub P Q = P \<oplus> neg Q" |
17516 | 137 |
|
138 |
text {* Square for Fast Exponentation *} |
|
22742
06165e40e7bd
switched from recdef to function package; constants add, mul, pow now curried; infix syntax for algebraic operations.
haftmann
parents:
22665
diff
changeset
|
139 |
fun |
06165e40e7bd
switched from recdef to function package; constants add, mul, pow now curried; infix syntax for algebraic operations.
haftmann
parents:
22665
diff
changeset
|
140 |
sqr :: "'a::{comm_ring,recpower} pol \<Rightarrow> 'a pol" |
06165e40e7bd
switched from recdef to function package; constants add, mul, pow now curried; infix syntax for algebraic operations.
haftmann
parents:
22665
diff
changeset
|
141 |
where |
06165e40e7bd
switched from recdef to function package; constants add, mul, pow now curried; infix syntax for algebraic operations.
haftmann
parents:
22665
diff
changeset
|
142 |
"sqr (Pc c) = Pc (c * c)" |
06165e40e7bd
switched from recdef to function package; constants add, mul, pow now curried; infix syntax for algebraic operations.
haftmann
parents:
22665
diff
changeset
|
143 |
| "sqr (Pinj i P) = mkPinj i (sqr P)" |
06165e40e7bd
switched from recdef to function package; constants add, mul, pow now curried; infix syntax for algebraic operations.
haftmann
parents:
22665
diff
changeset
|
144 |
| "sqr (PX A x B) = mkPX (sqr A) (x + x) (sqr B) \<oplus> |
06165e40e7bd
switched from recdef to function package; constants add, mul, pow now curried; infix syntax for algebraic operations.
haftmann
parents:
22665
diff
changeset
|
145 |
mkPX (Pc (1 + 1) \<otimes> A \<otimes> mkPinj 1 B) x (Pc 0)" |
17516 | 146 |
|
147 |
text {* Fast Exponentation *} |
|
22742
06165e40e7bd
switched from recdef to function package; constants add, mul, pow now curried; infix syntax for algebraic operations.
haftmann
parents:
22665
diff
changeset
|
148 |
fun |
06165e40e7bd
switched from recdef to function package; constants add, mul, pow now curried; infix syntax for algebraic operations.
haftmann
parents:
22665
diff
changeset
|
149 |
pow :: "nat \<Rightarrow> 'a::{comm_ring,recpower} pol \<Rightarrow> 'a pol" |
06165e40e7bd
switched from recdef to function package; constants add, mul, pow now curried; infix syntax for algebraic operations.
haftmann
parents:
22665
diff
changeset
|
150 |
where |
06165e40e7bd
switched from recdef to function package; constants add, mul, pow now curried; infix syntax for algebraic operations.
haftmann
parents:
22665
diff
changeset
|
151 |
"pow 0 P = Pc 1" |
06165e40e7bd
switched from recdef to function package; constants add, mul, pow now curried; infix syntax for algebraic operations.
haftmann
parents:
22665
diff
changeset
|
152 |
| "pow n P = (if even n then pow (n div 2) (sqr P) |
06165e40e7bd
switched from recdef to function package; constants add, mul, pow now curried; infix syntax for algebraic operations.
haftmann
parents:
22665
diff
changeset
|
153 |
else P \<otimes> pow (n div 2) (sqr P))" |
06165e40e7bd
switched from recdef to function package; constants add, mul, pow now curried; infix syntax for algebraic operations.
haftmann
parents:
22665
diff
changeset
|
154 |
|
17516 | 155 |
lemma pow_if: |
22742
06165e40e7bd
switched from recdef to function package; constants add, mul, pow now curried; infix syntax for algebraic operations.
haftmann
parents:
22665
diff
changeset
|
156 |
"pow n P = |
06165e40e7bd
switched from recdef to function package; constants add, mul, pow now curried; infix syntax for algebraic operations.
haftmann
parents:
22665
diff
changeset
|
157 |
(if n = 0 then Pc 1 else if even n then pow (n div 2) (sqr P) |
06165e40e7bd
switched from recdef to function package; constants add, mul, pow now curried; infix syntax for algebraic operations.
haftmann
parents:
22665
diff
changeset
|
158 |
else P \<otimes> pow (n div 2) (sqr P))" |
17516 | 159 |
by (cases n) simp_all |
160 |
||
161 |
||
162 |
text {* Normalization of polynomial expressions *} |
|
163 |
||
22742
06165e40e7bd
switched from recdef to function package; constants add, mul, pow now curried; infix syntax for algebraic operations.
haftmann
parents:
22665
diff
changeset
|
164 |
fun |
06165e40e7bd
switched from recdef to function package; constants add, mul, pow now curried; infix syntax for algebraic operations.
haftmann
parents:
22665
diff
changeset
|
165 |
norm :: "'a::{comm_ring,recpower} polex \<Rightarrow> 'a pol" |
06165e40e7bd
switched from recdef to function package; constants add, mul, pow now curried; infix syntax for algebraic operations.
haftmann
parents:
22665
diff
changeset
|
166 |
where |
06165e40e7bd
switched from recdef to function package; constants add, mul, pow now curried; infix syntax for algebraic operations.
haftmann
parents:
22665
diff
changeset
|
167 |
"norm (Pol P) = P" |
06165e40e7bd
switched from recdef to function package; constants add, mul, pow now curried; infix syntax for algebraic operations.
haftmann
parents:
22665
diff
changeset
|
168 |
| "norm (Add P Q) = norm P \<oplus> norm Q" |
06165e40e7bd
switched from recdef to function package; constants add, mul, pow now curried; infix syntax for algebraic operations.
haftmann
parents:
22665
diff
changeset
|
169 |
| "norm (Sub P Q) = norm P \<ominus> norm Q" |
06165e40e7bd
switched from recdef to function package; constants add, mul, pow now curried; infix syntax for algebraic operations.
haftmann
parents:
22665
diff
changeset
|
170 |
| "norm (Mul P Q) = norm P \<otimes> norm Q" |
06165e40e7bd
switched from recdef to function package; constants add, mul, pow now curried; infix syntax for algebraic operations.
haftmann
parents:
22665
diff
changeset
|
171 |
| "norm (Pow P n) = pow n (norm P)" |
06165e40e7bd
switched from recdef to function package; constants add, mul, pow now curried; infix syntax for algebraic operations.
haftmann
parents:
22665
diff
changeset
|
172 |
| "norm (Neg P) = neg (norm P)" |
17516 | 173 |
|
174 |
text {* mkPinj preserve semantics *} |
|
175 |
lemma mkPinj_ci: "Ipol l (mkPinj a B) = Ipol l (Pinj a B)" |
|
29667 | 176 |
by (induct B) (auto simp add: mkPinj_def algebra_simps) |
17516 | 177 |
|
178 |
text {* mkPX preserves semantics *} |
|
179 |
lemma mkPX_ci: "Ipol l (mkPX A b C) = Ipol l (PX A b C)" |
|
29667 | 180 |
by (cases A) (auto simp add: mkPX_def mkPinj_ci power_add algebra_simps) |
17516 | 181 |
|
182 |
text {* Correctness theorems for the implemented operations *} |
|
183 |
||
184 |
text {* Negation *} |
|
20622 | 185 |
lemma neg_ci: "Ipol l (neg P) = -(Ipol l P)" |
186 |
by (induct P arbitrary: l) auto |
|
17516 | 187 |
|
188 |
text {* Addition *} |
|
22742
06165e40e7bd
switched from recdef to function package; constants add, mul, pow now curried; infix syntax for algebraic operations.
haftmann
parents:
22665
diff
changeset
|
189 |
lemma add_ci: "Ipol l (P \<oplus> Q) = Ipol l P + Ipol l Q" |
20622 | 190 |
proof (induct P Q arbitrary: l rule: add.induct) |
17516 | 191 |
case (6 x P y Q) |
192 |
show ?case |
|
193 |
proof (rule linorder_cases) |
|
194 |
assume "x < y" |
|
29667 | 195 |
with 6 show ?case by (simp add: mkPinj_ci algebra_simps) |
17516 | 196 |
next |
197 |
assume "x = y" |
|
198 |
with 6 show ?case by (simp add: mkPinj_ci) |
|
199 |
next |
|
200 |
assume "x > y" |
|
29667 | 201 |
with 6 show ?case by (simp add: mkPinj_ci algebra_simps) |
17516 | 202 |
qed |
203 |
next |
|
204 |
case (7 x P Q y R) |
|
205 |
have "x = 0 \<or> x = 1 \<or> x > 1" by arith |
|
206 |
moreover |
|
207 |
{ assume "x = 0" with 7 have ?case by simp } |
|
208 |
moreover |
|
29667 | 209 |
{ assume "x = 1" with 7 have ?case by (simp add: algebra_simps) } |
17516 | 210 |
moreover |
211 |
{ assume "x > 1" from 7 have ?case by (cases x) simp_all } |
|
212 |
ultimately show ?case by blast |
|
213 |
next |
|
214 |
case (8 P x R y Q) |
|
215 |
have "y = 0 \<or> y = 1 \<or> y > 1" by arith |
|
216 |
moreover |
|
217 |
{ assume "y = 0" with 8 have ?case by simp } |
|
218 |
moreover |
|
219 |
{ assume "y = 1" with 8 have ?case by simp } |
|
220 |
moreover |
|
221 |
{ assume "y > 1" with 8 have ?case by simp } |
|
222 |
ultimately show ?case by blast |
|
223 |
next |
|
224 |
case (9 P1 x P2 Q1 y Q2) |
|
225 |
show ?case |
|
226 |
proof (rule linorder_cases) |
|
227 |
assume a: "x < y" hence "EX d. d + x = y" by arith |
|
29667 | 228 |
with 9 a show ?case by (auto simp add: mkPX_ci power_add algebra_simps) |
17516 | 229 |
next |
230 |
assume a: "y < x" hence "EX d. d + y = x" by arith |
|
29667 | 231 |
with 9 a show ?case by (auto simp add: power_add mkPX_ci algebra_simps) |
17516 | 232 |
next |
233 |
assume "x = y" |
|
29667 | 234 |
with 9 show ?case by (simp add: mkPX_ci algebra_simps) |
17516 | 235 |
qed |
29667 | 236 |
qed (auto simp add: algebra_simps) |
17516 | 237 |
|
238 |
text {* Multiplication *} |
|
22742
06165e40e7bd
switched from recdef to function package; constants add, mul, pow now curried; infix syntax for algebraic operations.
haftmann
parents:
22665
diff
changeset
|
239 |
lemma mul_ci: "Ipol l (P \<otimes> Q) = Ipol l P * Ipol l Q" |
20622 | 240 |
by (induct P Q arbitrary: l rule: mul.induct) |
29667 | 241 |
(simp_all add: mkPX_ci mkPinj_ci algebra_simps add_ci power_add) |
17516 | 242 |
|
243 |
text {* Substraction *} |
|
22742
06165e40e7bd
switched from recdef to function package; constants add, mul, pow now curried; infix syntax for algebraic operations.
haftmann
parents:
22665
diff
changeset
|
244 |
lemma sub_ci: "Ipol l (P \<ominus> Q) = Ipol l P - Ipol l Q" |
17516 | 245 |
by (simp add: add_ci neg_ci sub_def) |
246 |
||
247 |
text {* Square *} |
|
22742
06165e40e7bd
switched from recdef to function package; constants add, mul, pow now curried; infix syntax for algebraic operations.
haftmann
parents:
22665
diff
changeset
|
248 |
lemma sqr_ci: "Ipol ls (sqr P) = Ipol ls P * Ipol ls P" |
06165e40e7bd
switched from recdef to function package; constants add, mul, pow now curried; infix syntax for algebraic operations.
haftmann
parents:
22665
diff
changeset
|
249 |
by (induct P arbitrary: ls) |
29667 | 250 |
(simp_all add: add_ci mkPinj_ci mkPX_ci mul_ci algebra_simps power_add) |
17516 | 251 |
|
252 |
text {* Power *} |
|
22742
06165e40e7bd
switched from recdef to function package; constants add, mul, pow now curried; infix syntax for algebraic operations.
haftmann
parents:
22665
diff
changeset
|
253 |
lemma even_pow:"even n \<Longrightarrow> pow n P = pow (n div 2) (sqr P)" |
20622 | 254 |
by (induct n) simp_all |
17516 | 255 |
|
22742
06165e40e7bd
switched from recdef to function package; constants add, mul, pow now curried; infix syntax for algebraic operations.
haftmann
parents:
22665
diff
changeset
|
256 |
lemma pow_ci: "Ipol ls (pow n P) = Ipol ls P ^ n" |
06165e40e7bd
switched from recdef to function package; constants add, mul, pow now curried; infix syntax for algebraic operations.
haftmann
parents:
22665
diff
changeset
|
257 |
proof (induct n arbitrary: P rule: nat_less_induct) |
17516 | 258 |
case (1 k) |
259 |
show ?case |
|
260 |
proof (cases k) |
|
20622 | 261 |
case 0 |
262 |
then show ?thesis by simp |
|
263 |
next |
|
17516 | 264 |
case (Suc l) |
265 |
show ?thesis |
|
266 |
proof cases |
|
20622 | 267 |
assume "even l" |
268 |
then have "Suc l div 2 = l div 2" |
|
269 |
by (simp add: nat_number even_nat_plus_one_div_two) |
|
17516 | 270 |
moreover |
271 |
from Suc have "l < k" by simp |
|
22742
06165e40e7bd
switched from recdef to function package; constants add, mul, pow now curried; infix syntax for algebraic operations.
haftmann
parents:
22665
diff
changeset
|
272 |
with 1 have "\<And>P. Ipol ls (pow l P) = Ipol ls P ^ l" by simp |
17516 | 273 |
moreover |
20622 | 274 |
note Suc `even l` even_nat_plus_one_div_two |
17516 | 275 |
ultimately show ?thesis by (auto simp add: mul_ci power_Suc even_pow) |
276 |
next |
|
20622 | 277 |
assume "odd l" |
278 |
{ |
|
279 |
fix p |
|
22742
06165e40e7bd
switched from recdef to function package; constants add, mul, pow now curried; infix syntax for algebraic operations.
haftmann
parents:
22665
diff
changeset
|
280 |
have "Ipol ls (sqr P) ^ (Suc l div 2) = Ipol ls P ^ Suc l" |
20622 | 281 |
proof (cases l) |
282 |
case 0 |
|
283 |
with `odd l` show ?thesis by simp |
|
284 |
next |
|
285 |
case (Suc w) |
|
286 |
with `odd l` have "even w" by simp |
|
20678 | 287 |
have two_times: "2 * (w div 2) = w" |
288 |
by (simp only: numerals even_nat_div_two_times_two [OF `even w`]) |
|
22742
06165e40e7bd
switched from recdef to function package; constants add, mul, pow now curried; infix syntax for algebraic operations.
haftmann
parents:
22665
diff
changeset
|
289 |
have "Ipol ls P * Ipol ls P = Ipol ls P ^ Suc (Suc 0)" |
20622 | 290 |
by (simp add: power_Suc) |
22742
06165e40e7bd
switched from recdef to function package; constants add, mul, pow now curried; infix syntax for algebraic operations.
haftmann
parents:
22665
diff
changeset
|
291 |
then have "Ipol ls P * Ipol ls P = Ipol ls P ^ 2" |
20678 | 292 |
by (simp add: numerals) |
20622 | 293 |
with Suc show ?thesis |
20678 | 294 |
by (auto simp add: power_mult [symmetric, of _ 2 _] two_times mul_ci sqr_ci) |
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 - |
|
20622 | 310 |
from prems have "Ipol l (norm P1) = Ipol l (norm P2)" by simp |
311 |
then show ?thesis by (simp only: norm_ci) |
|
17516 | 312 |
qed |
313 |
||
314 |
||
315 |
use "comm_ring.ML" |
|
18708 | 316 |
setup CommRing.setup |
17516 | 317 |
|
318 |
end |