author | wenzelm |
Sat, 13 Jun 2009 11:33:50 +0200 | |
changeset 31584 | 91644309417c |
parent 30732 | afca5be252d6 |
child 35762 | af3ff2ba4c54 |
permissions | -rw-r--r-- |
7072 | 1 |
(* Title: Provers/Arith/assoc_fold.ML |
2 |
ID: $Id$ |
|
3 |
Author: Lawrence C Paulson, Cambridge University Computer Laboratory |
|
4 |
Copyright 1999 University of Cambridge |
|
5 |
||
20055
24c127b97ab5
simprocs: no theory argument -- use simpset context instead;
wenzelm
parents:
17956
diff
changeset
|
6 |
Simplification procedure for associative operators + and * on numeric |
24c127b97ab5
simprocs: no theory argument -- use simpset context instead;
wenzelm
parents:
17956
diff
changeset
|
7 |
types. Performs constant folding when the literals are separated, as |
24c127b97ab5
simprocs: no theory argument -- use simpset context instead;
wenzelm
parents:
17956
diff
changeset
|
8 |
in 3+n+4. |
7072 | 9 |
*) |
10 |
||
11 |
signature ASSOC_FOLD_DATA = |
|
12 |
sig |
|
20055
24c127b97ab5
simprocs: no theory argument -- use simpset context instead;
wenzelm
parents:
17956
diff
changeset
|
13 |
val assoc_ss: simpset |
24c127b97ab5
simprocs: no theory argument -- use simpset context instead;
wenzelm
parents:
17956
diff
changeset
|
14 |
val eq_reflection: thm |
30732
afca5be252d6
parameterize assoc_fold with is_numeral predicate
huffman
parents:
25919
diff
changeset
|
15 |
val is_numeral: term -> bool |
7072 | 16 |
end; |
17 |
||
20055
24c127b97ab5
simprocs: no theory argument -- use simpset context instead;
wenzelm
parents:
17956
diff
changeset
|
18 |
signature ASSOC_FOLD = |
24c127b97ab5
simprocs: no theory argument -- use simpset context instead;
wenzelm
parents:
17956
diff
changeset
|
19 |
sig |
24c127b97ab5
simprocs: no theory argument -- use simpset context instead;
wenzelm
parents:
17956
diff
changeset
|
20 |
val proc: simpset -> term -> thm option |
24c127b97ab5
simprocs: no theory argument -- use simpset context instead;
wenzelm
parents:
17956
diff
changeset
|
21 |
end; |
7072 | 22 |
|
20055
24c127b97ab5
simprocs: no theory argument -- use simpset context instead;
wenzelm
parents:
17956
diff
changeset
|
23 |
functor Assoc_Fold(Data: ASSOC_FOLD_DATA): ASSOC_FOLD = |
7072 | 24 |
struct |
25 |
||
20055
24c127b97ab5
simprocs: no theory argument -- use simpset context instead;
wenzelm
parents:
17956
diff
changeset
|
26 |
exception Assoc_fail; |
7072 | 27 |
|
20055
24c127b97ab5
simprocs: no theory argument -- use simpset context instead;
wenzelm
parents:
17956
diff
changeset
|
28 |
fun mk_sum plus [] = raise Assoc_fail |
24c127b97ab5
simprocs: no theory argument -- use simpset context instead;
wenzelm
parents:
17956
diff
changeset
|
29 |
| mk_sum plus tms = foldr1 (fn (x, y) => plus $ x $ y) tms; |
7072 | 30 |
|
20055
24c127b97ab5
simprocs: no theory argument -- use simpset context instead;
wenzelm
parents:
17956
diff
changeset
|
31 |
(*Separate the literals from the other terms being combined*) |
24c127b97ab5
simprocs: no theory argument -- use simpset context instead;
wenzelm
parents:
17956
diff
changeset
|
32 |
fun sift_terms plus (t, (lits,others)) = |
30732
afca5be252d6
parameterize assoc_fold with is_numeral predicate
huffman
parents:
25919
diff
changeset
|
33 |
if Data.is_numeral t then (t::lits, others) (*new literal*) else |
20055
24c127b97ab5
simprocs: no theory argument -- use simpset context instead;
wenzelm
parents:
17956
diff
changeset
|
34 |
(case t of |
30732
afca5be252d6
parameterize assoc_fold with is_numeral predicate
huffman
parents:
25919
diff
changeset
|
35 |
(f as Const _) $ x $ y => |
20055
24c127b97ab5
simprocs: no theory argument -- use simpset context instead;
wenzelm
parents:
17956
diff
changeset
|
36 |
if f = plus |
24c127b97ab5
simprocs: no theory argument -- use simpset context instead;
wenzelm
parents:
17956
diff
changeset
|
37 |
then sift_terms plus (x, sift_terms plus (y, (lits,others))) |
24c127b97ab5
simprocs: no theory argument -- use simpset context instead;
wenzelm
parents:
17956
diff
changeset
|
38 |
else (lits, t::others) (*arbitrary summand*) |
24c127b97ab5
simprocs: no theory argument -- use simpset context instead;
wenzelm
parents:
17956
diff
changeset
|
39 |
| _ => (lits, t::others)); |
7072 | 40 |
|
20055
24c127b97ab5
simprocs: no theory argument -- use simpset context instead;
wenzelm
parents:
17956
diff
changeset
|
41 |
(*A simproc to combine all literals in a associative nest*) |
24c127b97ab5
simprocs: no theory argument -- use simpset context instead;
wenzelm
parents:
17956
diff
changeset
|
42 |
fun proc ss lhs = |
24c127b97ab5
simprocs: no theory argument -- use simpset context instead;
wenzelm
parents:
17956
diff
changeset
|
43 |
let |
24c127b97ab5
simprocs: no theory argument -- use simpset context instead;
wenzelm
parents:
17956
diff
changeset
|
44 |
val plus = (case lhs of f $ _ $ _ => f | _ => error "Assoc_fold: bad pattern") |
24c127b97ab5
simprocs: no theory argument -- use simpset context instead;
wenzelm
parents:
17956
diff
changeset
|
45 |
val (lits, others) = sift_terms plus (lhs, ([],[])) |
24c127b97ab5
simprocs: no theory argument -- use simpset context instead;
wenzelm
parents:
17956
diff
changeset
|
46 |
val _ = length lits < 2 andalso raise Assoc_fail (*we can't reduce the number of terms*) |
24c127b97ab5
simprocs: no theory argument -- use simpset context instead;
wenzelm
parents:
17956
diff
changeset
|
47 |
val rhs = plus $ mk_sum plus lits $ mk_sum plus others |
24c127b97ab5
simprocs: no theory argument -- use simpset context instead;
wenzelm
parents:
17956
diff
changeset
|
48 |
val th = Goal.prove (Simplifier.the_context ss) [] [] (Logic.mk_equals (lhs, rhs)) |
24c127b97ab5
simprocs: no theory argument -- use simpset context instead;
wenzelm
parents:
17956
diff
changeset
|
49 |
(fn _ => rtac Data.eq_reflection 1 THEN |
24c127b97ab5
simprocs: no theory argument -- use simpset context instead;
wenzelm
parents:
17956
diff
changeset
|
50 |
simp_tac (Simplifier.inherit_context ss Data.assoc_ss) 1) |
24c127b97ab5
simprocs: no theory argument -- use simpset context instead;
wenzelm
parents:
17956
diff
changeset
|
51 |
in SOME th end handle Assoc_fail => NONE; |
13462 | 52 |
|
7072 | 53 |
end; |