paulson@8760: (* Title: Provers/Arith/cancel_numerals.ML paulson@8736: Author: Lawrence C Paulson, Cambridge University Computer Laboratory paulson@8736: Copyright 2000 University of Cambridge paulson@8736: paulson@8760: Cancel common coefficients in balanced expressions: paulson@8736: paulson@8760: i + #m*u + j ~~ i' + #m'*u + j' == #(m-m')*u + i + j ~~ i' + j' paulson@8736: paulson@8736: where ~~ is an appropriate balancing operation (e.g. =, <=, <, -). paulson@8760: paulson@8760: It works by (a) massaging both sides to bring the selected term to the front: paulson@8760: wenzelm@17223: #m*u + (i + j) ~~ #m'*u + (i' + j') paulson@8760: paulson@8760: (b) then using bal_add1 or bal_add2 to reach paulson@8760: paulson@8760: #(m-m')*u + i + j ~~ i' + j' (if m'<=m) paulson@8760: paulson@8760: or paulson@8760: paulson@8760: i + j ~~ #(m'-m)*u + i' + j' (otherwise) paulson@8736: *) paulson@8736: paulson@8736: signature CANCEL_NUMERALS_DATA = paulson@8736: sig paulson@8736: (*abstract syntax*) paulson@14387: val mk_sum: typ -> term list -> term paulson@8736: val dest_sum: term -> term list paulson@8736: val mk_bal: term * term -> term paulson@8736: val dest_bal: term -> term * term wenzelm@24630: val mk_coeff: int * term -> term wenzelm@24630: val dest_coeff: term -> int * term wenzelm@24630: val find_first_coeff: term -> term list -> int * term list paulson@8760: (*rules*) paulson@8760: val bal_add1: thm paulson@8760: val bal_add2: thm paulson@8736: (*proof tools*) wenzelm@20114: val prove_conv: tactic list -> Proof.context -> thm list -> term * term -> thm option haftmann@44947: val trans_tac: thm option -> tactic (*applies the initial lemma*) wenzelm@51717: val norm_tac: Proof.context -> tactic (*proves the initial lemma*) wenzelm@51717: val numeral_simp_tac: Proof.context -> tactic (*proves the final theorem*) wenzelm@51717: val simplify_meta_eq: Proof.context -> thm -> thm (*simplifies the final theorem*) paulson@8736: end; paulson@8736: wenzelm@42462: signature CANCEL_NUMERALS = wenzelm@42462: sig wenzelm@51717: val proc: Proof.context -> term -> thm option wenzelm@42462: end; paulson@8736: wenzelm@42462: functor CancelNumeralsFun(Data: CANCEL_NUMERALS_DATA): CANCEL_NUMERALS = paulson@8736: struct paulson@8736: paulson@8779: (*For t = #n*u then put u in the table*) wenzelm@17223: fun update_by_coeff t = wenzelm@17412: Termtab.update (#2 (Data.dest_coeff t), ()); paulson@8760: paulson@8760: (*a left-to-right scan of terms1, seeking a term of the form #n*u, where paulson@8760: #m*u is in terms2 for some m*) paulson@8760: fun find_common (terms1,terms2) = wenzelm@17223: let val tab2 = fold update_by_coeff terms2 Termtab.empty wenzelm@17223: fun seek [] = raise TERM("find_common", []) wenzelm@17223: | seek (t::terms) = wenzelm@17223: let val (_,u) = Data.dest_coeff t wenzelm@17223: in if Termtab.defined tab2 u then u else seek terms end paulson@8760: in seek terms1 end; paulson@8736: paulson@8736: (*the simplification procedure*) wenzelm@51717: fun proc ctxt t = wenzelm@15027: let wenzelm@51717: val prems = Simplifier.prems_of ctxt wenzelm@20114: val ([t'], ctxt') = Variable.import_terms true [t] ctxt wenzelm@20114: val export = singleton (Variable.export ctxt' ctxt) wenzelm@54565: (* FIXME ctxt vs. ctxt' (!?) *) wenzelm@20114: wenzelm@20114: val (t1,t2) = Data.dest_bal t' wenzelm@20114: val terms1 = Data.dest_sum t1 wenzelm@20114: and terms2 = Data.dest_sum t2 wenzelm@20114: wenzelm@20114: val u = find_common (terms1, terms2) wenzelm@20114: val (n1, terms1') = Data.find_first_coeff u terms1 wenzelm@20114: and (n2, terms2') = Data.find_first_coeff u terms2 wenzelm@20114: and T = Term.fastype_of u wenzelm@20114: wenzelm@20114: fun newshape (i,terms) = Data.mk_sum T (Data.mk_coeff(i,u)::terms) wenzelm@20114: val reshape = (*Move i*u to the front and put j*u into standard form wenzelm@17223: i + #m + j + k == #m + i + (j + k) *) wenzelm@20114: if n1=0 orelse n2=0 then (*trivial, so do nothing*) wenzelm@20114: raise TERM("cancel_numerals", []) wenzelm@51717: else Data.prove_conv [Data.norm_tac ctxt] ctxt prems wenzelm@20114: (t', Data.mk_bal (newshape(n1,terms1'), newshape(n2,terms2'))) paulson@8736: in wenzelm@51717: Option.map (export o Data.simplify_meta_eq ctxt) wenzelm@20114: (if n2 <= n1 then wenzelm@20114: Data.prove_conv wenzelm@58838: [Data.trans_tac reshape, resolve_tac [Data.bal_add1] 1, wenzelm@51717: Data.numeral_simp_tac ctxt] ctxt prems wenzelm@20114: (t', Data.mk_bal (newshape(n1-n2,terms1'), Data.mk_sum T terms2')) wenzelm@20114: else wenzelm@20114: Data.prove_conv wenzelm@58838: [Data.trans_tac reshape, resolve_tac [Data.bal_add2] 1, wenzelm@51717: Data.numeral_simp_tac ctxt] ctxt prems wenzelm@20114: (t', Data.mk_bal (Data.mk_sum T terms1', newshape(n2-n1,terms2')))) paulson@8736: end wenzelm@20114: (* FIXME avoid handling of generic exceptions *) skalberg@15531: handle TERM _ => NONE skalberg@15531: | TYPE _ => NONE; (*Typically (if thy doesn't include Numeral) wenzelm@17223: Undeclared type constructor "Numeral.bin"*) paulson@8736: paulson@8736: end;