author | nipkow |
Sun, 25 Jan 2004 00:42:22 +0100 | |
changeset 14360 | e654599b114e |
parent 14353 | 79f9fbef9106 |
child 14373 | 67a628beb981 |
permissions | -rw-r--r-- |
13957 | 1 |
(* Title: NSComplexBin.ML |
2 |
Author: Jacques D. Fleuriot |
|
3 |
Copyright: 2001 University of Edinburgh |
|
4 |
Descrition: Binary arithmetic for the nonstandard complex numbers |
|
5 |
*) |
|
6 |
||
7 |
(** hcomplex_of_complex (coercion from complex to nonstandard complex) **) |
|
8 |
||
9 |
Goal "hcomplex_of_complex (number_of w) = number_of w"; |
|
10 |
by (simp_tac (simpset() addsimps [hcomplex_number_of_def]) 1); |
|
11 |
qed "hcomplex_number_of"; |
|
12 |
Addsimps [hcomplex_number_of]; |
|
13 |
||
14 |
Goalw [hypreal_of_real_def] |
|
15 |
"hcomplex_of_hypreal (hypreal_of_real x) = \ |
|
16 |
\ hcomplex_of_complex(complex_of_real x)"; |
|
17 |
by (simp_tac (simpset() addsimps [hcomplex_of_hypreal, |
|
18 |
hcomplex_of_complex_def,complex_of_real_def]) 1); |
|
19 |
qed "hcomplex_of_hypreal_eq_hcomplex_of_complex"; |
|
20 |
||
21 |
Goalw [complex_number_of_def,hypreal_number_of_def] |
|
22 |
"hcomplex_of_complex (number_of w) = hcomplex_of_hypreal(number_of w)"; |
|
23 |
by (rtac (hcomplex_of_hypreal_eq_hcomplex_of_complex RS sym) 1); |
|
24 |
qed "hcomplex_hypreal_number_of"; |
|
25 |
||
26 |
Goalw [hcomplex_number_of_def] "Numeral0 = (0::hcomplex)"; |
|
14320 | 27 |
by(Simp_tac 1); |
13957 | 28 |
qed "hcomplex_numeral_0_eq_0"; |
29 |
||
30 |
Goalw [hcomplex_number_of_def] "Numeral1 = (1::hcomplex)"; |
|
14320 | 31 |
by(Simp_tac 1); |
13957 | 32 |
qed "hcomplex_numeral_1_eq_1"; |
33 |
||
34 |
(* |
|
35 |
Goal "z + hcnj z = \ |
|
36 |
\ hcomplex_of_hypreal (2 * hRe(z))"; |
|
37 |
by (res_inst_tac [("z","z")] eq_Abs_hcomplex 1); |
|
38 |
by (auto_tac (claset(),HOL_ss addsimps [hRe,hcnj,hcomplex_add, |
|
39 |
hypreal_mult,hcomplex_of_hypreal,complex_add_cnj])); |
|
40 |
qed "hcomplex_add_hcnj"; |
|
41 |
||
42 |
Goal "z - hcnj z = \ |
|
43 |
\ hcomplex_of_hypreal (hypreal_of_real #2 * hIm(z)) * iii"; |
|
44 |
by (res_inst_tac [("z","z")] eq_Abs_hcomplex 1); |
|
45 |
by (auto_tac (claset(),simpset() addsimps [hIm,hcnj,hcomplex_diff, |
|
46 |
hypreal_of_real_def,hypreal_mult,hcomplex_of_hypreal, |
|
47 |
complex_diff_cnj,iii_def,hcomplex_mult])); |
|
48 |
qed "hcomplex_diff_hcnj"; |
|
49 |
*) |
|
50 |
||
51 |
(** Addition **) |
|
52 |
||
53 |
Goal "(number_of v :: hcomplex) + number_of v' = number_of (bin_add v v')"; |
|
54 |
by (simp_tac |
|
55 |
(HOL_ss addsimps [hcomplex_number_of_def, |
|
56 |
hcomplex_of_complex_add RS sym, add_complex_number_of]) 1); |
|
57 |
qed "add_hcomplex_number_of"; |
|
58 |
Addsimps [add_hcomplex_number_of]; |
|
59 |
||
60 |
||
61 |
(** Subtraction **) |
|
62 |
||
63 |
Goalw [hcomplex_number_of_def] |
|
64 |
"- (number_of w :: hcomplex) = number_of (bin_minus w)"; |
|
65 |
by (simp_tac |
|
66 |
(HOL_ss addsimps [minus_complex_number_of, hcomplex_of_complex_minus RS sym]) 1); |
|
67 |
qed "minus_hcomplex_number_of"; |
|
68 |
Addsimps [minus_hcomplex_number_of]; |
|
69 |
||
70 |
Goalw [hcomplex_number_of_def, hcomplex_diff_def] |
|
71 |
"(number_of v :: hcomplex) - number_of w = \ |
|
72 |
\ number_of (bin_add v (bin_minus w))"; |
|
73 |
by (Simp_tac 1); |
|
74 |
qed "diff_hcomplex_number_of"; |
|
75 |
Addsimps [diff_hcomplex_number_of]; |
|
76 |
||
77 |
||
78 |
(** Multiplication **) |
|
79 |
||
80 |
Goal "(number_of v :: hcomplex) * number_of v' = number_of (bin_mult v v')"; |
|
81 |
by (simp_tac |
|
82 |
(HOL_ss addsimps [hcomplex_number_of_def, |
|
83 |
hcomplex_of_complex_mult RS sym, mult_complex_number_of]) 1); |
|
84 |
qed "mult_hcomplex_number_of"; |
|
85 |
Addsimps [mult_hcomplex_number_of]; |
|
86 |
||
87 |
Goal "(2::hcomplex) = 1 + 1"; |
|
88 |
by (simp_tac (simpset() addsimps [hcomplex_numeral_1_eq_1 RS sym]) 1); |
|
89 |
val lemma = result(); |
|
90 |
||
91 |
(*For specialist use: NOT as default simprules*) |
|
92 |
Goal "2 * z = (z+z::hcomplex)"; |
|
93 |
by (simp_tac (simpset() addsimps [lemma, hcomplex_add_mult_distrib]) 1); |
|
94 |
qed "hcomplex_mult_2"; |
|
95 |
||
96 |
Goal "z * 2 = (z+z::hcomplex)"; |
|
97 |
by (stac hcomplex_mult_commute 1 THEN rtac hcomplex_mult_2 1); |
|
98 |
qed "hcomplex_mult_2_right"; |
|
99 |
||
100 |
(** Equals (=) **) |
|
101 |
||
102 |
Goal "((number_of v :: hcomplex) = number_of v') = \ |
|
103 |
\ iszero (number_of (bin_add v (bin_minus v')))"; |
|
104 |
by (simp_tac |
|
105 |
(HOL_ss addsimps [hcomplex_number_of_def, |
|
106 |
hcomplex_of_complex_eq_iff, eq_complex_number_of]) 1); |
|
107 |
qed "eq_hcomplex_number_of"; |
|
108 |
Addsimps [eq_hcomplex_number_of]; |
|
109 |
||
110 |
(*** New versions of existing theorems involving 0, 1hc ***) |
|
111 |
||
112 |
Goal "- 1 = (-1::hcomplex)"; |
|
113 |
by (simp_tac (simpset() addsimps [hcomplex_numeral_1_eq_1 RS sym]) 1); |
|
114 |
qed "hcomplex_minus_1_eq_m1"; |
|
115 |
||
116 |
Goal "-1 * z = -(z::hcomplex)"; |
|
117 |
by (simp_tac (simpset() addsimps [hcomplex_minus_1_eq_m1 RS sym]) 1); |
|
118 |
qed "hcomplex_mult_minus1"; |
|
119 |
||
120 |
Goal "z * -1 = -(z::hcomplex)"; |
|
121 |
by (stac hcomplex_mult_commute 1 THEN rtac hcomplex_mult_minus1 1); |
|
122 |
qed "hcomplex_mult_minus1_right"; |
|
123 |
||
124 |
Addsimps [hcomplex_mult_minus1,hcomplex_mult_minus1_right]; |
|
125 |
||
126 |
(*Maps 0 to Numeral0 and 1 to Numeral1 and -Numeral1 to -1*) |
|
127 |
val hcomplex_numeral_ss = |
|
128 |
complex_numeral_ss addsimps [hcomplex_numeral_0_eq_0 RS sym, hcomplex_numeral_1_eq_1 RS sym, |
|
129 |
hcomplex_minus_1_eq_m1]; |
|
130 |
||
131 |
fun rename_numerals th = |
|
132 |
asm_full_simplify hcomplex_numeral_ss (Thm.transfer (the_context ()) th); |
|
133 |
||
134 |
||
135 |
(*Now insert some identities previously stated for 0 and 1hc*) |
|
136 |
||
137 |
||
138 |
Addsimps [hcomplex_numeral_0_eq_0,hcomplex_numeral_1_eq_1]; |
|
139 |
||
140 |
Goal "number_of v + (number_of w + z) = (number_of(bin_add v w) + z::hcomplex)"; |
|
141 |
by (auto_tac (claset(),simpset() addsimps [hcomplex_add_assoc RS sym])); |
|
142 |
qed "hcomplex_add_number_of_left"; |
|
143 |
||
144 |
Goal "number_of v *(number_of w * z) = (number_of(bin_mult v w) * z::hcomplex)"; |
|
145 |
by (simp_tac (simpset() addsimps [hcomplex_mult_assoc RS sym]) 1); |
|
146 |
qed "hcomplex_mult_number_of_left"; |
|
147 |
||
148 |
Goalw [hcomplex_diff_def] |
|
149 |
"number_of v + (number_of w - c) = number_of(bin_add v w) - (c::hcomplex)"; |
|
150 |
by (rtac hcomplex_add_number_of_left 1); |
|
151 |
qed "hcomplex_add_number_of_diff1"; |
|
152 |
||
153 |
Goal "number_of v + (c - number_of w) = \ |
|
154 |
\ number_of (bin_add v (bin_minus w)) + (c::hcomplex)"; |
|
14335 | 155 |
by (auto_tac (claset(),simpset() addsimps [hcomplex_diff_def]@ add_ac)); |
13957 | 156 |
qed "hcomplex_add_number_of_diff2"; |
157 |
||
158 |
Addsimps [hcomplex_add_number_of_left, hcomplex_mult_number_of_left, |
|
159 |
hcomplex_add_number_of_diff1, hcomplex_add_number_of_diff2]; |
|
160 |
||
161 |
||
162 |
(**** Simprocs for numeric literals ****) |
|
163 |
||
164 |
(** Combining of literal coefficients in sums of products **) |
|
165 |
||
166 |
Goal "(x = y) = (x-y = (0::hcomplex))"; |
|
167 |
by (simp_tac (simpset() addsimps [hcomplex_diff_eq_eq]) 1); |
|
168 |
qed "hcomplex_eq_iff_diff_eq_0"; |
|
169 |
||
170 |
(** For combine_numerals **) |
|
171 |
||
172 |
Goal "i*u + (j*u + k) = (i+j)*u + (k::hcomplex)"; |
|
173 |
by (asm_simp_tac (simpset() addsimps [hcomplex_add_mult_distrib] |
|
14335 | 174 |
@ add_ac) 1); |
13957 | 175 |
qed "left_hcomplex_add_mult_distrib"; |
176 |
||
177 |
(** For cancel_numerals **) |
|
178 |
||
179 |
Goal "((x::hcomplex) = u + v) = (x - (u + v) = 0)"; |
|
180 |
by (auto_tac (claset(),simpset() addsimps [hcomplex_diff_eq_eq])); |
|
181 |
qed "hcomplex_eq_add_diff_eq_0"; |
|
182 |
||
183 |
Goal "((x::hcomplex) = n) = (x - n = 0)"; |
|
184 |
by (auto_tac (claset(),simpset() addsimps [hcomplex_diff_eq_eq])); |
|
185 |
qed "hcomplex_eq_diff_eq_0"; |
|
186 |
||
187 |
val hcomplex_rel_iff_rel_0_rls = [hcomplex_eq_diff_eq_0,hcomplex_eq_add_diff_eq_0]; |
|
188 |
||
189 |
Goal "!!i::hcomplex. (i*u + m = j*u + n) = ((i-j)*u + m = n)"; |
|
190 |
by (auto_tac (claset(), simpset() addsimps [hcomplex_add_mult_distrib, |
|
14335 | 191 |
hcomplex_diff_def] @ add_ac)); |
13957 | 192 |
by (asm_simp_tac (simpset() addsimps [hcomplex_add_assoc RS sym]) 1); |
193 |
by (simp_tac (simpset() addsimps [hcomplex_add_assoc]) 1); |
|
194 |
qed "hcomplex_eq_add_iff1"; |
|
195 |
||
196 |
Goal "!!i::hcomplex. (i*u + m = j*u + n) = (m = (j-i)*u + n)"; |
|
197 |
by (res_inst_tac [("z","i")] eq_Abs_hcomplex 1); |
|
198 |
by (res_inst_tac [("z","j")] eq_Abs_hcomplex 1); |
|
199 |
by (res_inst_tac [("z","u")] eq_Abs_hcomplex 1); |
|
200 |
by (res_inst_tac [("z","m")] eq_Abs_hcomplex 1); |
|
201 |
by (res_inst_tac [("z","n")] eq_Abs_hcomplex 1); |
|
202 |
by (auto_tac (claset(), simpset() addsimps [hcomplex_diff,hcomplex_add, |
|
203 |
hcomplex_mult,complex_eq_add_iff2])); |
|
204 |
qed "hcomplex_eq_add_iff2"; |
|
205 |
||
206 |
||
207 |
structure HComplex_Numeral_Simprocs = |
|
208 |
struct |
|
209 |
||
210 |
(*Utilities*) |
|
211 |
||
212 |
val hcomplexT = Type("NSComplex.hcomplex",[]); |
|
213 |
||
214 |
fun mk_numeral n = HOLogic.number_of_const hcomplexT $ HOLogic.mk_bin n; |
|
215 |
||
216 |
val dest_numeral = Complex_Numeral_Simprocs.dest_numeral; |
|
217 |
||
218 |
val find_first_numeral = Complex_Numeral_Simprocs.find_first_numeral; |
|
219 |
||
220 |
val zero = mk_numeral 0; |
|
221 |
val mk_plus = HOLogic.mk_binop "op +"; |
|
222 |
||
223 |
val uminus_const = Const ("uminus", hcomplexT --> hcomplexT); |
|
224 |
||
225 |
(*Thus mk_sum[t] yields t+0; longer sums don't have a trailing zero*) |
|
226 |
fun mk_sum [] = zero |
|
227 |
| mk_sum [t,u] = mk_plus (t, u) |
|
228 |
| mk_sum (t :: ts) = mk_plus (t, mk_sum ts); |
|
229 |
||
230 |
(*this version ALWAYS includes a trailing zero*) |
|
231 |
fun long_mk_sum [] = zero |
|
232 |
| long_mk_sum (t :: ts) = mk_plus (t, mk_sum ts); |
|
233 |
||
234 |
val dest_plus = HOLogic.dest_bin "op +" hcomplexT; |
|
235 |
||
236 |
(*decompose additions AND subtractions as a sum*) |
|
237 |
fun dest_summing (pos, Const ("op +", _) $ t $ u, ts) = |
|
238 |
dest_summing (pos, t, dest_summing (pos, u, ts)) |
|
239 |
| dest_summing (pos, Const ("op -", _) $ t $ u, ts) = |
|
240 |
dest_summing (pos, t, dest_summing (not pos, u, ts)) |
|
241 |
| dest_summing (pos, t, ts) = |
|
242 |
if pos then t::ts else uminus_const$t :: ts; |
|
243 |
||
244 |
fun dest_sum t = dest_summing (true, t, []); |
|
245 |
||
246 |
val mk_diff = HOLogic.mk_binop "op -"; |
|
247 |
val dest_diff = HOLogic.dest_bin "op -" hcomplexT; |
|
248 |
||
249 |
val one = mk_numeral 1; |
|
250 |
val mk_times = HOLogic.mk_binop "op *"; |
|
251 |
||
252 |
fun mk_prod [] = one |
|
253 |
| mk_prod [t] = t |
|
254 |
| mk_prod (t :: ts) = if t = one then mk_prod ts |
|
255 |
else mk_times (t, mk_prod ts); |
|
256 |
||
257 |
val dest_times = HOLogic.dest_bin "op *" hcomplexT; |
|
258 |
||
259 |
fun dest_prod t = |
|
260 |
let val (t,u) = dest_times t |
|
261 |
in dest_prod t @ dest_prod u end |
|
262 |
handle TERM _ => [t]; |
|
263 |
||
264 |
(*DON'T do the obvious simplifications; that would create special cases*) |
|
265 |
fun mk_coeff (k, ts) = mk_times (mk_numeral k, ts); |
|
266 |
||
267 |
(*Express t as a product of (possibly) a numeral with other sorted terms*) |
|
268 |
fun dest_coeff sign (Const ("uminus", _) $ t) = dest_coeff (~sign) t |
|
269 |
| dest_coeff sign t = |
|
270 |
let val ts = sort Term.term_ord (dest_prod t) |
|
271 |
val (n, ts') = find_first_numeral [] ts |
|
272 |
handle TERM _ => (1, ts) |
|
273 |
in (sign*n, mk_prod ts') end; |
|
274 |
||
275 |
(*Find first coefficient-term THAT MATCHES u*) |
|
276 |
fun find_first_coeff past u [] = raise TERM("find_first_coeff", []) |
|
277 |
| find_first_coeff past u (t::terms) = |
|
278 |
let val (n,u') = dest_coeff 1 t |
|
279 |
in if u aconv u' then (n, rev past @ terms) |
|
280 |
else find_first_coeff (t::past) u terms |
|
281 |
end |
|
282 |
handle TERM _ => find_first_coeff (t::past) u terms; |
|
283 |
||
284 |
||
285 |
||
286 |
(*Simplify Numeral0+n, n+Numeral0, Numeral1*n, n*Numeral1*) |
|
287 |
val add_0s = map rename_numerals [hcomplex_add_zero_left, hcomplex_add_zero_right]; |
|
288 |
val mult_plus_1s = map rename_numerals [hcomplex_mult_one_left, hcomplex_mult_one_right]; |
|
289 |
val mult_minus_1s = map rename_numerals [hcomplex_mult_minus1, hcomplex_mult_minus1_right]; |
|
290 |
val mult_1s = mult_plus_1s @ mult_minus_1s; |
|
291 |
||
292 |
(*To perform binary arithmetic*) |
|
293 |
val bin_simps = |
|
294 |
[hcomplex_numeral_0_eq_0 RS sym, hcomplex_numeral_1_eq_1 RS sym, |
|
295 |
add_hcomplex_number_of, hcomplex_add_number_of_left, |
|
296 |
minus_hcomplex_number_of, diff_hcomplex_number_of, mult_hcomplex_number_of, |
|
297 |
hcomplex_mult_number_of_left] @ bin_arith_simps @ bin_rel_simps; |
|
298 |
||
14123 | 299 |
(*Binary arithmetic BUT NOT ADDITION since it may collapse adjacent terms |
300 |
during re-arrangement*) |
|
301 |
val non_add_bin_simps = |
|
302 |
bin_simps \\ [hcomplex_add_number_of_left, add_hcomplex_number_of]; |
|
303 |
||
13957 | 304 |
(*To evaluate binary negations of coefficients*) |
305 |
val hcomplex_minus_simps = NCons_simps @ |
|
306 |
[hcomplex_minus_1_eq_m1,minus_hcomplex_number_of, |
|
307 |
bin_minus_1, bin_minus_0, bin_minus_Pls, bin_minus_Min, |
|
308 |
bin_pred_1, bin_pred_0, bin_pred_Pls, bin_pred_Min]; |
|
309 |
||
310 |
||
311 |
(*To let us treat subtraction as addition*) |
|
14335 | 312 |
val diff_simps = [hcomplex_diff_def, minus_add_distrib, minus_minus]; |
13957 | 313 |
|
314 |
(*push the unary minus down: - x * y = x * - y *) |
|
315 |
val hcomplex_minus_mult_eq_1_to_2 = |
|
14335 | 316 |
[minus_mult_left RS sym, minus_mult_right] MRS trans |
13957 | 317 |
|> standard; |
318 |
||
319 |
(*to extract again any uncancelled minuses*) |
|
320 |
val hcomplex_minus_from_mult_simps = |
|
14335 | 321 |
[minus_minus, minus_mult_left RS sym, minus_mult_right RS sym]; |
13957 | 322 |
|
323 |
(*combine unary minus with numeric literals, however nested within a product*) |
|
324 |
val hcomplex_mult_minus_simps = |
|
14335 | 325 |
[hcomplex_mult_assoc, minus_mult_left, hcomplex_minus_mult_eq_1_to_2]; |
13957 | 326 |
|
327 |
(*Final simplification: cancel + and * *) |
|
328 |
val simplify_meta_eq = |
|
329 |
Int_Numeral_Simprocs.simplify_meta_eq |
|
14335 | 330 |
[add_zero_left, add_zero_right, |
14353
79f9fbef9106
Added lemmas to Ring_and_Field with slightly modified simplification rules
paulson
parents:
14335
diff
changeset
|
331 |
mult_zero_left, mult_zero_right, mult_1, mult_1_right]; |
13957 | 332 |
|
333 |
val prep_simproc = Complex_Numeral_Simprocs.prep_simproc; |
|
334 |
||
335 |
||
336 |
structure CancelNumeralsCommon = |
|
337 |
struct |
|
338 |
val mk_sum = mk_sum |
|
339 |
val dest_sum = dest_sum |
|
340 |
val mk_coeff = mk_coeff |
|
341 |
val dest_coeff = dest_coeff 1 |
|
342 |
val find_first_coeff = find_first_coeff [] |
|
343 |
val trans_tac = Real_Numeral_Simprocs.trans_tac |
|
344 |
val norm_tac = |
|
345 |
ALLGOALS (simp_tac (HOL_ss addsimps add_0s@mult_1s@diff_simps@ |
|
14335 | 346 |
hcomplex_minus_simps@add_ac)) |
14123 | 347 |
THEN ALLGOALS (simp_tac (HOL_ss addsimps non_add_bin_simps@hcomplex_mult_minus_simps)) |
13957 | 348 |
THEN ALLGOALS |
349 |
(simp_tac (HOL_ss addsimps hcomplex_minus_from_mult_simps@ |
|
14335 | 350 |
add_ac@mult_ac)) |
13957 | 351 |
val numeral_simp_tac = ALLGOALS (simp_tac (HOL_ss addsimps add_0s@bin_simps)) |
352 |
val simplify_meta_eq = simplify_meta_eq |
|
353 |
end; |
|
354 |
||
355 |
||
356 |
structure EqCancelNumerals = CancelNumeralsFun |
|
357 |
(open CancelNumeralsCommon |
|
358 |
val prove_conv = Bin_Simprocs.prove_conv |
|
359 |
val mk_bal = HOLogic.mk_eq |
|
360 |
val dest_bal = HOLogic.dest_bin "op =" hcomplexT |
|
361 |
val bal_add1 = hcomplex_eq_add_iff1 RS trans |
|
362 |
val bal_add2 = hcomplex_eq_add_iff2 RS trans |
|
363 |
); |
|
364 |
||
365 |
||
366 |
val cancel_numerals = |
|
367 |
map prep_simproc |
|
368 |
[("hcomplexeq_cancel_numerals", |
|
369 |
["(l::hcomplex) + m = n", "(l::hcomplex) = m + n", |
|
370 |
"(l::hcomplex) - m = n", "(l::hcomplex) = m - n", |
|
371 |
"(l::hcomplex) * m = n", "(l::hcomplex) = m * n"], |
|
372 |
EqCancelNumerals.proc)]; |
|
373 |
||
374 |
structure CombineNumeralsData = |
|
375 |
struct |
|
376 |
val add = op + : int*int -> int |
|
377 |
val mk_sum = long_mk_sum (*to work for e.g. #2*x + #3*x *) |
|
378 |
val dest_sum = dest_sum |
|
379 |
val mk_coeff = mk_coeff |
|
380 |
val dest_coeff = dest_coeff 1 |
|
381 |
val left_distrib = left_hcomplex_add_mult_distrib RS trans |
|
382 |
val prove_conv = Bin_Simprocs.prove_conv_nohyps |
|
383 |
val trans_tac = Real_Numeral_Simprocs.trans_tac |
|
384 |
val norm_tac = |
|
385 |
ALLGOALS (simp_tac (HOL_ss addsimps add_0s@mult_1s@diff_simps@ |
|
14335 | 386 |
hcomplex_minus_simps@add_ac)) |
14123 | 387 |
THEN ALLGOALS (simp_tac (HOL_ss addsimps non_add_bin_simps@hcomplex_mult_minus_simps)) |
13957 | 388 |
THEN ALLGOALS (simp_tac (HOL_ss addsimps hcomplex_minus_from_mult_simps@ |
14335 | 389 |
add_ac@mult_ac)) |
13957 | 390 |
val numeral_simp_tac = ALLGOALS |
391 |
(simp_tac (HOL_ss addsimps add_0s@bin_simps)) |
|
392 |
val simplify_meta_eq = simplify_meta_eq |
|
393 |
end; |
|
394 |
||
395 |
structure CombineNumerals = CombineNumeralsFun(CombineNumeralsData); |
|
396 |
||
397 |
val combine_numerals = |
|
398 |
prep_simproc ("hcomplex_combine_numerals", |
|
399 |
["(i::hcomplex) + j", "(i::hcomplex) - j"], |
|
400 |
CombineNumerals.proc); |
|
401 |
||
402 |
(** Declarations for ExtractCommonTerm **) |
|
403 |
||
404 |
(*this version ALWAYS includes a trailing one*) |
|
405 |
fun long_mk_prod [] = one |
|
406 |
| long_mk_prod (t :: ts) = mk_times (t, mk_prod ts); |
|
407 |
||
408 |
(*Find first term that matches u*) |
|
409 |
fun find_first past u [] = raise TERM("find_first", []) |
|
410 |
| find_first past u (t::terms) = |
|
411 |
if u aconv t then (rev past @ terms) |
|
412 |
else find_first (t::past) u terms |
|
413 |
handle TERM _ => find_first (t::past) u terms; |
|
414 |
||
415 |
(*Final simplification: cancel + and * *) |
|
416 |
fun cancel_simplify_meta_eq cancel_th th = |
|
417 |
Int_Numeral_Simprocs.simplify_meta_eq |
|
418 |
[hcomplex_mult_one_left, hcomplex_mult_one_right] |
|
419 |
(([th, cancel_th]) MRS trans); |
|
420 |
||
421 |
(*** Making constant folding work for 0 and 1 too ***) |
|
422 |
||
423 |
structure HComplexAbstractNumeralsData = |
|
424 |
struct |
|
425 |
val dest_eq = HOLogic.dest_eq o HOLogic.dest_Trueprop o concl_of |
|
426 |
val is_numeral = Bin_Simprocs.is_numeral |
|
427 |
val numeral_0_eq_0 = hcomplex_numeral_0_eq_0 |
|
428 |
val numeral_1_eq_1 = hcomplex_numeral_1_eq_1 |
|
429 |
val prove_conv = Bin_Simprocs.prove_conv_nohyps_novars |
|
430 |
fun norm_tac simps = ALLGOALS (simp_tac (HOL_ss addsimps simps)) |
|
431 |
val simplify_meta_eq = Bin_Simprocs.simplify_meta_eq |
|
432 |
end |
|
433 |
||
434 |
structure HComplexAbstractNumerals = AbstractNumeralsFun (HComplexAbstractNumeralsData) |
|
435 |
||
436 |
(*For addition, we already have rules for the operand 0. |
|
437 |
Multiplication is omitted because there are already special rules for |
|
438 |
both 0 and 1 as operands. Unary minus is trivial, just have - 1 = -1. |
|
439 |
For the others, having three patterns is a compromise between just having |
|
440 |
one (many spurious calls) and having nine (just too many!) *) |
|
441 |
val eval_numerals = |
|
442 |
map prep_simproc |
|
443 |
[("hcomplex_add_eval_numerals", |
|
444 |
["(m::hcomplex) + 1", "(m::hcomplex) + number_of v"], |
|
445 |
HComplexAbstractNumerals.proc add_hcomplex_number_of), |
|
446 |
("hcomplex_diff_eval_numerals", |
|
447 |
["(m::hcomplex) - 1", "(m::hcomplex) - number_of v"], |
|
448 |
HComplexAbstractNumerals.proc diff_hcomplex_number_of), |
|
449 |
("hcomplex_eq_eval_numerals", |
|
450 |
["(m::hcomplex) = 0", "(m::hcomplex) = 1", "(m::hcomplex) = number_of v"], |
|
451 |
HComplexAbstractNumerals.proc eq_hcomplex_number_of)] |
|
452 |
||
453 |
end; |
|
454 |
||
455 |
Addsimprocs HComplex_Numeral_Simprocs.eval_numerals; |
|
456 |
Addsimprocs HComplex_Numeral_Simprocs.cancel_numerals; |
|
457 |
Addsimprocs [HComplex_Numeral_Simprocs.combine_numerals]; |
|
458 |
||
459 |
||
460 |
(*examples: |
|
461 |
print_depth 22; |
|
462 |
set timing; |
|
463 |
set trace_simp; |
|
464 |
fun test s = (Goal s, by (Simp_tac 1)); |
|
465 |
||
466 |
test "l + 2 + 2 + 2 + (l + 2) + (oo + 2) = (uu::hcomplex)"; |
|
467 |
test " 2*u = (u::hcomplex)"; |
|
468 |
test "(i + j + 12 + (k::hcomplex)) - 15 = y"; |
|
469 |
test "(i + j + 12 + (k::hcomplex)) - 5 = y"; |
|
470 |
||
471 |
test "( 2*x - (u*v) + y) - v* 3*u = (w::hcomplex)"; |
|
472 |
test "( 2*x*u*v + (u*v)* 4 + y) - v*u* 4 = (w::hcomplex)"; |
|
473 |
test "( 2*x*u*v + (u*v)* 4 + y) - v*u = (w::hcomplex)"; |
|
474 |
test "u*v - (x*u*v + (u*v)* 4 + y) = (w::hcomplex)"; |
|
475 |
||
476 |
test "(i + j + 12 + (k::hcomplex)) = u + 15 + y"; |
|
477 |
test "(i + j* 2 + 12 + (k::hcomplex)) = j + 5 + y"; |
|
478 |
||
479 |
test " 2*y + 3*z + 6*w + 2*y + 3*z + 2*u = 2*y' + 3*z' + 6*w' + 2*y' + 3*z' + u + (vv::hcomplex)"; |
|
480 |
||
481 |
test "a + -(b+c) + b = (d::hcomplex)"; |
|
482 |
test "a + -(b+c) - b = (d::hcomplex)"; |
|
483 |
||
484 |
(*negative numerals*) |
|
485 |
test "(i + j + -2 + (k::hcomplex)) - (u + 5 + y) = zz"; |
|
486 |
||
487 |
test "(i + j + -12 + (k::hcomplex)) - 15 = y"; |
|
488 |
test "(i + j + 12 + (k::hcomplex)) - -15 = y"; |
|
489 |
test "(i + j + -12 + (k::hcomplex)) - -15 = y"; |
|
490 |
*) |
|
491 |
||
492 |
(** Constant folding for hcomplex plus and times **) |
|
493 |
||
494 |
structure HComplex_Times_Assoc_Data : ASSOC_FOLD_DATA = |
|
495 |
struct |
|
496 |
val ss = HOL_ss |
|
497 |
val eq_reflection = eq_reflection |
|
498 |
val sg_ref = Sign.self_ref (Theory.sign_of (the_context ())) |
|
499 |
val T = HComplex_Numeral_Simprocs.hcomplexT |
|
500 |
val plus = Const ("op *", [T,T] ---> T) |
|
14335 | 501 |
val add_ac = mult_ac |
13957 | 502 |
end; |
503 |
||
504 |
structure HComplex_Times_Assoc = Assoc_Fold (HComplex_Times_Assoc_Data); |
|
505 |
||
506 |
Addsimprocs [HComplex_Times_Assoc.conv]; |
|
507 |
||
508 |
Addsimps [hcomplex_of_complex_zero_iff]; |
|
509 |
||
510 |
(*Simplification of x-y = 0 *) |
|
511 |
||
512 |
AddIffs [hcomplex_eq_iff_diff_eq_0 RS sym]; |
|
513 |
||
514 |
(** extra thms **) |
|
515 |
||
516 |
Goal "(hcnj z = 0) = (z = 0)"; |
|
517 |
by (auto_tac (claset(),simpset() addsimps [hcomplex_hcnj_zero_iff])); |
|
518 |
qed "hcomplex_hcnj_num_zero_iff"; |
|
519 |
Addsimps [hcomplex_hcnj_num_zero_iff]; |
|
520 |
||
521 |
Goal "0 = Abs_hcomplex (hcomplexrel `` {%n. 0})"; |
|
522 |
by (simp_tac (simpset() addsimps [hcomplex_zero_def RS meta_eq_to_obj_eq RS sym]) 1); |
|
523 |
qed "hcomplex_zero_num"; |
|
524 |
||
525 |
Goal "1 = Abs_hcomplex (hcomplexrel `` {%n. 1})"; |
|
526 |
by (simp_tac (simpset() addsimps [hcomplex_one_def RS meta_eq_to_obj_eq RS sym]) 1); |
|
527 |
qed "hcomplex_one_num"; |
|
528 |
||
529 |
(*** Real and imaginary stuff ***) |
|
530 |
||
531 |
Goalw [hcomplex_number_of_def] |
|
532 |
"((number_of xa :: hcomplex) + iii * number_of ya = \ |
|
533 |
\ number_of xb + iii * number_of yb) = \ |
|
534 |
\ (((number_of xa :: hcomplex) = number_of xb) & \ |
|
535 |
\ ((number_of ya :: hcomplex) = number_of yb))"; |
|
536 |
by (auto_tac (claset(), HOL_ss addsimps [hcomplex_eq_cancel_iff, |
|
537 |
hcomplex_hypreal_number_of])); |
|
538 |
qed "hcomplex_number_of_eq_cancel_iff"; |
|
539 |
Addsimps [hcomplex_number_of_eq_cancel_iff]; |
|
540 |
||
541 |
Goalw [hcomplex_number_of_def] |
|
542 |
"((number_of xa :: hcomplex) + number_of ya * iii = \ |
|
543 |
\ number_of xb + number_of yb * iii) = \ |
|
544 |
\ (((number_of xa :: hcomplex) = number_of xb) & \ |
|
545 |
\ ((number_of ya :: hcomplex) = number_of yb))"; |
|
546 |
by (auto_tac (claset(), HOL_ss addsimps [hcomplex_eq_cancel_iffA, |
|
547 |
hcomplex_hypreal_number_of])); |
|
548 |
qed "hcomplex_number_of_eq_cancel_iffA"; |
|
549 |
Addsimps [hcomplex_number_of_eq_cancel_iffA]; |
|
550 |
||
551 |
Goalw [hcomplex_number_of_def] |
|
552 |
"((number_of xa :: hcomplex) + number_of ya * iii = \ |
|
553 |
\ number_of xb + iii * number_of yb) = \ |
|
554 |
\ (((number_of xa :: hcomplex) = number_of xb) & \ |
|
555 |
\ ((number_of ya :: hcomplex) = number_of yb))"; |
|
556 |
by (auto_tac (claset(), HOL_ss addsimps [hcomplex_eq_cancel_iffB, |
|
557 |
hcomplex_hypreal_number_of])); |
|
558 |
qed "hcomplex_number_of_eq_cancel_iffB"; |
|
559 |
Addsimps [hcomplex_number_of_eq_cancel_iffB]; |
|
560 |
||
561 |
Goalw [hcomplex_number_of_def] |
|
562 |
"((number_of xa :: hcomplex) + iii * number_of ya = \ |
|
563 |
\ number_of xb + number_of yb * iii) = \ |
|
564 |
\ (((number_of xa :: hcomplex) = number_of xb) & \ |
|
565 |
\ ((number_of ya :: hcomplex) = number_of yb))"; |
|
566 |
by (auto_tac (claset(), HOL_ss addsimps [hcomplex_eq_cancel_iffC, |
|
567 |
hcomplex_hypreal_number_of])); |
|
568 |
qed "hcomplex_number_of_eq_cancel_iffC"; |
|
569 |
Addsimps [hcomplex_number_of_eq_cancel_iffC]; |
|
570 |
||
571 |
Goalw [hcomplex_number_of_def] |
|
572 |
"((number_of xa :: hcomplex) + iii * number_of ya = \ |
|
573 |
\ number_of xb) = \ |
|
574 |
\ (((number_of xa :: hcomplex) = number_of xb) & \ |
|
575 |
\ ((number_of ya :: hcomplex) = 0))"; |
|
576 |
by (auto_tac (claset(), HOL_ss addsimps [hcomplex_eq_cancel_iff2, |
|
577 |
hcomplex_hypreal_number_of,hcomplex_of_hypreal_zero_iff])); |
|
578 |
qed "hcomplex_number_of_eq_cancel_iff2"; |
|
579 |
Addsimps [hcomplex_number_of_eq_cancel_iff2]; |
|
580 |
||
581 |
Goalw [hcomplex_number_of_def] |
|
582 |
"((number_of xa :: hcomplex) + number_of ya * iii = \ |
|
583 |
\ number_of xb) = \ |
|
584 |
\ (((number_of xa :: hcomplex) = number_of xb) & \ |
|
585 |
\ ((number_of ya :: hcomplex) = 0))"; |
|
586 |
by (auto_tac (claset(), HOL_ss addsimps [hcomplex_eq_cancel_iff2a, |
|
587 |
hcomplex_hypreal_number_of,hcomplex_of_hypreal_zero_iff])); |
|
588 |
qed "hcomplex_number_of_eq_cancel_iff2a"; |
|
589 |
Addsimps [hcomplex_number_of_eq_cancel_iff2a]; |
|
590 |
||
591 |
Goalw [hcomplex_number_of_def] |
|
592 |
"((number_of xa :: hcomplex) + iii * number_of ya = \ |
|
593 |
\ iii * number_of yb) = \ |
|
594 |
\ (((number_of xa :: hcomplex) = 0) & \ |
|
595 |
\ ((number_of ya :: hcomplex) = number_of yb))"; |
|
596 |
by (auto_tac (claset(), HOL_ss addsimps [hcomplex_eq_cancel_iff3, |
|
597 |
hcomplex_hypreal_number_of,hcomplex_of_hypreal_zero_iff])); |
|
598 |
qed "hcomplex_number_of_eq_cancel_iff3"; |
|
599 |
Addsimps [hcomplex_number_of_eq_cancel_iff3]; |
|
600 |
||
601 |
Goalw [hcomplex_number_of_def] |
|
602 |
"((number_of xa :: hcomplex) + number_of ya * iii= \ |
|
603 |
\ iii * number_of yb) = \ |
|
604 |
\ (((number_of xa :: hcomplex) = 0) & \ |
|
605 |
\ ((number_of ya :: hcomplex) = number_of yb))"; |
|
606 |
by (auto_tac (claset(), HOL_ss addsimps [hcomplex_eq_cancel_iff3a, |
|
607 |
hcomplex_hypreal_number_of,hcomplex_of_hypreal_zero_iff])); |
|
608 |
qed "hcomplex_number_of_eq_cancel_iff3a"; |
|
609 |
Addsimps [hcomplex_number_of_eq_cancel_iff3a]; |
|
610 |
||
611 |
Goalw [hcomplex_number_of_def] "hcnj (number_of v :: hcomplex) = number_of v"; |
|
612 |
by (rtac (hcomplex_hypreal_number_of RS ssubst) 1); |
|
613 |
by (rtac hcomplex_hcnj_hcomplex_of_hypreal 1); |
|
614 |
qed "hcomplex_number_of_hcnj"; |
|
615 |
Addsimps [hcomplex_number_of_hcnj]; |
|
616 |
||
617 |
Goalw [hcomplex_number_of_def] |
|
618 |
"hcmod(number_of v :: hcomplex) = abs (number_of v :: hypreal)"; |
|
619 |
by (rtac (hcomplex_hypreal_number_of RS ssubst) 1); |
|
620 |
by (auto_tac (claset(), HOL_ss addsimps [hcmod_hcomplex_of_hypreal])); |
|
621 |
qed "hcomplex_number_of_hcmod"; |
|
622 |
Addsimps [hcomplex_number_of_hcmod]; |
|
623 |
||
624 |
Goalw [hcomplex_number_of_def] |
|
625 |
"hRe(number_of v :: hcomplex) = number_of v"; |
|
626 |
by (rtac (hcomplex_hypreal_number_of RS ssubst) 1); |
|
627 |
by (auto_tac (claset(), HOL_ss addsimps [hRe_hcomplex_of_hypreal])); |
|
628 |
qed "hcomplex_number_of_hRe"; |
|
629 |
Addsimps [hcomplex_number_of_hRe]; |
|
630 |
||
631 |
Goalw [hcomplex_number_of_def] |
|
632 |
"hIm(number_of v :: hcomplex) = 0"; |
|
633 |
by (rtac (hcomplex_hypreal_number_of RS ssubst) 1); |
|
634 |
by (auto_tac (claset(), HOL_ss addsimps [hIm_hcomplex_of_hypreal])); |
|
635 |
qed "hcomplex_number_of_hIm"; |
|
636 |
Addsimps [hcomplex_number_of_hIm]; |
|
637 |
||
638 |
||
639 |