| author | ballarin | 
| Tue, 29 Jul 2008 16:19:19 +0200 | |
| changeset 27700 | ef4b26efa8b6 | 
| parent 24630 | 351a308ab58d | 
| child 35762 | af3ff2ba4c54 | 
| permissions | -rw-r--r-- | 
| 8760 | 1 | (* Title: Provers/Arith/cancel_numerals.ML | 
| 8736 | 2 | ID: $Id$ | 
| 3 | Author: Lawrence C Paulson, Cambridge University Computer Laboratory | |
| 4 | Copyright 2000 University of Cambridge | |
| 5 | ||
| 8760 | 6 | Cancel common coefficients in balanced expressions: | 
| 8736 | 7 | |
| 8760 | 8 | i + #m*u + j ~~ i' + #m'*u + j' == #(m-m')*u + i + j ~~ i' + j' | 
| 8736 | 9 | |
| 10 | where ~~ is an appropriate balancing operation (e.g. =, <=, <, -). | |
| 8760 | 11 | |
| 12 | It works by (a) massaging both sides to bring the selected term to the front: | |
| 13 | ||
| 17223 | 14 | #m*u + (i + j) ~~ #m'*u + (i' + j') | 
| 8760 | 15 | |
| 16 | (b) then using bal_add1 or bal_add2 to reach | |
| 17 | ||
| 18 | #(m-m')*u + i + j ~~ i' + j' (if m'<=m) | |
| 19 | ||
| 20 | or | |
| 21 | ||
| 22 | i + j ~~ #(m'-m)*u + i' + j' (otherwise) | |
| 8736 | 23 | *) | 
| 24 | ||
| 25 | signature CANCEL_NUMERALS_DATA = | |
| 26 | sig | |
| 27 | (*abstract syntax*) | |
| 14387 
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
 paulson parents: 
13484diff
changeset | 28 | val mk_sum: typ -> term list -> term | 
| 8736 | 29 | val dest_sum: term -> term list | 
| 30 | val mk_bal: term * term -> term | |
| 31 | val dest_bal: term -> term * term | |
| 24630 
351a308ab58d
simplified type int (eliminated IntInf.int, integer);
 wenzelm parents: 
20114diff
changeset | 32 | val mk_coeff: int * term -> term | 
| 
351a308ab58d
simplified type int (eliminated IntInf.int, integer);
 wenzelm parents: 
20114diff
changeset | 33 | val dest_coeff: term -> int * term | 
| 
351a308ab58d
simplified type int (eliminated IntInf.int, integer);
 wenzelm parents: 
20114diff
changeset | 34 | val find_first_coeff: term -> term list -> int * term list | 
| 8760 | 35 | (*rules*) | 
| 36 | val bal_add1: thm | |
| 37 | val bal_add2: thm | |
| 8736 | 38 | (*proof tools*) | 
| 20114 
a1bb4bc68ff3
prove_conv: Variable.import_terms instead of Term.addhoc_freeze_vars;
 wenzelm parents: 
20044diff
changeset | 39 | val prove_conv: tactic list -> Proof.context -> thm list -> term * term -> thm option | 
| 16973 | 40 | val trans_tac: simpset -> thm option -> tactic (*applies the initial lemma*) | 
| 41 | val norm_tac: simpset -> tactic (*proves the initial lemma*) | |
| 42 | val numeral_simp_tac: simpset -> tactic (*proves the final theorem*) | |
| 43 | val simplify_meta_eq: simpset -> thm -> thm (*simplifies the final theorem*) | |
| 8736 | 44 | end; | 
| 45 | ||
| 46 | ||
| 8760 | 47 | functor CancelNumeralsFun(Data: CANCEL_NUMERALS_DATA): | 
| 48 | sig | |
| 20044 
92cc2f4c7335
simprocs: no theory argument -- use simpset context instead;
 wenzelm parents: 
