| author | eberlm <eberlm@in.tum.de> | 
| Thu, 06 Apr 2017 10:22:03 +0200 | |
| changeset 65400 | feb83174a87a | 
| parent 60754 | 02924903a6fd | 
| permissions | -rw-r--r-- | 
| 7072 | 1  | 
(* Title: Provers/Arith/assoc_fold.ML  | 
2  | 
Author: Lawrence C Paulson, Cambridge University Computer Laboratory  | 
|
3  | 
Copyright 1999 University of Cambridge  | 
|
4  | 
||
| 
20055
 
24c127b97ab5
simprocs: no theory argument -- use simpset context instead;
 
wenzelm 
parents: 
17956 
diff
changeset
 | 
5  | 
Simplification procedure for associative operators + and * on numeric  | 
| 
 
24c127b97ab5
simprocs: no theory argument -- use simpset context instead;
 
wenzelm 
parents: 
17956 
diff
changeset
 | 
6  | 
types. Performs constant folding when the literals are separated, as  | 
| 
 
24c127b97ab5
simprocs: no theory argument -- use simpset context instead;
 
wenzelm 
parents: 
17956 
diff
changeset
 | 
7  | 
in 3+n+4.  | 
| 7072 | 8  | 
*)  | 
9  | 
||
10  | 
signature ASSOC_FOLD_DATA =  | 
|
11  | 
sig  | 
|
| 
20055
 
24c127b97ab5
simprocs: no theory argument -- use simpset context instead;
 
wenzelm 
parents: 
17956 
diff
changeset
 | 
12  | 
val assoc_ss: simpset  | 
| 
 
24c127b97ab5
simprocs: no theory argument -- use simpset context instead;
 
wenzelm 
parents: 
17956 
diff
changeset
 | 
13  | 
val eq_reflection: thm  | 
| 
30732
 
afca5be252d6
parameterize assoc_fold with is_numeral predicate
 
huffman 
parents: 
25919 
diff
changeset
 | 
14  | 
val is_numeral: term -> bool  | 
| 7072 | 15  | 
end;  | 
16  | 
||
| 
20055
 
24c127b97ab5
simprocs: no theory argument -- use simpset context instead;
 
wenzelm 
parents: 
17956 
diff
changeset
 | 
17  | 
signature ASSOC_FOLD =  | 
| 
 
24c127b97ab5
simprocs: no theory argument -- use simpset context instead;
 
wenzelm 
parents: 
17956 
diff
changeset
 | 
18  | 
sig  | 
| 
51717
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
35762 
diff
changeset
 | 
19  | 
val proc: Proof.context -> term -> thm option  | 
| 
20055
 
24c127b97ab5
simprocs: no theory argument -- use simpset context instead;
 
wenzelm 
parents: 
17956 
diff
changeset
 | 
20  | 
end;  | 
| 7072 | 21  | 
|
| 
20055
 
24c127b97ab5
simprocs: no theory argument -- use simpset context instead;
 
wenzelm 
parents: 
17956 
diff
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: 
17956 
diff
changeset
 | 
25  | 
exception Assoc_fail;  | 
| 7072 | 26  | 
|
| 
20055
 
24c127b97ab5
simprocs: no theory argument -- use simpset context instead;
 
wenzelm 
parents: 
17956 
diff
changeset
 | 
27  | 
fun mk_sum plus [] = raise Assoc_fail  | 
| 
 
24c127b97ab5
simprocs: no theory argument -- use simpset context instead;
 
wenzelm 
parents: 
17956 
diff
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: 
17956 
diff
changeset
 | 
30  | 
(*Separate the literals from the other terms being combined*)  | 
| 
 
24c127b97ab5
simprocs: no theory argument -- use simpset context instead;
 
wenzelm 
parents: 
17956 
diff
changeset
 | 
31  | 
fun sift_terms plus (t, (lits,others)) =  | 
| 
30732
 
afca5be252d6
parameterize assoc_fold with is_numeral predicate
 
huffman 
parents: 
25919 
diff
changeset
 | 
32  | 
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
 | 
33  | 
(case t of  | 
| 
30732
 
afca5be252d6
parameterize assoc_fold with is_numeral predicate
 
huffman 
parents: 
25919 
diff
changeset
 | 
34  | 
(f as Const _) $ x $ y =>  | 
| 
20055
 
24c127b97ab5
simprocs: no theory argument -- use simpset context instead;
 
wenzelm 
parents: 
17956 
diff
changeset
 | 
35  | 
if f = plus  | 
| 
 
24c127b97ab5
simprocs: no theory argument -- use simpset context instead;
 
wenzelm 
parents: 
17956 
diff
changeset
 | 
36  | 
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
 | 
37  | 
else (lits, t::others) (*arbitrary summand*)  | 
| 
 
24c127b97ab5
simprocs: no theory argument -- use simpset context instead;
 
wenzelm 
parents: 
17956 
diff
changeset
 | 
38  | 
| _ => (lits, t::others));  | 
| 7072 | 39  | 
|
| 
20055
 
24c127b97ab5
simprocs: no theory argument -- use simpset context instead;
 
wenzelm 
parents: 
17956 
diff
changeset
 | 
40  | 
(*A simproc to combine all literals in a associative nest*)  | 
| 
51717
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
35762 
diff
changeset
 | 
41  | 
fun proc ctxt lhs =  | 
| 
20055
 
24c127b97ab5
simprocs: no theory argument -- use simpset context instead;
 
wenzelm 
parents: 
17956 
diff
changeset
 | 
42  | 
let  | 
| 
 
24c127b97ab5
simprocs: no theory argument -- use simpset context instead;
 
wenzelm 
parents: 
17956 
diff
changeset
 | 
43  | 
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
 | 
44  | 
val (lits, others) = sift_terms plus (lhs, ([],[]))  | 
| 
 
24c127b97ab5
simprocs: no theory argument -- use simpset context instead;
 
wenzelm 
parents: 
17956 
diff
changeset
 | 
45  | 
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
 | 
46  | 
val rhs = plus $ mk_sum plus lits $ mk_sum plus others  | 
| 60754 | 47  | 
val th =  | 
48  | 
Goal.prove ctxt [] [] (Logic.mk_equals (lhs, rhs)) (fn _ =>  | 
|
49  | 
resolve_tac ctxt [Data.eq_reflection] 1 THEN  | 
|
50  | 
simp_tac (put_simpset Data.assoc_ss ctxt) 1)  | 
|
| 
20055
 
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;  |