src/Provers/Arith/cancel_numeral_factor.ML
author paulson
Mon, 16 May 2005 10:29:15 +0200
changeset 15965 f422f8283491
parent 15570 8d8c70b41bab
child 16973 b2a894562b8f
permissions -rw-r--r--
Use of IntInf.int instead of int in most numeric simprocs; avoids integer overflow in SML/NJ
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
10537
1d2f15504d38 simproc for cancelling common factors around = < <= div /
paulson
parents:
diff changeset
     1
(*  Title:      Provers/Arith/cancel_numeral_factor.ML
1d2f15504d38 simproc for cancelling common factors around = < <= div /
paulson
parents:
diff changeset
     2
    ID:         $Id$
1d2f15504d38 simproc for cancelling common factors around = < <= div /
paulson
parents:
diff changeset
     3
    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
1d2f15504d38 simproc for cancelling common factors around = < <= div /
paulson
parents:
diff changeset
     4
    Copyright   2000  University of Cambridge
1d2f15504d38 simproc for cancelling common factors around = < <= div /
paulson
parents:
diff changeset
     5
1d2f15504d38 simproc for cancelling common factors around = < <= div /
paulson
parents:
diff changeset
     6
Cancel common coefficients in balanced expressions:
1d2f15504d38 simproc for cancelling common factors around = < <= div /
paulson
parents:
diff changeset
     7
1d2f15504d38 simproc for cancelling common factors around = < <= div /
paulson
parents:
diff changeset
     8
     u*#m ~~ u'*#m'  ==  #n*u ~~ #n'*u'
1d2f15504d38 simproc for cancelling common factors around = < <= div /
paulson
parents:
diff changeset
     9
1d2f15504d38 simproc for cancelling common factors around = < <= div /
paulson
parents:
diff changeset
    10
where ~~ is an appropriate balancing operation (e.g. =, <=, <, div, /)
1d2f15504d38 simproc for cancelling common factors around = < <= div /
paulson
parents:
diff changeset
    11
and d = gcd(m,m') and n=m/d and n'=m'/d.
1d2f15504d38 simproc for cancelling common factors around = < <= div /
paulson
parents:
diff changeset
    12
1d2f15504d38 simproc for cancelling common factors around = < <= div /
paulson
parents:
diff changeset
    13
It works by (a) massaging both sides to bring gcd(m,m') to the front:
1d2f15504d38 simproc for cancelling common factors around = < <= div /
paulson
parents:
diff changeset
    14
1d2f15504d38 simproc for cancelling common factors around = < <= div /
paulson
parents:
diff changeset
    15
     u*#m ~~ u'*#m'  ==  #d*(#n*u) ~~ #d*(#n'*u')
1d2f15504d38 simproc for cancelling common factors around = < <= div /
paulson
parents:
diff changeset
    16
1d2f15504d38 simproc for cancelling common factors around = < <= div /
paulson
parents:
diff changeset
    17
(b) then using the rule "cancel" to reach #n*u ~~ #n'*u'.
1d2f15504d38 simproc for cancelling common factors around = < <= div /
paulson
parents:
diff changeset
    18
*)
1d2f15504d38 simproc for cancelling common factors around = < <= div /
paulson
parents:
diff changeset
    19
1d2f15504d38 simproc for cancelling common factors around = < <= div /
paulson
parents:
diff changeset
    20
signature CANCEL_NUMERAL_FACTOR_DATA =
1d2f15504d38 simproc for cancelling common factors around = < <= div /
paulson
parents:
diff changeset
    21
sig
1d2f15504d38 simproc for cancelling common factors around = < <= div /
paulson
parents:
diff changeset
    22
  (*abstract syntax*)
1d2f15504d38 simproc for cancelling common factors around = < <= div /
paulson
parents:
diff changeset
    23
  val mk_bal: term * term -> term
1d2f15504d38 simproc for cancelling common factors around = < <= div /
paulson
parents:
diff changeset
    24
  val dest_bal: term -> term * term
15965
f422f8283491 Use of IntInf.int instead of int in most numeric simprocs; avoids
paulson
parents: 15570
diff changeset
    25
  val mk_coeff: IntInf.int * term -> term
f422f8283491 Use of IntInf.int instead of int in most numeric simprocs; avoids
paulson
parents: 15570
diff changeset
    26
  val dest_coeff: term -> IntInf.int * term
10537
1d2f15504d38 simproc for cancelling common factors around = < <= div /
paulson
parents:
diff changeset
    27
  (*rules*)
1d2f15504d38 simproc for cancelling common factors around = < <= div /
paulson
parents:
diff changeset
    28
  val cancel: thm
1d2f15504d38 simproc for cancelling common factors around = < <= div /
paulson
parents:
diff changeset
    29
  val neg_exchanges: bool  (*true if a negative coeff swaps the two operands,
1d2f15504d38 simproc for cancelling common factors around = < <= div /
paulson
parents:
diff changeset
    30
                             as with < and <= but not = and div*)
1d2f15504d38 simproc for cancelling common factors around = < <= div /
paulson
parents:
diff changeset
    31
  (*proof tools*)
1d2f15504d38 simproc for cancelling common factors around = < <= div /
paulson
parents:
diff changeset
    32
  val prove_conv: tactic list -> Sign.sg -> 
13484
d8f5d3391766 adhoc_freeze_vars;
wenzelm
parents: 10656
diff changeset
    33
                  thm list -> string list -> term * term -> thm option
10537
1d2f15504d38 simproc for cancelling common factors around = < <= div /
paulson
parents:
diff changeset
    34
  val trans_tac: thm option -> tactic (*applies the initial lemma*)
1d2f15504d38 simproc for cancelling common factors around = < <= div /
paulson
parents:
diff changeset
    35
  val norm_tac: tactic                (*proves the initial lemma*)
1d2f15504d38 simproc for cancelling common factors around = < <= div /
paulson
parents:
diff changeset
    36
  val numeral_simp_tac: tactic        (*proves the final theorem*)
1d2f15504d38 simproc for cancelling common factors around = < <= div /
paulson
parents:
diff changeset
    37
  val simplify_meta_eq: thm -> thm    (*simplifies the final theorem*)
1d2f15504d38 simproc for cancelling common factors around = < <= div /
paulson
parents:
diff changeset
    38
end;
1d2f15504d38 simproc for cancelling common factors around = < <= div /
paulson
parents:
diff changeset
    39
1d2f15504d38 simproc for cancelling common factors around = < <= div /
paulson
parents:
diff changeset
    40
1d2f15504d38 simproc for cancelling common factors around = < <= div /
paulson
parents:
diff changeset
    41
functor CancelNumeralFactorFun(Data: CANCEL_NUMERAL_FACTOR_DATA):
1d2f15504d38 simproc for cancelling common factors around = < <= div /
paulson
parents:
diff changeset
    42
  sig
15027
d23887300b96 adapted type of simprocs;
wenzelm
parents: 13484
diff changeset
    43
  val proc: Sign.sg -> simpset -> term -> thm option
10537
1d2f15504d38 simproc for cancelling common factors around = < <= div /
paulson
parents:
diff changeset
    44
  end 
1d2f15504d38 simproc for cancelling common factors around = < <= div /
paulson
parents:
diff changeset
    45
=
1d2f15504d38 simproc for cancelling common factors around = < <= div /
paulson
parents:
diff changeset
    46
struct
1d2f15504d38 simproc for cancelling common factors around = < <= div /
paulson
parents:
diff changeset
    47
1d2f15504d38 simproc for cancelling common factors around = < <= div /
paulson
parents:
diff changeset
    48
(*the simplification procedure*)
15027
d23887300b96 adapted type of simprocs;
wenzelm
parents: 13484
diff changeset
    49
fun proc sg ss t =
d23887300b96 adapted type of simprocs;
wenzelm
parents: 13484
diff changeset
    50
  let
d23887300b96 adapted type of simprocs;
wenzelm
parents: 13484
diff changeset
    51
      val hyps = prems_of_ss ss;
d23887300b96 adapted type of simprocs;
wenzelm
parents: 13484
diff changeset
    52
      (*first freeze any Vars in the term to prevent flex-flex problems*)
13484
d8f5d3391766 adhoc_freeze_vars;
wenzelm
parents: 10656
diff changeset
    53
      val (t', xs) = Term.adhoc_freeze_vars t;
10537
1d2f15504d38 simproc for cancelling common factors around = < <= div /
paulson
parents:
diff changeset
    54
      val (t1,t2) = Data.dest_bal t' 
1d2f15504d38 simproc for cancelling common factors around = < <= div /
paulson
parents:
diff changeset
    55
      val (m1, t1') = Data.dest_coeff t1
1d2f15504d38 simproc for cancelling common factors around = < <= div /
paulson
parents:
diff changeset
    56
      and (m2, t2') = Data.dest_coeff t2
1d2f15504d38 simproc for cancelling common factors around = < <= div /
paulson
parents:
diff changeset
    57
      val d = (*if both are negative, also divide through by ~1*)
10656
68f3fddd6e24 tries harder to remove negative literals, e.g.
paulson
parents: 10537
diff changeset
    58
          if (m1<0 andalso m2<=0) orelse
15965
f422f8283491 Use of IntInf.int instead of int in most numeric simprocs; avoids
paulson
parents: 15570
diff changeset
    59
             (m1<=0 andalso m2<0) then ~ (abs(gcd(m1,m2))) else abs (gcd(m1,m2))
10537
1d2f15504d38 simproc for cancelling common factors around = < <= div /
paulson
parents:
diff changeset
    60
      val _ = if d=1 then   (*trivial, so do nothing*)
1d2f15504d38 simproc for cancelling common factors around = < <= div /
paulson
parents:
diff changeset
    61
		      raise TERM("cancel_numeral_factor", []) 
1d2f15504d38 simproc for cancelling common factors around = < <= div /
paulson
parents:
diff changeset
    62
              else ()
1d2f15504d38 simproc for cancelling common factors around = < <= div /
paulson
parents:
diff changeset
    63
      fun newshape (i,t) = Data.mk_coeff(d, Data.mk_coeff(i,t))
1d2f15504d38 simproc for cancelling common factors around = < <= div /
paulson
parents:
diff changeset
    64
      val n1 = m1 div d and n2 = m2 div d
1d2f15504d38 simproc for cancelling common factors around = < <= div /
paulson
parents:
diff changeset
    65
      val rhs = if d<0 andalso Data.neg_exchanges
1d2f15504d38 simproc for cancelling common factors around = < <= div /
paulson
parents:
diff changeset
    66
                then Data.mk_bal (Data.mk_coeff(n2,t2'), Data.mk_coeff(n1,t1'))
1d2f15504d38 simproc for cancelling common factors around = < <= div /
paulson
parents:
diff changeset
    67
                else Data.mk_bal (Data.mk_coeff(n1,t1'), Data.mk_coeff(n2,t2'))
1d2f15504d38 simproc for cancelling common factors around = < <= div /
paulson
parents:
diff changeset
    68
      val reshape =  (*Move d to the front and put the rest into standard form
1d2f15504d38 simproc for cancelling common factors around = < <= div /
paulson
parents:
diff changeset
    69
		       i * #m * j == #d * (#n * (j * k)) *)
13484
d8f5d3391766 adhoc_freeze_vars;
wenzelm
parents: 10656
diff changeset
    70
	    Data.prove_conv [Data.norm_tac] sg hyps xs
10537
1d2f15504d38 simproc for cancelling common factors around = < <= div /
paulson
parents:
diff changeset
    71
	      (t',   Data.mk_bal (newshape(n1,t1'), newshape(n2,t2')))
1d2f15504d38 simproc for cancelling common factors around = < <= div /
paulson
parents:
diff changeset
    72
  in
15570
8d8c70b41bab Move towards standard functions.
skalberg
parents: 15531
diff changeset
    73
      Option.map Data.simplify_meta_eq
10537
1d2f15504d38 simproc for cancelling common factors around = < <= div /
paulson
parents:
diff changeset
    74
       (Data.prove_conv 
1d2f15504d38 simproc for cancelling common factors around = < <= div /
paulson
parents:
diff changeset
    75
	       [Data.trans_tac reshape, rtac Data.cancel 1,
13484
d8f5d3391766 adhoc_freeze_vars;
wenzelm
parents: 10656
diff changeset
    76
		Data.numeral_simp_tac] sg hyps xs (t', rhs))
10537
1d2f15504d38 simproc for cancelling common factors around = < <= div /
paulson
parents:
diff changeset
    77
  end
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15027
diff changeset
    78
  handle TERM _ => NONE
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15027
diff changeset
    79
       | TYPE _ => NONE;   (*Typically (if thy doesn't include Numeral)
10537
1d2f15504d38 simproc for cancelling common factors around = < <= div /
paulson
parents:
diff changeset
    80
			     Undeclared type constructor "Numeral.bin"*)
1d2f15504d38 simproc for cancelling common factors around = < <= div /
paulson
parents:
diff changeset
    81
1d2f15504d38 simproc for cancelling common factors around = < <= div /
paulson
parents:
diff changeset
    82
end;