| author | wenzelm |
| Mon, 02 Sep 2024 20:54:00 +0200 | |
| changeset 80803 | 7e39c785ca5d |
| parent 78800 | 0b3700d31758 |
| 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 |
Author: Lawrence C Paulson, Cambridge University Computer Laboratory |
|
ad5026ff0c16
new simproc, replacing combine_coeffs and working for nat, int, real
paulson
parents:
diff
changeset
|
3 |
Copyright 2000 University of Cambridge |
|
ad5026ff0c16
new simproc, replacing combine_coeffs and working for nat, int, real
paulson
parents:
diff
changeset
|
4 |
|
|
ad5026ff0c16
new simproc, replacing combine_coeffs and working for nat, int, real
paulson
parents:
diff
changeset
|
5 |
Combine coefficients in expressions: |
|
ad5026ff0c16
new simproc, replacing combine_coeffs and working for nat, int, real
paulson
parents:
diff
changeset
|
6 |
|
|
ad5026ff0c16
new simproc, replacing combine_coeffs and working for nat, int, real
paulson
parents:
diff
changeset
|
7 |
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
|
8 |
|
|
ad5026ff0c16
new simproc, replacing combine_coeffs and working for nat, int, real
paulson
parents:
diff
changeset
|
9 |
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
|
10 |
|
|
ad5026ff0c16
new simproc, replacing combine_coeffs and working for nat, int, real
paulson
parents:
diff
changeset
|
11 |
#m*u + (#n*u + (i + (j + k))) |
|
ad5026ff0c16
new simproc, replacing combine_coeffs and working for nat, int, real
paulson
parents:
diff
changeset
|
12 |
|
|
ad5026ff0c16
new simproc, replacing combine_coeffs and working for nat, int, real
paulson
parents:
diff
changeset
|
13 |
(b) then using left_distrib to reach |
|
ad5026ff0c16
new simproc, replacing combine_coeffs and working for nat, int, real
paulson
parents:
diff
changeset
|
14 |
|
|
ad5026ff0c16
new simproc, replacing combine_coeffs and working for nat, int, real
paulson
parents:
diff
changeset
|
15 |
#(m+n)*u + (i + (j + k)) |
|
ad5026ff0c16
new simproc, replacing combine_coeffs and working for nat, int, real
paulson
parents:
diff
changeset
|
16 |
*) |
|
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 |
signature COMBINE_NUMERALS_DATA = |
|
ad5026ff0c16
new simproc, replacing combine_coeffs and working for nat, int, real
paulson
parents:
diff
changeset
|
19 |
sig |
|
ad5026ff0c16
new simproc, replacing combine_coeffs and working for nat, int, real
paulson
parents:
diff
changeset
|
20 |
(*abstract syntax*) |
|
23058
c722004c5a22
generalize CombineNumerals functor to allow coefficients with types other than IntInf.int
huffman
parents:
20114
diff
changeset
|
21 |
eqtype coeff |
|
c722004c5a22
generalize CombineNumerals functor to allow coefficients with types other than IntInf.int
huffman
parents:
20114
diff
changeset
|
22 |
val iszero: coeff -> bool |
|
c722004c5a22
generalize CombineNumerals functor to allow coefficients with types other than IntInf.int
huffman
parents:
20114
diff
changeset
|
23 |
val add: coeff * coeff -> coeff (*addition (or multiplication) *) |
|
14387
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
13484
diff
changeset
|
24 |
val mk_sum: typ -> term list -> term |
|
8774
ad5026ff0c16
new simproc, replacing combine_coeffs and working for nat, int, real
paulson
parents:
diff
changeset
|
25 |
val dest_sum: term -> term list |
|
23058
c722004c5a22
generalize CombineNumerals functor to allow coefficients with types other than IntInf.int
huffman
parents:
20114
diff
changeset
|
26 |
val mk_coeff: coeff * term -> term |
|
c722004c5a22
generalize CombineNumerals functor to allow coefficients with types other than IntInf.int
huffman
parents:
20114
diff
changeset
|
27 |
val dest_coeff: term -> coeff * term |
|
8774
ad5026ff0c16
new simproc, replacing combine_coeffs and working for nat, int, real
paulson
parents:
diff
changeset
|
28 |
(*rules*) |
|
ad5026ff0c16
new simproc, replacing combine_coeffs and working for nat, int, real
paulson
parents:
diff
changeset
|
29 |
val left_distrib: thm |
|
ad5026ff0c16
new simproc, replacing combine_coeffs and working for nat, int, real
paulson
parents:
diff
changeset
|
30 |
(*proof tools*) |
|
20114
a1bb4bc68ff3
prove_conv: Variable.import_terms instead of Term.addhoc_freeze_vars;
wenzelm
parents:
20044
diff
changeset
|
31 |
val prove_conv: tactic list -> Proof.context -> term * term -> thm option |
| 59530 | 32 |
val trans_tac: Proof.context -> thm option -> tactic (*applies the initial lemma*) |
|
51717
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
44947
diff
changeset
|
33 |
val norm_tac: Proof.context -> tactic (*proves the initial lemma*) |
|
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
44947
diff
changeset
|
34 |
val numeral_simp_tac: Proof.context -> tactic (*proves the final theorem*) |
|
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
44947
diff
changeset
|
35 |
val simplify_meta_eq: Proof.context -> thm -> thm (*simplifies the final theorem*) |
|
8774
ad5026ff0c16
new simproc, replacing combine_coeffs and working for nat, int, real
paulson
parents:
diff
changeset
|
36 |
end; |
|
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 |
|
|
ad5026ff0c16
new simproc, replacing combine_coeffs and working for nat, int, real
paulson
parents:
diff
changeset
|
39 |
functor CombineNumeralsFun(Data: COMBINE_NUMERALS_DATA): |
|
ad5026ff0c16
new simproc, replacing combine_coeffs and working for nat, int, real
paulson
parents:
diff
changeset
|
40 |
sig |
| 78800 | 41 |
val proc: Simplifier.proc |
|
20114
a1bb4bc68ff3
prove_conv: Variable.import_terms instead of Term.addhoc_freeze_vars;
wenzelm
parents:
20044
diff
changeset
|
42 |
end |
|
8774
ad5026ff0c16
new simproc, replacing combine_coeffs and working for nat, int, real
paulson
parents:
diff
changeset
|
43 |
= |
|
ad5026ff0c16
new simproc, replacing combine_coeffs and working for nat, int, real
paulson
parents:
diff
changeset
|
44 |
struct |
|
ad5026ff0c16
new simproc, replacing combine_coeffs and working for nat, int, real
paulson
parents:
diff
changeset
|
45 |
|
|
ad5026ff0c16
new simproc, replacing combine_coeffs and working for nat, int, real
paulson
parents:
diff
changeset
|
46 |
(*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
|
47 |
fun remove (_, _, []) = (*impossible, since #m*u was found by find_repeated*) |
|
20114
a1bb4bc68ff3
prove_conv: Variable.import_terms instead of Term.addhoc_freeze_vars;
wenzelm
parents:
20044
diff
changeset
|
48 |
raise TERM("combine_numerals: remove", [])
|
|
8774
ad5026ff0c16
new simproc, replacing combine_coeffs and working for nat, int, real
paulson
parents:
diff
changeset
|
49 |
| remove (m, u, t::terms) = |
| 9646 | 50 |
case try Data.dest_coeff t of |
|
20114
a1bb4bc68ff3
prove_conv: Variable.import_terms instead of Term.addhoc_freeze_vars;
wenzelm
parents:
20044
diff
changeset
|
51 |
SOME(n,v) => if m=n andalso u aconv v then terms |
|
a1bb4bc68ff3
prove_conv: Variable.import_terms instead of Term.addhoc_freeze_vars;
wenzelm
parents:
20044
diff
changeset
|
52 |
else t :: remove (m, u, terms) |
|
a1bb4bc68ff3
prove_conv: Variable.import_terms instead of Term.addhoc_freeze_vars;
wenzelm
parents:
20044
diff
changeset
|
53 |
| NONE => t :: remove (m, u, terms); |
|
8774
ad5026ff0c16
new simproc, replacing combine_coeffs and working for nat, int, real
paulson
parents:
diff
changeset
|
54 |
|
|
ad5026ff0c16
new simproc, replacing combine_coeffs and working for nat, int, real
paulson
parents:
diff
changeset
|
55 |
(*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
|
56 |
#m*u is already in terms for some m*) |
|
20114
a1bb4bc68ff3
prove_conv: Variable.import_terms instead of Term.addhoc_freeze_vars;
wenzelm
parents:
20044
diff
changeset
|
57 |
fun find_repeated (tab, _, []) = raise TERM("find_repeated", [])
|
|
8774
ad5026ff0c16
new simproc, replacing combine_coeffs and working for nat, int, real
paulson
parents:
diff
changeset
|
58 |
| find_repeated (tab, past, t::terms) = |
| 9646 | 59 |
case try Data.dest_coeff t of |
|
20114
a1bb4bc68ff3
prove_conv: Variable.import_terms instead of Term.addhoc_freeze_vars;
wenzelm
parents:
20044
diff
changeset
|
60 |
SOME(n,u) => |
|
a1bb4bc68ff3
prove_conv: Variable.import_terms instead of Term.addhoc_freeze_vars;
wenzelm
parents:
20044
diff
changeset
|
61 |
(case Termtab.lookup tab u of |
|
a1bb4bc68ff3
prove_conv: Variable.import_terms instead of Term.addhoc_freeze_vars;
wenzelm
parents:
20044
diff
changeset
|
62 |
SOME m => (u, m, n, rev (remove (m,u,past)) @ terms) |
|
a1bb4bc68ff3
prove_conv: Variable.import_terms instead of Term.addhoc_freeze_vars;
wenzelm
parents:
20044
diff
changeset
|
63 |
| NONE => find_repeated (Termtab.update (u, n) tab, |
|
a1bb4bc68ff3
prove_conv: Variable.import_terms instead of Term.addhoc_freeze_vars;
wenzelm
parents:
20044
diff
changeset
|
64 |
t::past, terms)) |
|
a1bb4bc68ff3
prove_conv: Variable.import_terms instead of Term.addhoc_freeze_vars;
wenzelm
parents:
20044
diff
changeset
|
65 |
| NONE => find_repeated (tab, t::past, terms); |
|
8774
ad5026ff0c16
new simproc, replacing combine_coeffs and working for nat, int, real
paulson
parents:
diff
changeset
|
66 |
|
|
ad5026ff0c16
new simproc, replacing combine_coeffs and working for nat, int, real
paulson
parents:
diff
changeset
|
67 |
(*the simplification procedure*) |
| 61144 | 68 |
fun proc ctxt ct = |
|
20044
92cc2f4c7335
simprocs: no theory argument -- use simpset context instead;
wenzelm
parents:
17412
diff
changeset
|
69 |
let |
| 61144 | 70 |
val t = Thm.term_of ct |
| 70326 | 71 |
val (t', ctxt') = yield_singleton (Variable.import_terms true) t ctxt |
|
20114
a1bb4bc68ff3
prove_conv: Variable.import_terms instead of Term.addhoc_freeze_vars;
wenzelm
parents:
20044
diff
changeset
|
72 |
|
|
a1bb4bc68ff3
prove_conv: Variable.import_terms instead of Term.addhoc_freeze_vars;
wenzelm
parents:
20044
diff
changeset
|
73 |
val (u,m,n,terms) = find_repeated (Termtab.empty, [], Data.dest_sum t') |
|
a1bb4bc68ff3
prove_conv: Variable.import_terms instead of Term.addhoc_freeze_vars;
wenzelm
parents:
20044
diff
changeset
|
74 |
val T = Term.fastype_of u |
|
a1bb4bc68ff3
prove_conv: Variable.import_terms instead of Term.addhoc_freeze_vars;
wenzelm
parents:
20044
diff
changeset
|
75 |
|
|
a1bb4bc68ff3
prove_conv: Variable.import_terms instead of Term.addhoc_freeze_vars;
wenzelm
parents:
20044
diff
changeset
|
76 |
val reshape = (*Move i*u to the front and put j*u into standard form |
|
a1bb4bc68ff3
prove_conv: Variable.import_terms instead of Term.addhoc_freeze_vars;
wenzelm
parents:
20044
diff
changeset
|
77 |
i + #m + j + k == #m + i + (j + k) *) |
|
23058
c722004c5a22
generalize CombineNumerals functor to allow coefficients with types other than IntInf.int
huffman
parents:
20114
diff
changeset
|
78 |
if Data.iszero m orelse Data.iszero n then (*trivial, so do nothing*) |
|
20114
a1bb4bc68ff3
prove_conv: Variable.import_terms instead of Term.addhoc_freeze_vars;
wenzelm
parents:
20044
diff
changeset
|
79 |
raise TERM("combine_numerals", [])
|
| 70315 | 80 |
else Data.prove_conv [Data.norm_tac ctxt'] ctxt' |
|
20114
a1bb4bc68ff3
prove_conv: Variable.import_terms instead of Term.addhoc_freeze_vars;
wenzelm
parents:
20044
diff
changeset
|
81 |
(t', Data.mk_sum T ([Data.mk_coeff (m, u), Data.mk_coeff (n, u)] @ terms)) |
|
8774
ad5026ff0c16
new simproc, replacing combine_coeffs and working for nat, int, real
paulson
parents:
diff
changeset
|
82 |
in |
| 70527 | 83 |
Data.prove_conv |
84 |
[Data.trans_tac ctxt' reshape, resolve_tac ctxt' [Data.left_distrib] 1, |
|
85 |
Data.numeral_simp_tac ctxt'] ctxt' |
|
86 |
(t', Data.mk_sum T (Data.mk_coeff(Data.add(m,n), u) :: terms)) |
|
87 |
|> Option.map |
|
88 |
(Data.simplify_meta_eq ctxt' #> |
|
89 |
singleton (Variable.export ctxt' ctxt)) |
|
|
8774
ad5026ff0c16
new simproc, replacing combine_coeffs and working for nat, int, real
paulson
parents:
diff
changeset
|
90 |
end |
|
20114
a1bb4bc68ff3
prove_conv: Variable.import_terms instead of Term.addhoc_freeze_vars;
wenzelm
parents:
20044
diff
changeset
|
91 |
(* FIXME avoid handling of generic exceptions *) |
| 15531 | 92 |
handle TERM _ => NONE |
93 |
| 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
|
94 |
Undeclared type constructor "Numeral.bin"*) |
|
8774
ad5026ff0c16
new simproc, replacing combine_coeffs and working for nat, int, real
paulson
parents:
diff
changeset
|
95 |
|
|
ad5026ff0c16
new simproc, replacing combine_coeffs and working for nat, int, real
paulson
parents:
diff
changeset
|
96 |
end; |