File ‹~~/src/Provers/Arith/cancel_numerals.ML›
signature CANCEL_NUMERALS_DATA =
sig
  
  val mk_sum: typ -> term list -> term
  val dest_sum: term -> term list
  val mk_bal: term * term -> term
  val dest_bal: term -> term * term
  val mk_coeff: int * term -> term
  val dest_coeff: term -> int * term
  val find_first_coeff: term -> term list -> int * term list
  
  val bal_add1: thm
  val bal_add2: thm
  
  val prove_conv: tactic list -> Proof.context -> thm list -> 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;
signature CANCEL_NUMERALS =
sig
  val proc: Simplifier.proc
end;
functor CancelNumeralsFun(Data: CANCEL_NUMERALS_DATA): CANCEL_NUMERALS =
struct
fun update_by_coeff t =
  Termtab.update (#2 (Data.dest_coeff t), ());
fun find_common (terms1,terms2) =
  let val tab2 = fold update_by_coeff terms2 Termtab.empty
      fun seek [] = raise TERM("find_common", [])
        | seek (t::terms) =
              let val (_,u) = Data.dest_coeff t
              in if Termtab.defined tab2 u then u else seek terms end
  in  seek terms1 end;
fun proc ctxt ct =
  let
    val prems = Simplifier.prems_of ctxt
    val t = Thm.term_of ct
    val (t', ctxt') = yield_singleton (Variable.import_terms true) t ctxt
    val (t1,t2) = Data.dest_bal t'
    val terms1 = Data.dest_sum t1
    and terms2 = Data.dest_sum t2
    val u = find_common (terms1, terms2)
    val (n1, terms1') = Data.find_first_coeff u terms1
    and (n2, terms2') = Data.find_first_coeff u terms2
    and T = Term.fastype_of u
    fun newshape (i,terms) = Data.mk_sum T (Data.mk_coeff(i,u)::terms)
    val reshape =  
        if n1=0 orelse n2=0 then   
          raise TERM("cancel_numerals", [])
        else Data.prove_conv [Data.norm_tac ctxt'] ctxt' prems
          (t', Data.mk_bal (newshape(n1,terms1'), newshape(n2,terms2')))
  in
    (if n2 <= n1 then
       Data.prove_conv
         [Data.trans_tac ctxt' reshape, resolve_tac ctxt' [Data.bal_add1] 1,
          Data.numeral_simp_tac ctxt'] ctxt' prems
         (t', Data.mk_bal (newshape(n1-n2,terms1'), Data.mk_sum T terms2'))
     else
       Data.prove_conv
         [Data.trans_tac ctxt' reshape, resolve_tac ctxt' [Data.bal_add2] 1,
          Data.numeral_simp_tac ctxt'] ctxt' prems
         (t', Data.mk_bal (Data.mk_sum T terms1', newshape(n2-n1,terms2'))))
    |> Option.map
      (Data.simplify_meta_eq ctxt' #>
        singleton (Variable.export ctxt' ctxt))
  end
  
  handle TERM _ => NONE
       | TYPE _ => NONE;   
end;