File ‹~~/src/Provers/Arith/combine_numerals.ML›
signature COMBINE_NUMERALS_DATA =
sig
  
  eqtype coeff
  val iszero: coeff -> bool
  val add: coeff * coeff -> coeff     
  val mk_sum: typ -> term list -> term
  val dest_sum: term -> term list
  val mk_coeff: coeff * term -> term
  val dest_coeff: term -> coeff * term
  
  val left_distrib: thm
  
  val prove_conv: tactic list -> Proof.context -> term * term -> thm option
  val trans_tac: Proof.context -> thm option -> tactic            
  val norm_tac: Proof.context -> tactic          
  val numeral_simp_tac: Proof.context -> tactic  
  val simplify_meta_eq: Proof.context -> thm -> thm  
end;
functor CombineNumeralsFun(Data: COMBINE_NUMERALS_DATA):
  sig
  val proc: Simplifier.proc
  end
=
struct
fun remove (_, _, []) = 
      raise TERM("combine_numerals: remove", [])
  | remove (m, u, t::terms) =
      case try Data.dest_coeff t of
          SOME(n,v) => if m=n andalso u aconv v then terms
                       else t :: remove (m, u, terms)
        | NONE      =>  t :: remove (m, u, terms);
fun find_repeated (tab, _, []) = raise TERM("find_repeated", [])
  | find_repeated (tab, past, t::terms) =
      case try Data.dest_coeff t of
          SOME(n,u) =>
              (case Termtab.lookup tab u of
                  SOME m => (u, m, n, rev (remove (m,u,past)) @ terms)
                | NONE => find_repeated (Termtab.update (u, n) tab,
                                         t::past,  terms))
        | NONE => find_repeated (tab, t::past, terms);
fun proc ctxt ct =
  let
    val t = Thm.term_of ct
    val (t', ctxt') = yield_singleton (Variable.import_terms true) t ctxt
    val (u,m,n,terms) = find_repeated (Termtab.empty, [], Data.dest_sum t')
    val T = Term.fastype_of u
    val reshape =  
      if Data.iszero m orelse Data.iszero n then   
        raise TERM("combine_numerals", [])
      else Data.prove_conv [Data.norm_tac ctxt'] ctxt'
        (t', Data.mk_sum T ([Data.mk_coeff (m, u), Data.mk_coeff (n, u)] @ terms))
  in
    Data.prove_conv
       [Data.trans_tac ctxt' reshape, resolve_tac ctxt' [Data.left_distrib] 1,
        Data.numeral_simp_tac ctxt'] ctxt'
       (t', Data.mk_sum T (Data.mk_coeff(Data.add(m,n), u) :: terms))
    |> Option.map
      (Data.simplify_meta_eq ctxt' #>
        singleton (Variable.export ctxt' ctxt))
  end
  
  handle TERM _ => NONE
       | TYPE _ => NONE;   
end;