17412diff
changeset | 49 | val proc: simpset -> term -> thm option | 
| 17223 | 50 | end | 
| 8760 | 51 | = | 
| 8736 | 52 | struct | 
| 53 | ||
| 8779 | 54 | (*For t = #n*u then put u in the table*) | 
| 17223 | 55 | fun update_by_coeff t = | 
| 17412 | 56 | Termtab.update (#2 (Data.dest_coeff t), ()); | 
| 8760 | 57 | |
| 58 | (*a left-to-right scan of terms1, seeking a term of the form #n*u, where | |
| 59 | #m*u is in terms2 for some m*) | |
| 60 | fun find_common (terms1,terms2) = | |
| 17223 | 61 | let val tab2 = fold update_by_coeff terms2 Termtab.empty | 
| 62 |       fun seek [] = raise TERM("find_common", [])
 | |
| 63 | | seek (t::terms) = | |
| 64 | let val (_,u) = Data.dest_coeff t | |
| 65 | in if Termtab.defined tab2 u then u else seek terms end | |
| 8760 | 66 | in seek terms1 end; | 
| 8736 | 67 | |
| 68 | (*the simplification procedure*) | |
| 20044 
92cc2f4c7335
simprocs: no theory argument -- use simpset context instead;
 wenzelm parents: 
17412diff
changeset | 69 | fun proc ss t = | 
| 15027 | 70 | let | 
| 20114 
a1bb4bc68ff3
prove_conv: Variable.import_terms instead of Term.addhoc_freeze_vars;
 wenzelm parents: 
20044diff
changeset | 71 | val ctxt = Simplifier.the_context ss | 
| 
a1bb4bc68ff3
prove_conv: Variable.import_terms instead of Term.addhoc_freeze_vars;
 wenzelm parents: 
20044diff
changeset | 72 | val prems = prems_of_ss ss | 
| 
a1bb4bc68ff3
prove_conv: Variable.import_terms instead of Term.addhoc_freeze_vars;
 wenzelm parents: 
20044diff
changeset | 73 | val ([t'], ctxt') = Variable.import_terms true [t] ctxt | 
| 
a1bb4bc68ff3
prove_conv: Variable.import_terms instead of Term.addhoc_freeze_vars;
 wenzelm parents: 
20044diff
changeset | 74 | val export = singleton (Variable.export ctxt' ctxt) | 
| 
a1bb4bc68ff3
prove_conv: Variable.import_terms instead of Term.addhoc_freeze_vars;
 wenzelm parents: 
20044diff
changeset | 75 | |
| 
a1bb4bc68ff3
prove_conv: Variable.import_terms instead of Term.addhoc_freeze_vars;
 wenzelm parents: 
20044diff
changeset | 76 | val (t1,t2) = Data.dest_bal t' | 
| 
a1bb4bc68ff3
prove_conv: Variable.import_terms instead of Term.addhoc_freeze_vars;
 wenzelm parents: 
20044diff
changeset | 77 | val terms1 = Data.dest_sum t1 | 
| 
a1bb4bc68ff3
prove_conv: Variable.import_terms instead of Term.addhoc_freeze_vars;
 wenzelm parents: 
20044diff
changeset | 78 | and terms2 = Data.dest_sum t2 | 
| 
a1bb4bc68ff3
prove_conv: Variable.import_terms instead of Term.addhoc_freeze_vars;
 wenzelm parents: 
20044diff
changeset | 79 | |
| 
a1bb4bc68ff3
prove_conv: Variable.import_terms instead of Term.addhoc_freeze_vars;
 wenzelm parents: 
20044diff
changeset | 80 | val u = find_common (terms1, terms2) | 
| 
a1bb4bc68ff3
prove_conv: Variable.import_terms instead of Term.addhoc_freeze_vars;
 wenzelm parents: 
20044diff
changeset | 81 | val (n1, terms1') = Data.find_first_coeff u terms1 | 
| 
a1bb4bc68ff3
prove_conv: Variable.import_terms instead of Term.addhoc_freeze_vars;
 wenzelm parents: 
20044diff
changeset | 82 | and (n2, terms2') = Data.find_first_coeff u terms2 | 
| 
a1bb4bc68ff3
prove_conv: Variable.import_terms instead of Term.addhoc_freeze_vars;
 wenzelm parents: 
20044diff
changeset | 83 | and T = Term.fastype_of u | 
| 
a1bb4bc68ff3
prove_conv: Variable.import_terms instead of Term.addhoc_freeze_vars;
 wenzelm parents: 
20044diff
changeset | 84 | |
| 
a1bb4bc68ff3
prove_conv: Variable.import_terms instead of Term.addhoc_freeze_vars;
 wenzelm parents: 
20044diff
changeset | 85 | fun newshape (i,terms) = Data.mk_sum T (Data.mk_coeff(i,u)::terms) | 
| 
a1bb4bc68ff3
prove_conv: Variable.import_terms instead of Term.addhoc_freeze_vars;
 wenzelm parents: 
20044diff
changeset | 86 | val reshape = (*Move i*u to the front and put j*u into standard form | 
| 17223 | 87 | i + #m + j + k == #m + i + (j + k) *) | 
| 20114 
a1bb4bc68ff3
prove_conv: Variable.import_terms instead of Term.addhoc_freeze_vars;
 wenzelm parents: 
20044diff
changeset | 88 | if n1=0 orelse n2=0 then (*trivial, so do nothing*) | 
| 
a1bb4bc68ff3
prove_conv: Variable.import_terms instead of Term.addhoc_freeze_vars;
 wenzelm parents: 
20044diff
changeset | 89 |           raise TERM("cancel_numerals", [])
 | 
| 
a1bb4bc68ff3
prove_conv: Variable.import_terms instead of Term.addhoc_freeze_vars;
 wenzelm parents: 
20044diff
changeset | 90 | else Data.prove_conv [Data.norm_tac ss] ctxt prems | 
| 
a1bb4bc68ff3
prove_conv: Variable.import_terms instead of Term.addhoc_freeze_vars;
 wenzelm parents: 
20044diff
changeset | 91 | (t', Data.mk_bal (newshape(n1,terms1'), newshape(n2,terms2'))) | 
| 8736 | 92 | in | 
| 20114 
a1bb4bc68ff3
prove_conv: Variable.import_terms instead of Term.addhoc_freeze_vars;
 wenzelm parents: 
20044diff
changeset | 93 | Option.map (export o Data.simplify_meta_eq ss) | 
| 
a1bb4bc68ff3
prove_conv: Variable.import_terms instead of Term.addhoc_freeze_vars;
 wenzelm parents: 
20044diff
changeset | 94 | (if n2 <= n1 then | 
| 
a1bb4bc68ff3
prove_conv: Variable.import_terms instead of Term.addhoc_freeze_vars;
 wenzelm parents: 
20044diff
changeset | 95 | Data.prove_conv | 
| 
a1bb4bc68ff3
prove_conv: Variable.import_terms instead of Term.addhoc_freeze_vars;
 wenzelm parents: 
20044diff
changeset | 96 | [Data.trans_tac ss reshape, rtac Data.bal_add1 1, | 
| 
a1bb4bc68ff3
prove_conv: Variable.import_terms instead of Term.addhoc_freeze_vars;
 wenzelm parents: 
20044diff
changeset | 97 | Data.numeral_simp_tac ss] ctxt prems | 
| 
a1bb4bc68ff3
prove_conv: Variable.import_terms instead of Term.addhoc_freeze_vars;
 wenzelm parents: 
20044diff
changeset | 98 | (t', Data.mk_bal (newshape(n1-n2,terms1'), Data.mk_sum T terms2')) | 
| 
a1bb4bc68ff3
prove_conv: Variable.import_terms instead of Term.addhoc_freeze_vars;
 wenzelm parents: 
20044diff
changeset | 99 | else | 
| 
a1bb4bc68ff3
prove_conv: Variable.import_terms instead of Term.addhoc_freeze_vars;
 wenzelm parents: 
20044diff
changeset | 100 | Data.prove_conv | 
| 
a1bb4bc68ff3
prove_conv: Variable.import_terms instead of Term.addhoc_freeze_vars;
 wenzelm parents: 
20044diff
changeset | 101 | [Data.trans_tac ss reshape, rtac Data.bal_add2 1, | 
| 
a1bb4bc68ff3
prove_conv: Variable.import_terms instead of Term.addhoc_freeze_vars;
 wenzelm parents: 
20044diff
changeset | 102 | Data.numeral_simp_tac ss] ctxt prems | 
| 
a1bb4bc68ff3
prove_conv: Variable.import_terms instead of Term.addhoc_freeze_vars;
 wenzelm parents: 
20044diff
changeset | 103 | (t', Data.mk_bal (Data.mk_sum T terms1', newshape(n2-n1,terms2')))) | 
| 8736 | 104 | end | 
| 20114 
a1bb4bc68ff3
prove_conv: Variable.import_terms instead of Term.addhoc_freeze_vars;
 wenzelm parents: 
20044diff
changeset | 105 | (* FIXME avoid handling of generic exceptions *) | 
| 15531 | 106 | handle TERM _ => NONE | 
| 107 | | TYPE _ => NONE; (*Typically (if thy doesn't include Numeral) | |
| 17223 | 108 | Undeclared type constructor "Numeral.bin"*) | 
| 8736 | 109 | |
| 110 | end; |