author | haftmann |
Sat, 01 Oct 2022 13:08:34 +0000 | |
changeset 76232 | a7ccb744047b |
parent 67123 | 3fe40ff1b921 |
child 80105 | 2fa018321400 |
permissions | -rw-r--r-- |
31021 | 1 |
(* Author: Bernhard Haeupler |
17378
105519771c67
The oracle for Presburger has been changer: It is automatically generated form a verified formaliztion of Cooper's Algorithm ex/Reflected_Presburger.thy
chaieb
parents:
diff
changeset
|
2 |
|
64962
bf41e1109db3
Added new / improved tactics for fields and rings
berghofe
parents:
61586
diff
changeset
|
3 |
This theory is about of the relative completeness of the ring |
bf41e1109db3
Added new / improved tactics for fields and rings
berghofe
parents:
61586
diff
changeset
|
4 |
method. As long as the reified atomic polynomials of type pol are |
bf41e1109db3
Added new / improved tactics for fields and rings
berghofe
parents:
61586
diff
changeset
|
5 |
in normal form, the ring method is complete. |
17388 | 6 |
*) |
7 |
||
60533 | 8 |
section \<open>Proof of the relative completeness of method comm-ring\<close> |
17378
105519771c67
The oracle for Presburger has been changer: It is automatically generated form a verified formaliztion of Cooper's Algorithm ex/Reflected_Presburger.thy
chaieb
parents:
diff
changeset
|
9 |
|
105519771c67
The oracle for Presburger has been changer: It is automatically generated form a verified formaliztion of Cooper's Algorithm ex/Reflected_Presburger.thy
chaieb
parents:
diff
changeset
|
10 |
theory Commutative_Ring_Complete |
67123 | 11 |
imports Commutative_Ring |
17378
105519771c67
The oracle for Presburger has been changer: It is automatically generated form a verified formaliztion of Cooper's Algorithm ex/Reflected_Presburger.thy
chaieb
parents:
diff
changeset
|
12 |
begin |
22742
06165e40e7bd
switched from recdef to function package; constants add, mul, pow now curried; infix syntax for algebraic operations.
haftmann
parents:
17508
diff
changeset
|
13 |
|
60533 | 14 |
text \<open>Formalization of normal form\<close> |
64962
bf41e1109db3
Added new / improved tactics for fields and rings
berghofe
parents:
61586
diff
changeset
|
15 |
fun isnorm :: "pol \<Rightarrow> bool" |
67123 | 16 |
where |
17 |
"isnorm (Pc c) \<longleftrightarrow> True" |
|
18 |
| "isnorm (Pinj i (Pc c)) \<longleftrightarrow> False" |
|
19 |
| "isnorm (Pinj i (Pinj j Q)) \<longleftrightarrow> False" |
|
20 |
| "isnorm (Pinj 0 P) \<longleftrightarrow> False" |
|
21 |
| "isnorm (Pinj i (PX Q1 j Q2)) \<longleftrightarrow> isnorm (PX Q1 j Q2)" |
|
22 |
| "isnorm (PX P 0 Q) \<longleftrightarrow> False" |
|
23 |
| "isnorm (PX (Pc c) i Q) \<longleftrightarrow> c \<noteq> 0 \<and> isnorm Q" |
|
24 |
| "isnorm (PX (PX P1 j (Pc c)) i Q) \<longleftrightarrow> c \<noteq> 0 \<and> isnorm (PX P1 j (Pc c)) \<and> isnorm Q" |
|
25 |
| "isnorm (PX P i Q) \<longleftrightarrow> isnorm P \<and> isnorm Q" |
|
17378
105519771c67
The oracle for Presburger has been changer: It is automatically generated form a verified formaliztion of Cooper's Algorithm ex/Reflected_Presburger.thy
chaieb
parents:
diff
changeset
|
26 |
|
67123 | 27 |
|
28 |
subsection \<open>Some helpful lemmas\<close> |
|
29 |
||
44779 | 30 |
lemma norm_Pinj_0_False: "isnorm (Pinj 0 P) = False" |
31 |
by (cases P) auto |
|
17378
105519771c67
The oracle for Presburger has been changer: It is automatically generated form a verified formaliztion of Cooper's Algorithm ex/Reflected_Presburger.thy
chaieb
parents:
diff
changeset
|
32 |
|
44779 | 33 |
lemma norm_PX_0_False: "isnorm (PX (Pc 0) i Q) = False" |
34 |
by (cases i) auto |
|
17378
105519771c67
The oracle for Presburger has been changer: It is automatically generated form a verified formaliztion of Cooper's Algorithm ex/Reflected_Presburger.thy
chaieb
parents:
diff
changeset
|
35 |
|
44779 | 36 |
lemma norm_Pinj: "isnorm (Pinj i Q) \<Longrightarrow> isnorm Q" |
37 |
by (cases i) (simp add: norm_Pinj_0_False norm_PX_0_False, cases Q, auto) |
|
17378
105519771c67
The oracle for Presburger has been changer: It is automatically generated form a verified formaliztion of Cooper's Algorithm ex/Reflected_Presburger.thy
chaieb
parents:
diff
changeset
|
38 |
|
44779 | 39 |
lemma norm_PX2: "isnorm (PX P i Q) \<Longrightarrow> isnorm Q" |
55754 | 40 |
apply (cases i) |
67123 | 41 |
apply auto |
55754 | 42 |
apply (cases P) |
67123 | 43 |
apply auto |
60708 | 44 |
subgoal for \<dots> pol2 by (cases pol2) auto |
55754 | 45 |
done |
44779 | 46 |
|
47 |
lemma norm_PX1: "isnorm (PX P i Q) \<Longrightarrow> isnorm P" |
|
55754 | 48 |
apply (cases i) |
67123 | 49 |
apply auto |
55754 | 50 |
apply (cases P) |
67123 | 51 |
apply auto |
60708 | 52 |
subgoal for \<dots> pol2 by (cases pol2) auto |
55754 | 53 |
done |
17378
105519771c67
The oracle for Presburger has been changer: It is automatically generated form a verified formaliztion of Cooper's Algorithm ex/Reflected_Presburger.thy
chaieb
parents:
diff
changeset
|
54 |
|
55754 | 55 |
lemma mkPinj_cn: "y \<noteq> 0 \<Longrightarrow> isnorm Q \<Longrightarrow> isnorm (mkPinj y Q)" |
44779 | 56 |
apply (auto simp add: mkPinj_def norm_Pinj_0_False split: pol.split) |
67123 | 57 |
apply (rename_tac nat a, case_tac nat, auto simp add: norm_Pinj_0_False) |
58 |
apply (rename_tac pol a, case_tac pol, auto) |
|
44779 | 59 |
apply (case_tac y, auto) |
60 |
done |
|
17378
105519771c67
The oracle for Presburger has been changer: It is automatically generated form a verified formaliztion of Cooper's Algorithm ex/Reflected_Presburger.thy
chaieb
parents:
diff
changeset
|
61 |
|
55793 | 62 |
lemma norm_PXtrans: |
55754 | 63 |
assumes A: "isnorm (PX P x Q)" |
55793 | 64 |
and "isnorm Q2" |
17378
105519771c67
The oracle for Presburger has been changer: It is automatically generated form a verified formaliztion of Cooper's Algorithm ex/Reflected_Presburger.thy
chaieb
parents:
diff
changeset
|
65 |
shows "isnorm (PX P x Q2)" |
44779 | 66 |
proof (cases P) |
67 |
case (PX p1 y p2) |
|
55754 | 68 |
with assms show ?thesis |
69 |
apply (cases x) |
|
67123 | 70 |
apply auto |
55754 | 71 |
apply (cases p2) |
67123 | 72 |
apply auto |
55754 | 73 |
done |
17378
105519771c67
The oracle for Presburger has been changer: It is automatically generated form a verified formaliztion of Cooper's Algorithm ex/Reflected_Presburger.thy
chaieb
parents:
diff
changeset
|
74 |
next |
44779 | 75 |
case Pc |
60535 | 76 |
with assms show ?thesis |
77 |
by (cases x) auto |
|
17378
105519771c67
The oracle for Presburger has been changer: It is automatically generated form a verified formaliztion of Cooper's Algorithm ex/Reflected_Presburger.thy
chaieb
parents:
diff
changeset
|
78 |
next |
44779 | 79 |
case Pinj |
60535 | 80 |
with assms show ?thesis |
81 |
by (cases x) auto |
|
17378
105519771c67
The oracle for Presburger has been changer: It is automatically generated form a verified formaliztion of Cooper's Algorithm ex/Reflected_Presburger.thy
chaieb
parents:
diff
changeset
|
82 |
qed |
105519771c67
The oracle for Presburger has been changer: It is automatically generated form a verified formaliztion of Cooper's Algorithm ex/Reflected_Presburger.thy
chaieb
parents:
diff
changeset
|
83 |
|
55793 | 84 |
lemma norm_PXtrans2: |
85 |
assumes "isnorm (PX P x Q)" |
|
86 |
and "isnorm Q2" |
|
87 |
shows "isnorm (PX P (Suc (n + x)) Q2)" |
|
88 |
proof (cases P) |
|
89 |
case (PX p1 y p2) |
|
90 |
with assms show ?thesis |
|
91 |
apply (cases x) |
|
67123 | 92 |
apply auto |
55793 | 93 |
apply (cases p2) |
67123 | 94 |
apply auto |
55793 | 95 |
done |
96 |
next |
|
97 |
case Pc |
|
98 |
with assms show ?thesis |
|
99 |
by (cases x) auto |
|
100 |
next |
|
101 |
case Pinj |
|
102 |
with assms show ?thesis |
|
103 |
by (cases x) auto |
|
104 |
qed |
|
105 |
||
67123 | 106 |
text \<open>\<open>mkPX\<close> conserves normalizedness (\<open>_cn\<close>)\<close> |
55793 | 107 |
lemma mkPX_cn: |
55754 | 108 |
assumes "x \<noteq> 0" |
109 |
and "isnorm P" |
|
55793 | 110 |
and "isnorm Q" |
17378
105519771c67
The oracle for Presburger has been changer: It is automatically generated form a verified formaliztion of Cooper's Algorithm ex/Reflected_Presburger.thy
chaieb
parents:
diff
changeset
|
111 |
shows "isnorm (mkPX P x Q)" |
55754 | 112 |
proof (cases P) |
17378
105519771c67
The oracle for Presburger has been changer: It is automatically generated form a verified formaliztion of Cooper's Algorithm ex/Reflected_Presburger.thy
chaieb
parents:
diff
changeset
|
113 |
case (Pc c) |
55754 | 114 |
with assms show ?thesis |
115 |
by (cases x) (auto simp add: mkPinj_cn mkPX_def) |
|
17378
105519771c67
The oracle for Presburger has been changer: It is automatically generated form a verified formaliztion of Cooper's Algorithm ex/Reflected_Presburger.thy
chaieb
parents:
diff
changeset
|
116 |
next |
105519771c67
The oracle for Presburger has been changer: It is automatically generated form a verified formaliztion of Cooper's Algorithm ex/Reflected_Presburger.thy
chaieb
parents:
diff
changeset
|
117 |
case (Pinj i Q) |
55754 | 118 |
with assms show ?thesis |
119 |
by (cases x) (auto simp add: mkPinj_cn mkPX_def) |
|
17378
105519771c67
The oracle for Presburger has been changer: It is automatically generated form a verified formaliztion of Cooper's Algorithm ex/Reflected_Presburger.thy
chaieb
parents:
diff
changeset
|
120 |
next |
105519771c67
The oracle for Presburger has been changer: It is automatically generated form a verified formaliztion of Cooper's Algorithm ex/Reflected_Presburger.thy
chaieb
parents:
diff
changeset
|
121 |
case (PX P1 y P2) |
55793 | 122 |
with assms have Y0: "y > 0" |
123 |
by (cases y) auto |
|
41807 | 124 |
from assms PX have "isnorm P1" "isnorm P2" |
125 |
by (auto simp add: norm_PX1[of P1 y P2] norm_PX2[of P1 y P2]) |
|
126 |
from assms PX Y0 show ?thesis |
|
60535 | 127 |
apply (cases x) |
67123 | 128 |
apply (auto simp add: mkPX_def norm_PXtrans2[of P1 y _ Q _]) |
60535 | 129 |
apply (cases P2) |
67123 | 130 |
apply auto |
60535 | 131 |
done |
17378
105519771c67
The oracle for Presburger has been changer: It is automatically generated form a verified formaliztion of Cooper's Algorithm ex/Reflected_Presburger.thy
chaieb
parents:
diff
changeset
|
132 |
qed |
105519771c67
The oracle for Presburger has been changer: It is automatically generated form a verified formaliztion of Cooper's Algorithm ex/Reflected_Presburger.thy
chaieb
parents:
diff
changeset
|
133 |
|
67123 | 134 |
text \<open>\<open>add\<close> conserves normalizedness\<close> |
64962
bf41e1109db3
Added new / improved tactics for fields and rings
berghofe
parents:
61586
diff
changeset
|
135 |
lemma add_cn: "isnorm P \<Longrightarrow> isnorm Q \<Longrightarrow> isnorm (P \<langle>+\<rangle> Q)" |
44779 | 136 |
proof (induct P Q rule: add.induct) |
60535 | 137 |
case 1 |
138 |
then show ?case by simp |
|
139 |
next |
|
44779 | 140 |
case (2 c i P2) |
55793 | 141 |
then show ?case |
60535 | 142 |
apply (cases P2) |
67123 | 143 |
apply simp_all |
60535 | 144 |
apply (cases i) |
67123 | 145 |
apply simp_all |
60535 | 146 |
done |
17378
105519771c67
The oracle for Presburger has been changer: It is automatically generated form a verified formaliztion of Cooper's Algorithm ex/Reflected_Presburger.thy
chaieb
parents:
diff
changeset
|
147 |
next |
44779 | 148 |
case (3 i P2 c) |
55793 | 149 |
then show ?case |
60535 | 150 |
apply (cases P2) |
67123 | 151 |
apply simp_all |
60535 | 152 |
apply (cases i) |
67123 | 153 |
apply simp_all |
60535 | 154 |
done |
17378
105519771c67
The oracle for Presburger has been changer: It is automatically generated form a verified formaliztion of Cooper's Algorithm ex/Reflected_Presburger.thy
chaieb
parents:
diff
changeset
|
155 |
next |
105519771c67
The oracle for Presburger has been changer: It is automatically generated form a verified formaliztion of Cooper's Algorithm ex/Reflected_Presburger.thy
chaieb
parents:
diff
changeset
|
156 |
case (4 c P2 i Q2) |
44779 | 157 |
then have "isnorm P2" "isnorm Q2" |
158 |
by (auto simp only: norm_PX1[of P2 i Q2] norm_PX2[of P2 i Q2]) |
|
159 |
with 4 show ?case |
|
60535 | 160 |
apply (cases i) |
67123 | 161 |
apply simp |
60535 | 162 |
apply (cases P2) |
67123 | 163 |
apply auto |
60708 | 164 |
subgoal for \<dots> pol2 by (cases pol2) auto |
60535 | 165 |
done |
17378
105519771c67
The oracle for Presburger has been changer: It is automatically generated form a verified formaliztion of Cooper's Algorithm ex/Reflected_Presburger.thy
chaieb
parents:
diff
changeset
|
166 |
next |
105519771c67
The oracle for Presburger has been changer: It is automatically generated form a verified formaliztion of Cooper's Algorithm ex/Reflected_Presburger.thy
chaieb
parents:
diff
changeset
|
167 |
case (5 P2 i Q2 c) |
44779 | 168 |
then have "isnorm P2" "isnorm Q2" |
169 |
by (auto simp only: norm_PX1[of P2 i Q2] norm_PX2[of P2 i Q2]) |
|
170 |
with 5 show ?case |
|
60535 | 171 |
apply (cases i) |
67123 | 172 |
apply simp |
60535 | 173 |
apply (cases P2) |
67123 | 174 |
apply auto |
60708 | 175 |
subgoal for \<dots> pol2 by (cases pol2) auto |
60535 | 176 |
done |
17378
105519771c67
The oracle for Presburger has been changer: It is automatically generated form a verified formaliztion of Cooper's Algorithm ex/Reflected_Presburger.thy
chaieb
parents:
diff
changeset
|
177 |
next |
105519771c67
The oracle for Presburger has been changer: It is automatically generated form a verified formaliztion of Cooper's Algorithm ex/Reflected_Presburger.thy
chaieb
parents:
diff
changeset
|
178 |
case (6 x P2 y Q2) |
55793 | 179 |
then have Y0: "y > 0" |
180 |
by (cases y) (auto simp add: norm_Pinj_0_False) |
|
181 |
with 6 have X0: "x > 0" |
|
182 |
by (cases x) (auto simp add: norm_Pinj_0_False) |
|
60535 | 183 |
consider "x < y" | "x = y" | "x > y" by arith |
184 |
then show ?case |
|
185 |
proof cases |
|
60567 | 186 |
case xy: 1 |
60535 | 187 |
then obtain d where y: "y = d + x" |
188 |
by atomize_elim arith |
|
189 |
from 6 have *: "isnorm P2" "isnorm Q2" |
|
44779 | 190 |
by (auto simp add: norm_Pinj[of _ P2] norm_Pinj[of _ Q2]) |
60567 | 191 |
from 6 xy y have "isnorm (Pinj d Q2)" |
44779 | 192 |
by (cases d, simp, cases Q2, auto) |
60535 | 193 |
with 6 X0 y * show ?thesis |
194 |
by (simp add: mkPinj_cn) |
|
195 |
next |
|
60567 | 196 |
case xy: 2 |
44779 | 197 |
from 6 have "isnorm P2" "isnorm Q2" |
198 |
by (auto simp add: norm_Pinj[of _ P2] norm_Pinj[of _ Q2]) |
|
60567 | 199 |
with xy 6 Y0 show ?thesis |
60535 | 200 |
by (simp add: mkPinj_cn) |
201 |
next |
|
60567 | 202 |
case xy: 3 |
60535 | 203 |
then obtain d where x: "x = d + y" |
204 |
by atomize_elim arith |
|
205 |
from 6 have *: "isnorm P2" "isnorm Q2" |
|
44779 | 206 |
by (auto simp add: norm_Pinj[of _ P2] norm_Pinj[of _ Q2]) |
60567 | 207 |
from 6 xy x have "isnorm (Pinj d P2)" |
44779 | 208 |
by (cases d) (simp, cases P2, auto) |
60567 | 209 |
with xy 6 Y0 * x show ?thesis by (simp add: mkPinj_cn) |
60535 | 210 |
qed |
17378
105519771c67
The oracle for Presburger has been changer: It is automatically generated form a verified formaliztion of Cooper's Algorithm ex/Reflected_Presburger.thy
chaieb
parents:
diff
changeset
|
211 |
next |
105519771c67
The oracle for Presburger has been changer: It is automatically generated form a verified formaliztion of Cooper's Algorithm ex/Reflected_Presburger.thy
chaieb
parents:
diff
changeset
|
212 |
case (7 x P2 Q2 y R) |
60535 | 213 |
consider "x = 0" | "x = 1" | d where "x = Suc (Suc d)" |
214 |
by atomize_elim arith |
|
215 |
then show ?case |
|
216 |
proof cases |
|
60767 | 217 |
case 1 |
60535 | 218 |
with 7 show ?thesis |
219 |
by (auto simp add: norm_Pinj_0_False) |
|
220 |
next |
|
60567 | 221 |
case x: 2 |
44779 | 222 |
from 7 have "isnorm R" "isnorm P2" |
223 |
by (auto simp add: norm_Pinj[of _ P2] norm_PX2[of Q2 y R]) |
|
64962
bf41e1109db3
Added new / improved tactics for fields and rings
berghofe
parents:
61586
diff
changeset
|
224 |
with 7 x have "isnorm (R \<langle>+\<rangle> P2)" |
60535 | 225 |
by simp |
60567 | 226 |
with 7 x show ?thesis |
55793 | 227 |
by (simp add: norm_PXtrans[of Q2 y _]) |
60535 | 228 |
next |
60567 | 229 |
case x: 3 |
41807 | 230 |
with 7 have NR: "isnorm R" "isnorm P2" |
231 |
by (auto simp add: norm_Pinj[of _ P2] norm_PX2[of Q2 y R]) |
|
60567 | 232 |
with 7 x have "isnorm (Pinj (x - 1) P2)" |
60535 | 233 |
by (cases P2) auto |
64962
bf41e1109db3
Added new / improved tactics for fields and rings
berghofe
parents:
61586
diff
changeset
|
234 |
with 7 x NR have "isnorm (R \<langle>+\<rangle> Pinj (x - 1) P2)" |
60535 | 235 |
by simp |
60567 | 236 |
with \<open>isnorm (PX Q2 y R)\<close> x show ?thesis |
55793 | 237 |
by (simp add: norm_PXtrans[of Q2 y _]) |
60535 | 238 |
qed |
17378
105519771c67
The oracle for Presburger has been changer: It is automatically generated form a verified formaliztion of Cooper's Algorithm ex/Reflected_Presburger.thy
chaieb
parents:
diff
changeset
|
239 |
next |
105519771c67
The oracle for Presburger has been changer: It is automatically generated form a verified formaliztion of Cooper's Algorithm ex/Reflected_Presburger.thy
chaieb
parents:
diff
changeset
|
240 |
case (8 Q2 y R x P2) |
60535 | 241 |
consider "x = 0" | "x = 1" | "x > 1" |
242 |
by arith |
|
243 |
then show ?case |
|
244 |
proof cases |
|
245 |
case 1 |
|
246 |
with 8 show ?thesis |
|
55793 | 247 |
by (auto simp add: norm_Pinj_0_False) |
60535 | 248 |
next |
60567 | 249 |
case x: 2 |
55793 | 250 |
with 8 have "isnorm R" "isnorm P2" |
251 |
by (auto simp add: norm_Pinj[of _ P2] norm_PX2[of Q2 y R]) |
|
64962
bf41e1109db3
Added new / improved tactics for fields and rings
berghofe
parents:
61586
diff
changeset
|
252 |
with 8 x have "isnorm (R \<langle>+\<rangle> P2)" |
55793 | 253 |
by simp |
60567 | 254 |
with 8 x show ?thesis |
55793 | 255 |
by (simp add: norm_PXtrans[of Q2 y _]) |
60535 | 256 |
next |
60567 | 257 |
case x: 3 |
60535 | 258 |
then obtain d where x: "x = Suc (Suc d)" by atomize_elim arith |
41807 | 259 |
with 8 have NR: "isnorm R" "isnorm P2" |
260 |
by (auto simp add: norm_Pinj[of _ P2] norm_PX2[of Q2 y R]) |
|
60535 | 261 |
with 8 x have "isnorm (Pinj (x - 1) P2)" |
55793 | 262 |
by (cases P2) auto |
64962
bf41e1109db3
Added new / improved tactics for fields and rings
berghofe
parents:
61586
diff
changeset
|
263 |
with 8 x NR have "isnorm (R \<langle>+\<rangle> Pinj (x - 1) P2)" |
55793 | 264 |
by simp |
60535 | 265 |
with \<open>isnorm (PX Q2 y R)\<close> x show ?thesis |
55793 | 266 |
by (simp add: norm_PXtrans[of Q2 y _]) |
60535 | 267 |
qed |
17378
105519771c67
The oracle for Presburger has been changer: It is automatically generated form a verified formaliztion of Cooper's Algorithm ex/Reflected_Presburger.thy
chaieb
parents:
diff
changeset
|
268 |
next |
105519771c67
The oracle for Presburger has been changer: It is automatically generated form a verified formaliztion of Cooper's Algorithm ex/Reflected_Presburger.thy
chaieb
parents:
diff
changeset
|
269 |
case (9 P1 x P2 Q1 y Q2) |
55793 | 270 |
then have Y0: "y > 0" by (cases y) auto |
271 |
with 9 have X0: "x > 0" by (cases x) auto |
|
41807 | 272 |
with 9 have NP1: "isnorm P1" and NP2: "isnorm P2" |
273 |
by (auto simp add: norm_PX1[of P1 _ P2] norm_PX2[of P1 _ P2]) |
|
44779 | 274 |
with 9 have NQ1: "isnorm Q1" and NQ2: "isnorm Q2" |
41807 | 275 |
by (auto simp add: norm_PX1[of Q1 _ Q2] norm_PX2[of Q1 _ Q2]) |
60535 | 276 |
consider "y < x" | "x = y" | "x < y" by arith |
277 |
then show ?case |
|
278 |
proof cases |
|
60567 | 279 |
case xy: 1 |
60535 | 280 |
then obtain d where x: "x = d + y" |
281 |
by atomize_elim arith |
|
55793 | 282 |
have "isnorm (PX P1 d (Pc 0))" |
41807 | 283 |
proof (cases P1) |
17378
105519771c67
The oracle for Presburger has been changer: It is automatically generated form a verified formaliztion of Cooper's Algorithm ex/Reflected_Presburger.thy
chaieb
parents:
diff
changeset
|
284 |
case (PX p1 y p2) |
60567 | 285 |
with 9 xy x show ?thesis |
55793 | 286 |
by (cases d) (simp, cases p2, auto) |
41807 | 287 |
next |
55793 | 288 |
case Pc |
60567 | 289 |
with 9 xy x show ?thesis |
55793 | 290 |
by (cases d) auto |
41807 | 291 |
next |
55793 | 292 |
case Pinj |
60567 | 293 |
with 9 xy x show ?thesis |
55793 | 294 |
by (cases d) auto |
17378
105519771c67
The oracle for Presburger has been changer: It is automatically generated form a verified formaliztion of Cooper's Algorithm ex/Reflected_Presburger.thy
chaieb
parents:
diff
changeset
|
295 |
qed |
64962
bf41e1109db3
Added new / improved tactics for fields and rings
berghofe
parents:
61586
diff
changeset
|
296 |
with 9 NQ1 NP1 NP2 NQ2 xy x have "isnorm (P2 \<langle>+\<rangle> Q2)" "isnorm (PX P1 (x - y) (Pc 0) \<langle>+\<rangle> Q1)" |
55793 | 297 |
by auto |
60567 | 298 |
with Y0 xy x show ?thesis |
55793 | 299 |
by (simp add: mkPX_cn) |
60535 | 300 |
next |
60567 | 301 |
case xy: 2 |
64962
bf41e1109db3
Added new / improved tactics for fields and rings
berghofe
parents:
61586
diff
changeset
|
302 |
with 9 NP1 NP2 NQ1 NQ2 have "isnorm (P2 \<langle>+\<rangle> Q2)" "isnorm (P1 \<langle>+\<rangle> Q1)" |
55793 | 303 |
by auto |
60567 | 304 |
with xy Y0 show ?thesis |
55793 | 305 |
by (simp add: mkPX_cn) |
60535 | 306 |
next |
60567 | 307 |
case xy: 3 |
60535 | 308 |
then obtain d where y: "y = d + x" |
309 |
by atomize_elim arith |
|
55793 | 310 |
have "isnorm (PX Q1 d (Pc 0))" |
41807 | 311 |
proof (cases Q1) |
17378
105519771c67
The oracle for Presburger has been changer: It is automatically generated form a verified formaliztion of Cooper's Algorithm ex/Reflected_Presburger.thy
chaieb
parents:
diff
changeset
|
312 |
case (PX p1 y p2) |
60567 | 313 |
with 9 xy y show ?thesis |
55793 | 314 |
by (cases d) (simp, cases p2, auto) |
41807 | 315 |
next |
55793 | 316 |
case Pc |
60567 | 317 |
with 9 xy y show ?thesis |
55793 | 318 |
by (cases d) auto |
41807 | 319 |
next |
55793 | 320 |
case Pinj |
60567 | 321 |
with 9 xy y show ?thesis |
55793 | 322 |
by (cases d) auto |
17378
105519771c67
The oracle for Presburger has been changer: It is automatically generated form a verified formaliztion of Cooper's Algorithm ex/Reflected_Presburger.thy
chaieb
parents:
diff
changeset
|
323 |
qed |
64962
bf41e1109db3
Added new / improved tactics for fields and rings
berghofe
parents:
61586
diff
changeset
|
324 |
with 9 NQ1 NP1 NP2 NQ2 xy y have "isnorm (P2 \<langle>+\<rangle> Q2)" "isnorm (PX Q1 (y - x) (Pc 0) \<langle>+\<rangle> P1)" |
55793 | 325 |
by auto |
60567 | 326 |
with X0 xy y show ?thesis |
55793 | 327 |
by (simp add: mkPX_cn) |
60535 | 328 |
qed |
329 |
qed |
|
17378
105519771c67
The oracle for Presburger has been changer: It is automatically generated form a verified formaliztion of Cooper's Algorithm ex/Reflected_Presburger.thy
chaieb
parents:
diff
changeset
|
330 |
|
67123 | 331 |
text \<open>\<open>mul\<close> concerves normalizedness\<close> |
64962
bf41e1109db3
Added new / improved tactics for fields and rings
berghofe
parents:
61586
diff
changeset
|
332 |
lemma mul_cn: "isnorm P \<Longrightarrow> isnorm Q \<Longrightarrow> isnorm (P \<langle>*\<rangle> Q)" |
44779 | 333 |
proof (induct P Q rule: mul.induct) |
60535 | 334 |
case 1 |
335 |
then show ?case by simp |
|
336 |
next |
|
55793 | 337 |
case (2 c i P2) |
338 |
then show ?case |
|
60535 | 339 |
apply (cases P2) |
67123 | 340 |
apply simp_all |
60535 | 341 |
apply (cases i) |
67123 | 342 |
apply (simp_all add: mkPinj_cn) |
60535 | 343 |
done |
17378
105519771c67
The oracle for Presburger has been changer: It is automatically generated form a verified formaliztion of Cooper's Algorithm ex/Reflected_Presburger.thy
chaieb
parents:
diff
changeset
|
344 |
next |
55793 | 345 |
case (3 i P2 c) |
346 |
then show ?case |
|
60535 | 347 |
apply (cases P2) |
67123 | 348 |
apply simp_all |
60535 | 349 |
apply (cases i) |
67123 | 350 |
apply (simp_all add: mkPinj_cn) |
60535 | 351 |
done |
17378
105519771c67
The oracle for Presburger has been changer: It is automatically generated form a verified formaliztion of Cooper's Algorithm ex/Reflected_Presburger.thy
chaieb
parents:
diff
changeset
|
352 |
next |
105519771c67
The oracle for Presburger has been changer: It is automatically generated form a verified formaliztion of Cooper's Algorithm ex/Reflected_Presburger.thy
chaieb
parents:
diff
changeset
|
353 |
case (4 c P2 i Q2) |
44779 | 354 |
then have "isnorm P2" "isnorm Q2" |
355 |
by (auto simp only: norm_PX1[of P2 i Q2] norm_PX2[of P2 i Q2]) |
|
55793 | 356 |
with 4 show ?case |
60535 | 357 |
apply (cases "c = 0") |
67123 | 358 |
apply simp_all |
60535 | 359 |
apply (cases "i = 0") |
67123 | 360 |
apply (simp_all add: mkPX_cn) |
60535 | 361 |
done |
17378
105519771c67
The oracle for Presburger has been changer: It is automatically generated form a verified formaliztion of Cooper's Algorithm ex/Reflected_Presburger.thy
chaieb
parents:
diff
changeset
|
362 |
next |
105519771c67
The oracle for Presburger has been changer: It is automatically generated form a verified formaliztion of Cooper's Algorithm ex/Reflected_Presburger.thy
chaieb
parents:
diff
changeset
|
363 |
case (5 P2 i Q2 c) |
44779 | 364 |
then have "isnorm P2" "isnorm Q2" |
365 |
by (auto simp only: norm_PX1[of P2 i Q2] norm_PX2[of P2 i Q2]) |
|
41807 | 366 |
with 5 show ?case |
60535 | 367 |
apply (cases "c = 0") |
67123 | 368 |
apply simp_all |
60535 | 369 |
apply (cases "i = 0") |
67123 | 370 |
apply (simp_all add: mkPX_cn) |
60535 | 371 |
done |
17378
105519771c67
The oracle for Presburger has been changer: It is automatically generated form a verified formaliztion of Cooper's Algorithm ex/Reflected_Presburger.thy
chaieb
parents:
diff
changeset
|
372 |
next |
105519771c67
The oracle for Presburger has been changer: It is automatically generated form a verified formaliztion of Cooper's Algorithm ex/Reflected_Presburger.thy
chaieb
parents:
diff
changeset
|
373 |
case (6 x P2 y Q2) |
60535 | 374 |
consider "x < y" | "x = y" | "x > y" by arith |
375 |
then show ?case |
|
376 |
proof cases |
|
60567 | 377 |
case xy: 1 |
60535 | 378 |
then obtain d where y: "y = d + x" |
379 |
by atomize_elim arith |
|
380 |
from 6 have *: "x > 0" |
|
55793 | 381 |
by (cases x) (auto simp add: norm_Pinj_0_False) |
60535 | 382 |
from 6 have **: "isnorm P2" "isnorm Q2" |
55793 | 383 |
by (auto simp add: norm_Pinj[of _ P2] norm_Pinj[of _ Q2]) |
60567 | 384 |
from 6 xy y have "isnorm (Pinj d Q2)" |
60535 | 385 |
apply (cases d) |
67123 | 386 |
apply simp |
60535 | 387 |
apply (cases Q2) |
67123 | 388 |
apply auto |
60535 | 389 |
done |
390 |
with 6 * ** y show ?thesis |
|
391 |
by (simp add: mkPinj_cn) |
|
392 |
next |
|
60567 | 393 |
case xy: 2 |
60535 | 394 |
from 6 have *: "isnorm P2" "isnorm Q2" |
55793 | 395 |
by (auto simp add: norm_Pinj[of _ P2] norm_Pinj[of _ Q2]) |
60535 | 396 |
from 6 have **: "y > 0" |
397 |
by (cases y) (auto simp add: norm_Pinj_0_False) |
|
60567 | 398 |
with 6 xy * ** show ?thesis |
60535 | 399 |
by (simp add: mkPinj_cn) |
400 |
next |
|
60567 | 401 |
case xy: 3 |
60535 | 402 |
then obtain d where x: "x = d + y" |
403 |
by atomize_elim arith |
|
404 |
from 6 have *: "y > 0" |
|
55793 | 405 |
by (cases y) (auto simp add: norm_Pinj_0_False) |
60535 | 406 |
from 6 have **: "isnorm P2" "isnorm Q2" |
55793 | 407 |
by (auto simp add: norm_Pinj[of _ P2] norm_Pinj[of _ Q2]) |
60567 | 408 |
from 6 xy x have "isnorm (Pinj d P2)" |
60535 | 409 |
apply (cases d) |
410 |
apply simp |
|
411 |
apply (cases P2) |
|
412 |
apply auto |
|
413 |
done |
|
60567 | 414 |
with 6 xy * ** x show ?thesis |
60535 | 415 |
by (simp add: mkPinj_cn) |
416 |
qed |
|
17378
105519771c67
The oracle for Presburger has been changer: It is automatically generated form a verified formaliztion of Cooper's Algorithm ex/Reflected_Presburger.thy
chaieb
parents:
diff
changeset
|
417 |
next |
105519771c67
The oracle for Presburger has been changer: It is automatically generated form a verified formaliztion of Cooper's Algorithm ex/Reflected_Presburger.thy
chaieb
parents:
diff
changeset
|
418 |
case (7 x P2 Q2 y R) |
41807 | 419 |
then have Y0: "y > 0" by (cases y) auto |
60535 | 420 |
consider "x = 0" | "x = 1" | d where "x = Suc (Suc d)" |
421 |
by atomize_elim arith |
|
422 |
then show ?case |
|
423 |
proof cases |
|
424 |
case 1 |
|
425 |
with 7 show ?thesis |
|
55793 | 426 |
by (auto simp add: norm_Pinj_0_False) |
60535 | 427 |
next |
60567 | 428 |
case x: 2 |
55793 | 429 |
from 7 have "isnorm R" "isnorm P2" |
430 |
by (auto simp add: norm_Pinj[of _ P2] norm_PX2[of Q2 y R]) |
|
64962
bf41e1109db3
Added new / improved tactics for fields and rings
berghofe
parents:
61586
diff
changeset
|
431 |
with 7 x have "isnorm (R \<langle>*\<rangle> P2)" "isnorm Q2" |
55793 | 432 |
by (auto simp add: norm_PX1[of Q2 y R]) |
60567 | 433 |
with 7 x Y0 show ?thesis |
55793 | 434 |
by (simp add: mkPX_cn) |
60535 | 435 |
next |
60567 | 436 |
case x: 3 |
60535 | 437 |
from 7 have *: "isnorm R" "isnorm Q2" |
41807 | 438 |
by (auto simp add: norm_PX2[of Q2 y R] norm_PX1[of Q2 y R]) |
60567 | 439 |
from 7 x have "isnorm (Pinj (x - 1) P2)" |
55793 | 440 |
by (cases P2) auto |
64962
bf41e1109db3
Added new / improved tactics for fields and rings
berghofe
parents:
61586
diff
changeset
|
441 |
with 7 x * have "isnorm (R \<langle>*\<rangle> Pinj (x - 1) P2)" "isnorm (Pinj x P2 \<langle>*\<rangle> Q2)" |
55793 | 442 |
by auto |
60567 | 443 |
with Y0 x show ?thesis |
55793 | 444 |
by (simp add: mkPX_cn) |
60535 | 445 |
qed |
17378
105519771c67
The oracle for Presburger has been changer: It is automatically generated form a verified formaliztion of Cooper's Algorithm ex/Reflected_Presburger.thy
chaieb
parents:
diff
changeset
|
446 |
next |
105519771c67
The oracle for Presburger has been changer: It is automatically generated form a verified formaliztion of Cooper's Algorithm ex/Reflected_Presburger.thy
chaieb
parents:
diff
changeset
|
447 |
case (8 Q2 y R x P2) |
55793 | 448 |
then have Y0: "y > 0" |
449 |
by (cases y) auto |
|
60535 | 450 |
consider "x = 0" | "x = 1" | d where "x = Suc (Suc d)" |
451 |
by atomize_elim arith |
|
452 |
then show ?case |
|
453 |
proof cases |
|
454 |
case 1 |
|
455 |
with 8 show ?thesis |
|
55793 | 456 |
by (auto simp add: norm_Pinj_0_False) |
60535 | 457 |
next |
60567 | 458 |
case x: 2 |
55793 | 459 |
from 8 have "isnorm R" "isnorm P2" |
460 |
by (auto simp add: norm_Pinj[of _ P2] norm_PX2[of Q2 y R]) |
|
64962
bf41e1109db3
Added new / improved tactics for fields and rings
berghofe
parents:
61586
diff
changeset
|
461 |
with 8 x have "isnorm (R \<langle>*\<rangle> P2)" "isnorm Q2" |
55793 | 462 |
by (auto simp add: norm_PX1[of Q2 y R]) |
60567 | 463 |
with 8 x Y0 show ?thesis |
55793 | 464 |
by (simp add: mkPX_cn) |
60535 | 465 |
next |
60567 | 466 |
case x: 3 |
60535 | 467 |
from 8 have *: "isnorm R" "isnorm Q2" |
41807 | 468 |
by (auto simp add: norm_PX2[of Q2 y R] norm_PX1[of Q2 y R]) |
60567 | 469 |
from 8 x have "isnorm (Pinj (x - 1) P2)" |
55793 | 470 |
by (cases P2) auto |
64962
bf41e1109db3
Added new / improved tactics for fields and rings
berghofe
parents:
61586
diff
changeset
|
471 |
with 8 x * have "isnorm (R \<langle>*\<rangle> Pinj (x - 1) P2)" "isnorm (Pinj x P2 \<langle>*\<rangle> Q2)" |
55793 | 472 |
by auto |
60567 | 473 |
with Y0 x show ?thesis |
60535 | 474 |
by (simp add: mkPX_cn) |
475 |
qed |
|
17378
105519771c67
The oracle for Presburger has been changer: It is automatically generated form a verified formaliztion of Cooper's Algorithm ex/Reflected_Presburger.thy
chaieb
parents:
diff
changeset
|
476 |
next |
105519771c67
The oracle for Presburger has been changer: It is automatically generated form a verified formaliztion of Cooper's Algorithm ex/Reflected_Presburger.thy
chaieb
parents:
diff
changeset
|
477 |
case (9 P1 x P2 Q1 y Q2) |
41807 | 478 |
from 9 have X0: "x > 0" by (cases x) auto |
479 |
from 9 have Y0: "y > 0" by (cases y) auto |
|
60535 | 480 |
from 9 have *: "isnorm P1" "isnorm P2" |
55793 | 481 |
by (auto simp add: norm_PX1[of P1 x P2] norm_PX2[of P1 x P2]) |
482 |
from 9 have "isnorm Q1" "isnorm Q2" |
|
483 |
by (auto simp add: norm_PX1[of Q1 y Q2] norm_PX2[of Q1 y Q2]) |
|
64962
bf41e1109db3
Added new / improved tactics for fields and rings
berghofe
parents:
61586
diff
changeset
|
484 |
with 9 * have "isnorm (P1 \<langle>*\<rangle> Q1)" "isnorm (P2 \<langle>*\<rangle> Q2)" |
bf41e1109db3
Added new / improved tactics for fields and rings
berghofe
parents:
61586
diff
changeset
|
485 |
"isnorm (P1 \<langle>*\<rangle> mkPinj 1 Q2)" "isnorm (Q1 \<langle>*\<rangle> mkPinj 1 P2)" |
17378
105519771c67
The oracle for Presburger has been changer: It is automatically generated form a verified formaliztion of Cooper's Algorithm ex/Reflected_Presburger.thy
chaieb
parents:
diff
changeset
|
486 |
by (auto simp add: mkPinj_cn) |
41807 | 487 |
with 9 X0 Y0 have |
64962
bf41e1109db3
Added new / improved tactics for fields and rings
berghofe
parents:
61586
diff
changeset
|
488 |
"isnorm (mkPX (P1 \<langle>*\<rangle> Q1) (x + y) (P2 \<langle>*\<rangle> Q2))" |
bf41e1109db3
Added new / improved tactics for fields and rings
berghofe
parents:
61586
diff
changeset
|
489 |
"isnorm (mkPX (P1 \<langle>*\<rangle> mkPinj (Suc 0) Q2) x (Pc 0))" |
bf41e1109db3
Added new / improved tactics for fields and rings
berghofe
parents:
61586
diff
changeset
|
490 |
"isnorm (mkPX (Q1 \<langle>*\<rangle> mkPinj (Suc 0) P2) y (Pc 0))" |
17378
105519771c67
The oracle for Presburger has been changer: It is automatically generated form a verified formaliztion of Cooper's Algorithm ex/Reflected_Presburger.thy
chaieb
parents:
diff
changeset
|
491 |
by (auto simp add: mkPX_cn) |
60535 | 492 |
then show ?case |
493 |
by (simp add: add_cn) |
|
494 |
qed |
|
17378
105519771c67
The oracle for Presburger has been changer: It is automatically generated form a verified formaliztion of Cooper's Algorithm ex/Reflected_Presburger.thy
chaieb
parents:
diff
changeset
|
495 |
|
60533 | 496 |
text \<open>neg conserves normalizedness\<close> |
17378
105519771c67
The oracle for Presburger has been changer: It is automatically generated form a verified formaliztion of Cooper's Algorithm ex/Reflected_Presburger.thy
chaieb
parents:
diff
changeset
|
497 |
lemma neg_cn: "isnorm P \<Longrightarrow> isnorm (neg P)" |
22742
06165e40e7bd
switched from recdef to function package; constants add, mul, pow now curried; infix syntax for algebraic operations.
haftmann
parents:
17508
diff
changeset
|
498 |
proof (induct P) |
60535 | 499 |
case Pc |
500 |
then show ?case by simp |
|
501 |
next |
|
17378
105519771c67
The oracle for Presburger has been changer: It is automatically generated form a verified formaliztion of Cooper's Algorithm ex/Reflected_Presburger.thy
chaieb
parents:
diff
changeset
|
502 |
case (Pinj i P2) |
55793 | 503 |
then have "isnorm P2" |
504 |
by (simp add: norm_Pinj[of i P2]) |
|
505 |
with Pinj show ?case |
|
506 |
by (cases P2) (auto, cases i, auto) |
|
17378
105519771c67
The oracle for Presburger has been changer: It is automatically generated form a verified formaliztion of Cooper's Algorithm ex/Reflected_Presburger.thy
chaieb
parents:
diff
changeset
|
507 |
next |
41807 | 508 |
case (PX P1 x P2) note PX1 = this |
509 |
from PX have "isnorm P2" "isnorm P1" |
|
510 |
by (auto simp add: norm_PX1[of P1 x P2] norm_PX2[of P1 x P2]) |
|
511 |
with PX show ?case |
|
512 |
proof (cases P1) |
|
17378
105519771c67
The oracle for Presburger has been changer: It is automatically generated form a verified formaliztion of Cooper's Algorithm ex/Reflected_Presburger.thy
chaieb
parents:
diff
changeset
|
513 |
case (PX p1 y p2) |
55793 | 514 |
with PX1 show ?thesis |
515 |
by (cases x) (auto, cases p2, auto) |
|
17378
105519771c67
The oracle for Presburger has been changer: It is automatically generated form a verified formaliztion of Cooper's Algorithm ex/Reflected_Presburger.thy
chaieb
parents:
diff
changeset
|
516 |
next |
105519771c67
The oracle for Presburger has been changer: It is automatically generated form a verified formaliztion of Cooper's Algorithm ex/Reflected_Presburger.thy
chaieb
parents:
diff
changeset
|
517 |
case Pinj |
55793 | 518 |
with PX1 show ?thesis |
519 |
by (cases x) auto |
|
41807 | 520 |
qed (cases x, auto) |
60535 | 521 |
qed |
17378
105519771c67
The oracle for Presburger has been changer: It is automatically generated form a verified formaliztion of Cooper's Algorithm ex/Reflected_Presburger.thy
chaieb
parents:
diff
changeset
|
522 |
|
60533 | 523 |
text \<open>sub conserves normalizedness\<close> |
64962
bf41e1109db3
Added new / improved tactics for fields and rings
berghofe
parents:
61586
diff
changeset
|
524 |
lemma sub_cn: "isnorm p \<Longrightarrow> isnorm q \<Longrightarrow> isnorm (p \<langle>-\<rangle> q)" |
44779 | 525 |
by (simp add: sub_def add_cn neg_cn) |
17378
105519771c67
The oracle for Presburger has been changer: It is automatically generated form a verified formaliztion of Cooper's Algorithm ex/Reflected_Presburger.thy
chaieb
parents:
diff
changeset
|
526 |
|
60533 | 527 |
text \<open>sqr conserves normalizizedness\<close> |
44779 | 528 |
lemma sqr_cn: "isnorm P \<Longrightarrow> isnorm (sqr P)" |
41807 | 529 |
proof (induct P) |
44779 | 530 |
case Pc |
531 |
then show ?case by simp |
|
532 |
next |
|
17378
105519771c67
The oracle for Presburger has been changer: It is automatically generated form a verified formaliztion of Cooper's Algorithm ex/Reflected_Presburger.thy
chaieb
parents:
diff
changeset
|
533 |
case (Pinj i Q) |
41807 | 534 |
then show ?case |
60535 | 535 |
apply (cases Q) |
67123 | 536 |
apply (auto simp add: mkPX_cn mkPinj_cn) |
60535 | 537 |
apply (cases i) |
67123 | 538 |
apply (auto simp add: mkPX_cn mkPinj_cn) |
60535 | 539 |
done |
55793 | 540 |
next |
17378
105519771c67
The oracle for Presburger has been changer: It is automatically generated form a verified formaliztion of Cooper's Algorithm ex/Reflected_Presburger.thy
chaieb
parents:
diff
changeset
|
541 |
case (PX P1 x P2) |
55754 | 542 |
then have "x + x \<noteq> 0" "isnorm P2" "isnorm P1" |
67123 | 543 |
by (cases x) (use PX in \<open>auto simp add: norm_PX1[of P1 x P2] norm_PX2[of P1 x P2]\<close>) |
64962
bf41e1109db3
Added new / improved tactics for fields and rings
berghofe
parents:
61586
diff
changeset
|
544 |
with PX have "isnorm (mkPX (Pc (1 + 1) \<langle>*\<rangle> P1 \<langle>*\<rangle> mkPinj (Suc 0) P2) x (Pc 0))" |
60535 | 545 |
and "isnorm (mkPX (sqr P1) (x + x) (sqr P2))" |
41807 | 546 |
by (auto simp add: add_cn mkPX_cn mkPinj_cn mul_cn) |
55793 | 547 |
then show ?case |
548 |
by (auto simp add: add_cn mkPX_cn mkPinj_cn mul_cn) |
|
44779 | 549 |
qed |
17378
105519771c67
The oracle for Presburger has been changer: It is automatically generated form a verified formaliztion of Cooper's Algorithm ex/Reflected_Presburger.thy
chaieb
parents:
diff
changeset
|
550 |
|
67123 | 551 |
text \<open>\<open>pow\<close> conserves normalizedness\<close> |
44779 | 552 |
lemma pow_cn: "isnorm P \<Longrightarrow> isnorm (pow n P)" |
553 |
proof (induct n arbitrary: P rule: less_induct) |
|
554 |
case (less k) |
|
55793 | 555 |
show ?case |
41807 | 556 |
proof (cases "k = 0") |
44779 | 557 |
case True |
558 |
then show ?thesis by simp |
|
559 |
next |
|
17378
105519771c67
The oracle for Presburger has been changer: It is automatically generated form a verified formaliztion of Cooper's Algorithm ex/Reflected_Presburger.thy
chaieb
parents:
diff
changeset
|
560 |
case False |
60535 | 561 |
then have K2: "k div 2 < k" |
562 |
by (cases k) auto |
|
563 |
from less have "isnorm (sqr P)" |
|
564 |
by (simp add: sqr_cn) |
|
44779 | 565 |
with less False K2 show ?thesis |
58712 | 566 |
by (cases "even k") (auto simp add: mul_cn elim: evenE oddE) |
44779 | 567 |
qed |
17378
105519771c67
The oracle for Presburger has been changer: It is automatically generated form a verified formaliztion of Cooper's Algorithm ex/Reflected_Presburger.thy
chaieb
parents:
diff
changeset
|
568 |
qed |
105519771c67
The oracle for Presburger has been changer: It is automatically generated form a verified formaliztion of Cooper's Algorithm ex/Reflected_Presburger.thy
chaieb
parents:
diff
changeset
|
569 |
|
17388 | 570 |
end |