| author | urbanc | 
| Wed, 01 Nov 2006 19:03:30 +0100 | |
| changeset 21142 | a56a839e9feb | 
| parent 20055 | 24c127b97ab5 | 
| child 22997 | d4f3b015b50b | 
| 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: 
17956diff
changeset | 6 | Simplification procedure for associative operators + and * on numeric | 
| 
24c127b97ab5
simprocs: no theory argument -- use simpset context instead;
 wenzelm parents: 
17956diff
changeset | 7 | types. Performs constant folding when the literals are separated, as | 
| 
24c127b97ab5
simprocs: no theory argument -- use simpset context instead;
 wenzelm parents: 
17956diff
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: 
17956diff
changeset | 13 | val assoc_ss: simpset | 
| 
24c127b97ab5
simprocs: no theory argument -- use simpset context instead;
 wenzelm parents: 
17956diff
changeset | 14 | val eq_reflection: thm | 
| 7072 | 15 | end; | 
| 16 | ||
| 20055 
24c127b97ab5
simprocs: no theory argument -- use simpset context instead;
 wenzelm parents: 
17956diff
changeset | 17 | signature ASSOC_FOLD = | 
| 
24c127b97ab5
simprocs: no theory argument -- use simpset context instead;
 wenzelm parents: 
17956diff
changeset | 18 | sig | 
| 
24c127b97ab5
simprocs: no theory argument -- use simpset context instead;
 wenzelm parents: 
17956diff
changeset | 19 | val proc: simpset -> term -> thm option | 
| 
24c127b97ab5
simprocs: no theory argument -- use simpset context instead;
 wenzelm parents: 
17956diff
changeset | 20 | end; | 
| 7072 | 21 | |
| 20055 
24c127b97ab5
simprocs: no theory argument -- use simpset context instead;
 wenzelm parents: 
17956diff
changeset | 22 | functor Assoc_Fold(Data: ASSOC_FOLD_DATA): ASSOC_FOLD = | 
| 7072 | 23 | struct | 
| 24 | ||
| 20055 
24c127b97ab5
simprocs: no theory argument -- use simpset context instead;
 wenzelm parents: 
17956diff
changeset | 25 | exception Assoc_fail; | 
| 7072 | 26 | |
| 20055 
24c127b97ab5
simprocs: no theory argument -- use simpset context instead;
 wenzelm parents: 
17956diff
changeset | 27 | fun mk_sum plus [] = raise Assoc_fail | 
| 
24c127b97ab5
simprocs: no theory argument -- use simpset context instead;
 wenzelm parents: 
17956diff
changeset | 28 | | mk_sum plus tms = foldr1 (fn (x, y) => plus $ x $ y) tms; | 
| 7072 | 29 | |
| 20055 
24c127b97ab5
simprocs: no theory argument -- use simpset context instead;
 wenzelm parents: 
17956diff
changeset | 30 | (*Separate the literals from the other terms being combined*) | 
| 
24c127b97ab5
simprocs: no theory argument -- use simpset context instead;
 wenzelm parents: 
17956diff
changeset | 31 | fun sift_terms plus (t, (lits,others)) = | 
| 
24c127b97ab5
simprocs: no theory argument -- use simpset context instead;
 wenzelm parents: 
17956diff
changeset | 32 | (case t of | 
| 
24c127b97ab5
simprocs: no theory argument -- use simpset context instead;
 wenzelm parents: 
17956diff
changeset | 33 |     Const("Numeral.number_of", _) $ _ =>       (* FIXME logic dependent *)
 | 
| 
24c127b97ab5
simprocs: no theory argument -- use simpset context instead;
 wenzelm parents: 
17956diff
changeset | 34 | (t::lits, others) (*new literal*) | 
| 
24c127b97ab5
simprocs: no theory argument -- use simpset context instead;
 wenzelm parents: 
17956diff
changeset | 35 | | (f as Const _) $ x $ y => | 
| 
24c127b97ab5
simprocs: no theory argument -- use simpset context instead;
 wenzelm parents: 
17956diff
changeset | 36 | if f = plus | 
| 
24c127b97ab5
simprocs: no theory argument -- use simpset context instead;
 wenzelm parents: 
17956diff
changeset | 37 | then sift_terms plus (x, sift_terms plus (y, (lits,others))) | 
| 
24c127b97ab5
simprocs: no theory argument -- use simpset context instead;
 wenzelm parents: 
17956diff
changeset | 38 | else (lits, t::others) (*arbitrary summand*) | 
| 
24c127b97ab5
simprocs: no theory argument -- use simpset context instead;
 wenzelm parents: 
17956diff
changeset | 39 | | _ => (lits, t::others)); | 
| 7072 | 40 | |
| 20055 
24c127b97ab5
simprocs: no theory argument -- use simpset context instead;
 wenzelm parents: 
17956diff
changeset | 41 | (*A simproc to combine all literals in a associative nest*) | 
| 
24c127b97ab5
simprocs: no theory argument -- use simpset context instead;
 wenzelm parents: 
17956diff
changeset | 42 | fun proc ss lhs = | 
| 
24c127b97ab5
simprocs: no theory argument -- use simpset context instead;
 wenzelm parents: 
17956diff
changeset | 43 | let | 
| 
24c127b97ab5
simprocs: no theory argument -- use simpset context instead;
 wenzelm parents: 
17956diff
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: 
17956diff
changeset | 45 | val (lits, others) = sift_terms plus (lhs, ([],[])) | 
| 
24c127b97ab5
simprocs: no theory argument -- use simpset context instead;
 wenzelm parents: 
17956diff
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: 
17956diff
changeset | 47 | val rhs = plus $ mk_sum plus lits $ mk_sum plus others | 
| 
24c127b97ab5
simprocs: no theory argument -- use simpset context instead;
 wenzelm parents: 
17956diff
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: 
17956diff
changeset | 49 | (fn _ => rtac Data.eq_reflection 1 THEN | 
| 
24c127b97ab5
simprocs: no theory argument -- use simpset context instead;
 wenzelm parents: 
17956diff
changeset | 50 | simp_tac (Simplifier.inherit_context ss Data.assoc_ss) 1) | 
| 
24c127b97ab5
simprocs: no theory argument -- use simpset context instead;
 wenzelm parents: 
17956diff
changeset | 51 | in SOME th end handle Assoc_fail => NONE; | 
| 13462 | 52 | |
| 7072 | 53 | end; |