author | huffman |
Mon, 11 May 2009 12:41:46 -0700 | |
changeset 31114 | 2e9cc546e5b0 |
parent 30649 | 57753e0ec1d4 |
child 35762 | af3ff2ba4c54 |
permissions | -rw-r--r-- |
10694 | 1 |
(* Title: Provers/Arith/extract_common_term.ML |
2 |
ID: $Id$ |
|
3 |
Author: Lawrence C Paulson, Cambridge University Computer Laboratory |
|
4 |
Copyright 2000 University of Cambridge |
|
5 |
||
6 |
Extract common terms in balanced expressions: |
|
7 |
||
8 |
i + u + j ~~ i' + u + j' == u + (i + j) ~~ u + (i' + j') |
|
9 |
i + u ~~ u == u + i ~~ u + 0 |
|
10 |
||
20114
a1bb4bc68ff3
prove_conv: Variable.import_terms instead of Term.addhoc_freeze_vars;
wenzelm
parents:
20044
diff
changeset
|
11 |
where ~~ is an appropriate balancing operation (e.g. =, <=, <, -) and 0 is a |
10694 | 12 |
suitable identity for +. |
13 |
||
14 |
This massaged formula is then simplified in a user-specified way. |
|
15 |
*) |
|
16 |
||
17 |
signature EXTRACT_COMMON_TERM_DATA = |
|
18 |
sig |
|
19 |
(*abstract syntax*) |
|
14387
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
13484
diff
changeset
|
20 |
val mk_sum: typ -> term list -> term |
10694 | 21 |
val dest_sum: term -> term list |
22 |
val mk_bal: term * term -> term |
|
23 |
val dest_bal: term -> term * term |
|
24 |
val find_first: term -> term list -> term list |
|
25 |
(*proof tools*) |
|
20114
a1bb4bc68ff3
prove_conv: Variable.import_terms instead of Term.addhoc_freeze_vars;
wenzelm
parents:
20044
diff
changeset
|
26 |
val prove_conv: tactic list -> Proof.context -> thm list -> term * term -> thm option |
16973 | 27 |
val norm_tac: simpset -> tactic (*proves the result*) |
30649
57753e0ec1d4
1. New cancellation simprocs for common factors in inequations
nipkow
parents:
20114
diff
changeset
|
28 |
val simplify_meta_eq: simpset -> thm -> thm -> thm (*simplifies the result*) |
57753e0ec1d4
1. New cancellation simprocs for common factors in inequations
nipkow
parents:
20114
diff
changeset
|
29 |
val simp_conv: simpset -> term -> thm option (*proves simp thm*) |
10694 | 30 |
end; |
31 |
||
32 |
||
33 |
functor ExtractCommonTermFun(Data: EXTRACT_COMMON_TERM_DATA): |
|
34 |
sig |
|
20044
92cc2f4c7335
simprocs: no theory argument -- use simpset context instead;
wenzelm
parents:
17412
diff
changeset
|
35 |
val proc: simpset -> term -> thm option |
20114
a1bb4bc68ff3
prove_conv: Variable.import_terms instead of Term.addhoc_freeze_vars;
wenzelm
parents:
20044
diff
changeset
|
36 |
end |
10694 | 37 |
= |
38 |
struct |
|
39 |
||
40 |
(*a left-to-right scan of terms1, seeking a term u that is also in terms2*) |
|
41 |
fun find_common (terms1,terms2) = |
|
17412 | 42 |
let val tab2 = fold (Termtab.update o rpair ()) terms2 Termtab.empty |
20114
a1bb4bc68ff3
prove_conv: Variable.import_terms instead of Term.addhoc_freeze_vars;
wenzelm
parents:
20044
diff
changeset
|
43 |
fun seek [] = raise TERM("find_common", []) |
a1bb4bc68ff3
prove_conv: Variable.import_terms instead of Term.addhoc_freeze_vars;
wenzelm
parents:
20044
diff
changeset
|
44 |
| seek (u::terms) = |
a1bb4bc68ff3
prove_conv: Variable.import_terms instead of Term.addhoc_freeze_vars;
wenzelm
parents:
20044
diff
changeset
|
45 |
if Termtab.defined tab2 u then u |
a1bb4bc68ff3
prove_conv: Variable.import_terms instead of Term.addhoc_freeze_vars;
wenzelm
parents:
20044
diff
changeset
|
46 |
else seek terms |
16973 | 47 |
in seek terms1 end; |
10694 | 48 |
|
49 |
(*the simplification procedure*) |
|
20044
92cc2f4c7335
simprocs: no theory argument -- use simpset context instead;
wenzelm
parents:
17412
diff
changeset
|
50 |
fun proc ss t = |
15027 | 51 |
let |
20114
a1bb4bc68ff3
prove_conv: Variable.import_terms instead of Term.addhoc_freeze_vars;
wenzelm
parents:
20044
diff
changeset
|
52 |
val ctxt = Simplifier.the_context ss; |
a1bb4bc68ff3
prove_conv: Variable.import_terms instead of Term.addhoc_freeze_vars;
wenzelm
parents:
20044
diff
changeset
|
53 |
val prems = prems_of_ss ss; |
a1bb4bc68ff3
prove_conv: Variable.import_terms instead of Term.addhoc_freeze_vars;
wenzelm
parents:
20044
diff
changeset
|
54 |
val ([t'], ctxt') = Variable.import_terms true [t] ctxt |
a1bb4bc68ff3
prove_conv: Variable.import_terms instead of Term.addhoc_freeze_vars;
wenzelm
parents:
20044
diff
changeset
|
55 |
val export = singleton (Variable.export ctxt' ctxt) |
a1bb4bc68ff3
prove_conv: Variable.import_terms instead of Term.addhoc_freeze_vars;
wenzelm
parents:
20044
diff
changeset
|
56 |
|
a1bb4bc68ff3
prove_conv: Variable.import_terms instead of Term.addhoc_freeze_vars;
wenzelm
parents:
20044
diff
changeset
|
57 |
val (t1,t2) = Data.dest_bal t' |
a1bb4bc68ff3
prove_conv: Variable.import_terms instead of Term.addhoc_freeze_vars;
wenzelm
parents:
20044
diff
changeset
|
58 |
val terms1 = Data.dest_sum t1 |
a1bb4bc68ff3
prove_conv: Variable.import_terms instead of Term.addhoc_freeze_vars;
wenzelm
parents:
20044
diff
changeset
|
59 |
and terms2 = Data.dest_sum t2 |
a1bb4bc68ff3
prove_conv: Variable.import_terms instead of Term.addhoc_freeze_vars;
wenzelm
parents:
20044
diff
changeset
|
60 |
|
a1bb4bc68ff3
prove_conv: Variable.import_terms instead of Term.addhoc_freeze_vars;
wenzelm
parents:
20044
diff
changeset
|
61 |
val u = find_common (terms1,terms2) |
30649
57753e0ec1d4
1. New cancellation simprocs for common factors in inequations
nipkow
parents:
20114
diff
changeset
|
62 |
val simp_th = |
57753e0ec1d4
1. New cancellation simprocs for common factors in inequations
nipkow
parents:
20114
diff
changeset
|
63 |
case Data.simp_conv ss u of NONE => raise TERM("no simp", []) |
57753e0ec1d4
1. New cancellation simprocs for common factors in inequations
nipkow
parents:
20114
diff
changeset
|
64 |
| SOME th => th |
20114
a1bb4bc68ff3
prove_conv: Variable.import_terms instead of Term.addhoc_freeze_vars;
wenzelm
parents:
20044
diff
changeset
|
65 |
val terms1' = Data.find_first u terms1 |
a1bb4bc68ff3
prove_conv: Variable.import_terms instead of Term.addhoc_freeze_vars;
wenzelm
parents:
20044
diff
changeset
|
66 |
and terms2' = Data.find_first u terms2 |
a1bb4bc68ff3
prove_conv: Variable.import_terms instead of Term.addhoc_freeze_vars;
wenzelm
parents:
20044
diff
changeset
|
67 |
and T = Term.fastype_of u |
a1bb4bc68ff3
prove_conv: Variable.import_terms instead of Term.addhoc_freeze_vars;
wenzelm
parents:
20044
diff
changeset
|
68 |
|
a1bb4bc68ff3
prove_conv: Variable.import_terms instead of Term.addhoc_freeze_vars;
wenzelm
parents:
20044
diff
changeset
|
69 |
val reshape = |
a1bb4bc68ff3
prove_conv: Variable.import_terms instead of Term.addhoc_freeze_vars;
wenzelm
parents:
20044
diff
changeset
|
70 |
Data.prove_conv [Data.norm_tac ss] ctxt prems |
a1bb4bc68ff3
prove_conv: Variable.import_terms instead of Term.addhoc_freeze_vars;
wenzelm
parents:
20044
diff
changeset
|
71 |
(t', Data.mk_bal (Data.mk_sum T (u::terms1'), Data.mk_sum T (u::terms2'))) |
10694 | 72 |
in |
30649
57753e0ec1d4
1. New cancellation simprocs for common factors in inequations
nipkow
parents:
20114
diff
changeset
|
73 |
Option.map (export o Data.simplify_meta_eq ss simp_th) reshape |
10694 | 74 |
end |
20114
a1bb4bc68ff3
prove_conv: Variable.import_terms instead of Term.addhoc_freeze_vars;
wenzelm
parents:
20044
diff
changeset
|
75 |
(* FIXME avoid handling of generic exceptions *) |
15531 | 76 |
handle TERM _ => NONE |
77 |
| TYPE _ => NONE; (*Typically (if thy doesn't include Numeral) |
|
20114
a1bb4bc68ff3
prove_conv: Variable.import_terms instead of Term.addhoc_freeze_vars;
wenzelm
parents:
20044
diff
changeset
|
78 |
Undeclared type constructor "Numeral.bin"*) |
10694 | 79 |
|
80 |
end; |