1(*  Title:      Provers/Arith/cancel_numeral_factor.ML
2    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
3    Copyright   2000  University of Cambridge
4
5Cancel common coefficients in balanced expressions:
6
7     u*#m ~~ u'*#m'  ==  #n*u ~~ #n'*u'
8
9where ~~ is an appropriate balancing operation (e.g. =, <=, <, div, /)
10and d = gcd(m,m') and n=m/d and n'=m'/d.
11
12It works by (a) massaging both sides to bring gcd(m,m') to the front:
13
14     u*#m ~~ u'*#m'  ==  #d*(#n*u) ~~ #d*(#n'*u')
15
16(b) then using the rule "cancel" to reach #n*u ~~ #n'*u'.
17*)
18
19signature CANCEL_NUMERAL_FACTOR_DATA =
20sig
21  (*abstract syntax*)
22  val mk_bal: term * term -> term
23  val dest_bal: term -> term * term
24  val mk_coeff: int * term -> term
25  val dest_coeff: term -> int * term
26  (*rules*)
27  val cancel: thm
28  val neg_exchanges: bool  (*true if a negative coeff swaps the two operands,
29                             as with < and <= but not = and div*)
30  (*proof tools*)
31  val prove_conv: tactic list -> Proof.context -> thm list -> term * term -> thm option
32  val trans_tac: Proof.context -> thm option -> tactic  (*applies the initial lemma*)
33  val norm_tac: Proof.context -> tactic              (*proves the initial lemma*)
34  val numeral_simp_tac: Proof.context -> tactic      (*proves the final theorem*)
35  val simplify_meta_eq: Proof.context -> thm -> thm  (*simplifies the final theorem*)
36end;
37
38
39functor CancelNumeralFactorFun(Data: CANCEL_NUMERAL_FACTOR_DATA):
40  sig val proc: Proof.context -> cterm -> thm option end =
41struct
42
43(*the simplification procedure*)
44fun proc ctxt ct =
45  let
46    val prems = Simplifier.prems_of ctxt;
47    val t = Thm.term_of ct;
48    val (t', ctxt') = yield_singleton (Variable.import_terms true) t ctxt
49
50    val (t1,t2) = Data.dest_bal t'
51    val (m1, t1') = Data.dest_coeff t1
52    and (m2, t2') = Data.dest_coeff t2
53    val d = (*if both are negative, also divide through by ~1*)
54      if (m1<0 andalso m2<=0) orelse
55         (m1<=0 andalso m2<0) then ~ (abs (Integer.gcd m1 m2)) else abs (Integer.gcd m1 m2)
56    val _ = if d=1 then   (*trivial, so do nothing*)
57              raise TERM("cancel_numeral_factor", [])
58            else ()
59    fun newshape (i,t) = Data.mk_coeff(d, Data.mk_coeff(i,t))
60    val n1 = m1 div d and n2 = m2 div d
61    val rhs = if d<0 andalso Data.neg_exchanges
62              then Data.mk_bal (Data.mk_coeff(n2,t2'), Data.mk_coeff(n1,t1'))
63              else Data.mk_bal (Data.mk_coeff(n1,t1'), Data.mk_coeff(n2,t2'))
64    val reshape =  (*Move d to the front and put the rest into standard form
65                       i * #m * j == #d * (#n * (j * k)) *)
66      Data.prove_conv [Data.norm_tac ctxt'] ctxt' prems
67        (t', Data.mk_bal (newshape(n1,t1'), newshape(n2,t2')))
68  in
69    Data.prove_conv
70       [Data.trans_tac ctxt' reshape, resolve_tac ctxt' [Data.cancel] 1,
71        Data.numeral_simp_tac ctxt'] ctxt' prems (t', rhs)
72    |> Option.map
73      (Data.simplify_meta_eq ctxt' #>
74        singleton (Variable.export ctxt' ctxt))
75  end
76  (* FIXME avoid handling of generic exceptions *)
77  handle TERM _ => NONE
78       | TYPE _ => NONE;   (*Typically (if thy doesn't include Numeral)
79                             Undeclared type constructor "Numeral.bin"*)
80
81end;
82