author | wenzelm |
Fri, 01 Dec 2000 19:43:06 +0100 | |
changeset 10569 | e8346dad78e1 |
parent 9646 | aa3b82085e07 |
child 13484 | d8f5d3391766 |
permissions | -rw-r--r-- |
8774
ad5026ff0c16
new simproc, replacing combine_coeffs and working for nat, int, real
paulson
parents:
diff
changeset
|
1 |
(* Title: Provers/Arith/combine_numerals.ML |
ad5026ff0c16
new simproc, replacing combine_coeffs and working for nat, int, real
paulson
parents:
diff
changeset
|
2 |
ID: $Id$ |
ad5026ff0c16
new simproc, replacing combine_coeffs and working for nat, int, real
paulson
parents:
diff
changeset
|
3 |
Author: Lawrence C Paulson, Cambridge University Computer Laboratory |
ad5026ff0c16
new simproc, replacing combine_coeffs and working for nat, int, real
paulson
parents:
diff
changeset
|
4 |
Copyright 2000 University of Cambridge |
ad5026ff0c16
new simproc, replacing combine_coeffs and working for nat, int, real
paulson
parents:
diff
changeset
|
5 |
|
ad5026ff0c16
new simproc, replacing combine_coeffs and working for nat, int, real
paulson
parents:
diff
changeset
|
6 |
Combine coefficients in expressions: |
ad5026ff0c16
new simproc, replacing combine_coeffs and working for nat, int, real
paulson
parents:
diff
changeset
|
7 |
|
ad5026ff0c16
new simproc, replacing combine_coeffs and working for nat, int, real
paulson
parents:
diff
changeset
|
8 |
i + #m*u + j ... + #n*u + k == #(m+n)*u + (i + (j + k)) |
ad5026ff0c16
new simproc, replacing combine_coeffs and working for nat, int, real
paulson
parents:
diff
changeset
|
9 |
|
ad5026ff0c16
new simproc, replacing combine_coeffs and working for nat, int, real
paulson
parents:
diff
changeset
|
10 |
It works by (a) massaging the sum to bring the selected terms to the front: |
ad5026ff0c16
new simproc, replacing combine_coeffs and working for nat, int, real
paulson
parents:
diff
changeset
|
11 |
|
ad5026ff0c16
new simproc, replacing combine_coeffs and working for nat, int, real
paulson
parents:
diff
changeset
|
12 |
#m*u + (#n*u + (i + (j + k))) |
ad5026ff0c16
new simproc, replacing combine_coeffs and working for nat, int, real
paulson
parents:
diff
changeset
|
13 |
|
ad5026ff0c16
new simproc, replacing combine_coeffs and working for nat, int, real
paulson
parents:
diff
changeset
|
14 |
(b) then using left_distrib to reach |
ad5026ff0c16
new simproc, replacing combine_coeffs and working for nat, int, real
paulson
parents:
diff
changeset
|
15 |
|
ad5026ff0c16
new simproc, replacing combine_coeffs and working for nat, int, real
paulson
parents:
diff
changeset
|
16 |
#(m+n)*u + (i + (j + k)) |
ad5026ff0c16
new simproc, replacing combine_coeffs and working for nat, int, real
paulson
parents:
diff
changeset
|
17 |
*) |
ad5026ff0c16
new simproc, replacing combine_coeffs and working for nat, int, real
paulson
parents:
diff
changeset
|
18 |
|
ad5026ff0c16
new simproc, replacing combine_coeffs and working for nat, int, real
paulson
parents:
diff
changeset
|
19 |
signature COMBINE_NUMERALS_DATA = |
ad5026ff0c16
new simproc, replacing combine_coeffs and working for nat, int, real
paulson
parents:
diff
changeset
|
20 |
sig |
ad5026ff0c16
new simproc, replacing combine_coeffs and working for nat, int, real
paulson
parents:
diff
changeset
|
21 |
(*abstract syntax*) |
9571 | 22 |
val add: int * int -> int (*addition (or multiplication) *) |
8774
ad5026ff0c16
new simproc, replacing combine_coeffs and working for nat, int, real
paulson
parents:
diff
changeset
|
23 |
val mk_sum: term list -> term |
ad5026ff0c16
new simproc, replacing combine_coeffs and working for nat, int, real
paulson
parents:
diff
changeset
|
24 |
val dest_sum: term -> term list |
ad5026ff0c16
new simproc, replacing combine_coeffs and working for nat, int, real
paulson
parents:
diff
changeset
|
25 |
val mk_coeff: int * term -> term |
ad5026ff0c16
new simproc, replacing combine_coeffs and working for nat, int, real
paulson
parents:
diff
changeset
|
26 |
val dest_coeff: term -> int * term |
ad5026ff0c16
new simproc, replacing combine_coeffs and working for nat, int, real
paulson
parents:
diff
changeset
|
27 |
(*rules*) |
ad5026ff0c16
new simproc, replacing combine_coeffs and working for nat, int, real
paulson
parents:
diff
changeset
|
28 |
val left_distrib: thm |
ad5026ff0c16
new simproc, replacing combine_coeffs and working for nat, int, real
paulson
parents:
diff
changeset
|
29 |
(*proof tools*) |
ad5026ff0c16
new simproc, replacing combine_coeffs and working for nat, int, real
paulson
parents:
diff
changeset
|
30 |
val prove_conv: tactic list -> Sign.sg -> term * term -> thm option |
8799 | 31 |
val trans_tac: thm option -> tactic (*applies the initial lemma*) |
32 |
val norm_tac: tactic (*proves the initial lemma*) |
|
33 |
val numeral_simp_tac: tactic (*proves the final theorem*) |
|
34 |
val simplify_meta_eq: thm -> thm (*simplifies the final theorem*) |
|
8774
ad5026ff0c16
new simproc, replacing combine_coeffs and working for nat, int, real
paulson
parents:
diff
changeset
|
35 |
end; |
ad5026ff0c16
new simproc, replacing combine_coeffs and working for nat, int, real
paulson
parents:
diff
changeset
|
36 |
|
ad5026ff0c16
new simproc, replacing combine_coeffs and working for nat, int, real
paulson
parents:
diff
changeset
|
37 |
|
ad5026ff0c16
new simproc, replacing combine_coeffs and working for nat, int, real
paulson
parents:
diff
changeset
|
38 |
functor CombineNumeralsFun(Data: COMBINE_NUMERALS_DATA): |
ad5026ff0c16
new simproc, replacing combine_coeffs and working for nat, int, real
paulson
parents:
diff
changeset
|
39 |
sig |
ad5026ff0c16
new simproc, replacing combine_coeffs and working for nat, int, real
paulson
parents:
diff
changeset
|
40 |
val proc: Sign.sg -> thm list -> term -> thm option |
ad5026ff0c16
new simproc, replacing combine_coeffs and working for nat, int, real
paulson
parents:
diff
changeset
|
41 |
end |
ad5026ff0c16
new simproc, replacing combine_coeffs and working for nat, int, real
paulson
parents:
diff
changeset
|
42 |
= |
ad5026ff0c16
new simproc, replacing combine_coeffs and working for nat, int, real
paulson
parents:
diff
changeset
|
43 |
struct |
ad5026ff0c16
new simproc, replacing combine_coeffs and working for nat, int, real
paulson
parents:
diff
changeset
|
44 |
|
ad5026ff0c16
new simproc, replacing combine_coeffs and working for nat, int, real
paulson
parents:
diff
changeset
|
45 |
(*Remove the first occurrence of #m*u from the term list*) |
ad5026ff0c16
new simproc, replacing combine_coeffs and working for nat, int, real
paulson
parents:
diff
changeset
|
46 |
fun remove (_, _, []) = (*impossible, since #m*u was found by find_repeated*) |
ad5026ff0c16
new simproc, replacing combine_coeffs and working for nat, int, real
paulson
parents:
diff
changeset
|
47 |
raise TERM("combine_numerals: remove", []) |
ad5026ff0c16
new simproc, replacing combine_coeffs and working for nat, int, real
paulson
parents:
diff
changeset
|
48 |
| remove (m, u, t::terms) = |
9646 | 49 |
case try Data.dest_coeff t of |
50 |
Some(n,v) => if m=n andalso u aconv v then terms |
|
51 |
else t :: remove (m, u, terms) |
|
52 |
| None => t :: remove (m, u, terms); |
|
8774
ad5026ff0c16
new simproc, replacing combine_coeffs and working for nat, int, real
paulson
parents:
diff
changeset
|
53 |
|
ad5026ff0c16
new simproc, replacing combine_coeffs and working for nat, int, real
paulson
parents:
diff
changeset
|
54 |
(*a left-to-right scan of terms, seeking another term of the form #n*u, where |
ad5026ff0c16
new simproc, replacing combine_coeffs and working for nat, int, real
paulson
parents:
diff
changeset
|
55 |
#m*u is already in terms for some m*) |
ad5026ff0c16
new simproc, replacing combine_coeffs and working for nat, int, real
paulson
parents:
diff
changeset
|
56 |
fun find_repeated (tab, _, []) = raise TERM("find_repeated", []) |
ad5026ff0c16
new simproc, replacing combine_coeffs and working for nat, int, real
paulson
parents:
diff
changeset
|
57 |
| find_repeated (tab, past, t::terms) = |
9646 | 58 |
case try Data.dest_coeff t of |
59 |
Some(n,u) => |
|
60 |
(case Termtab.lookup (tab, u) of |
|
61 |
Some m => (u, m, n, rev (remove (m,u,past)) @ terms) |
|
62 |
| None => find_repeated (Termtab.update ((u,n), tab), |
|
63 |
t::past, terms)) |
|
64 |
| None => find_repeated (tab, t::past, terms); |
|
8774
ad5026ff0c16
new simproc, replacing combine_coeffs and working for nat, int, real
paulson
parents:
diff
changeset
|
65 |
|
ad5026ff0c16
new simproc, replacing combine_coeffs and working for nat, int, real
paulson
parents:
diff
changeset
|
66 |
(*the simplification procedure*) |
ad5026ff0c16
new simproc, replacing combine_coeffs and working for nat, int, real
paulson
parents:
diff
changeset
|
67 |
fun proc sg _ t = |
9191
300bd596d696
now freezes Vars in order to prevent errors in cases like these:
paulson
parents:
8816
diff
changeset
|
68 |
let (*first freeze any Vars in the term to prevent flex-flex problems*) |
300bd596d696
now freezes Vars in order to prevent errors in cases like these:
paulson
parents:
8816
diff
changeset
|
69 |
val rand_s = gensym"_" |
300bd596d696
now freezes Vars in order to prevent errors in cases like these:
paulson
parents:
8816
diff
changeset
|
70 |
fun mk_inst (var as Var((a,i),T)) = |
300bd596d696
now freezes Vars in order to prevent errors in cases like these:
paulson
parents:
8816
diff
changeset
|
71 |
(var, Free((a ^ rand_s ^ string_of_int i), T)) |
300bd596d696
now freezes Vars in order to prevent errors in cases like these:
paulson
parents:
8816
diff
changeset
|
72 |
val t' = subst_atomic (map mk_inst (term_vars t)) t |
300bd596d696
now freezes Vars in order to prevent errors in cases like these:
paulson
parents:
8816
diff
changeset
|
73 |
val (u,m,n,terms) = find_repeated (Termtab.empty, [], Data.dest_sum t') |
8774
ad5026ff0c16
new simproc, replacing combine_coeffs and working for nat, int, real
paulson
parents:
diff
changeset
|
74 |
val reshape = (*Move i*u to the front and put j*u into standard form |
ad5026ff0c16
new simproc, replacing combine_coeffs and working for nat, int, real
paulson
parents:
diff
changeset
|
75 |
i + #m + j + k == #m + i + (j + k) *) |
ad5026ff0c16
new simproc, replacing combine_coeffs and working for nat, int, real
paulson
parents:
diff
changeset
|
76 |
if m=0 orelse n=0 then (*trivial, so do nothing*) |
ad5026ff0c16
new simproc, replacing combine_coeffs and working for nat, int, real
paulson
parents:
diff
changeset
|
77 |
raise TERM("combine_numerals", []) |
ad5026ff0c16
new simproc, replacing combine_coeffs and working for nat, int, real
paulson
parents:
diff
changeset
|
78 |
else Data.prove_conv [Data.norm_tac] sg |
9191
300bd596d696
now freezes Vars in order to prevent errors in cases like these:
paulson
parents:
8816
diff
changeset
|
79 |
(t', |
8774
ad5026ff0c16
new simproc, replacing combine_coeffs and working for nat, int, real
paulson
parents:
diff
changeset
|
80 |
Data.mk_sum ([Data.mk_coeff(m,u), |
ad5026ff0c16
new simproc, replacing combine_coeffs and working for nat, int, real
paulson
parents:
diff
changeset
|
81 |
Data.mk_coeff(n,u)] @ terms)) |
ad5026ff0c16
new simproc, replacing combine_coeffs and working for nat, int, real
paulson
parents:
diff
changeset
|
82 |
in |
8799 | 83 |
apsome Data.simplify_meta_eq |
84 |
(Data.prove_conv |
|
85 |
[Data.trans_tac reshape, rtac Data.left_distrib 1, |
|
86 |
Data.numeral_simp_tac] sg |
|
9571 | 87 |
(t', Data.mk_sum (Data.mk_coeff(Data.add(m,n), u) :: terms))) |
8774
ad5026ff0c16
new simproc, replacing combine_coeffs and working for nat, int, real
paulson
parents:
diff
changeset
|
88 |
end |
ad5026ff0c16
new simproc, replacing combine_coeffs and working for nat, int, real
paulson
parents:
diff
changeset
|
89 |
handle TERM _ => None |
ad5026ff0c16
new simproc, replacing combine_coeffs and working for nat, int, real
paulson
parents:
diff
changeset
|
90 |
| TYPE _ => None; (*Typically (if thy doesn't include Numeral) |
ad5026ff0c16
new simproc, replacing combine_coeffs and working for nat, int, real
paulson
parents:
diff
changeset
|
91 |
Undeclared type constructor "Numeral.bin"*) |
ad5026ff0c16
new simproc, replacing combine_coeffs and working for nat, int, real
paulson
parents:
diff
changeset
|
92 |
|
ad5026ff0c16
new simproc, replacing combine_coeffs and working for nat, int, real
paulson
parents:
diff
changeset
|
93 |
end; |