src/Provers/Arith/abel_cancel.ML
author nipkow
Sat, 25 Jun 2005 16:07:13 +0200
changeset 16569 a12992c34c12
parent 16423 24abe4c0e4b4
child 16973 b2a894562b8f
permissions -rw-r--r--
cancels completely within terms as well now.
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
14756
9c8cc63714f4 modified abel_cancel.ML for polymorphic types
obua
parents: 13480
diff changeset
     1
(*  Title:      HOL/OrderedGroup.ML
5589
94c05305fb29 new simproc functor
paulson
parents:
diff changeset
     2
    ID:         $Id$
94c05305fb29 new simproc functor
paulson
parents:
diff changeset
     3
    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
94c05305fb29 new simproc functor
paulson
parents:
diff changeset
     4
    Copyright   1998  University of Cambridge
94c05305fb29 new simproc functor
paulson
parents:
diff changeset
     5
94c05305fb29 new simproc functor
paulson
parents:
diff changeset
     6
Simplification procedures for abelian groups (e.g. integers, reals)
94c05305fb29 new simproc functor
paulson
parents:
diff changeset
     7
13462
56610e2ba220 sane interface for simprocs;
wenzelm
parents: 12262
diff changeset
     8
- Cancel complementary terms in sums
5589
94c05305fb29 new simproc functor
paulson
parents:
diff changeset
     9
- Cancel like terms on opposite sides of relations
14756
9c8cc63714f4 modified abel_cancel.ML for polymorphic types
obua
parents: 13480
diff changeset
    10
16569
a12992c34c12 cancels completely within terms as well now.
nipkow
parents: 16423
diff changeset
    11
Modification in May 2004 by Steven Obua: polymorphic types work also now
a12992c34c12 cancels completely within terms as well now.
nipkow
parents: 16423
diff changeset
    12
Modification in June 2005 by Tobias Nipkow: cancel_sums works in general now
a12992c34c12 cancels completely within terms as well now.
nipkow
parents: 16423
diff changeset
    13
(using Clemens Ballarin's code for ordered rewriting in abelian groups)
a12992c34c12 cancels completely within terms as well now.
nipkow
parents: 16423
diff changeset
    14
and the whole file is reduced to a fraction of its former size.
5589
94c05305fb29 new simproc functor
paulson
parents:
diff changeset
    15
*)
94c05305fb29 new simproc functor
paulson
parents:
diff changeset
    16
94c05305fb29 new simproc functor
paulson
parents:
diff changeset
    17
signature ABEL_CANCEL =
94c05305fb29 new simproc functor
paulson
parents:
diff changeset
    18
sig
16569
a12992c34c12 cancels completely within terms as well now.
nipkow
parents: 16423
diff changeset
    19
  val cancel_ss       : simpset       (*abelian group cancel simpset*)
a12992c34c12 cancels completely within terms as well now.
nipkow
parents: 16423
diff changeset
    20
  val thy_ref         : theory_ref    (*signature of the theory of the group*)
a12992c34c12 cancels completely within terms as well now.
nipkow
parents: 16423
diff changeset
    21
  val T               : typ           (*the type of group elements*)
a12992c34c12 cancels completely within terms as well now.
nipkow
parents: 16423
diff changeset
    22
  val eq_reflection   : thm           (*object-equality to meta-equality*)
a12992c34c12 cancels completely within terms as well now.
nipkow
parents: 16423
diff changeset
    23
  val eqI_rules       : thm list
a12992c34c12 cancels completely within terms as well now.
nipkow
parents: 16423
diff changeset
    24
  val dest_eqI        : thm -> term
5589
94c05305fb29 new simproc functor
paulson
parents:
diff changeset
    25
end;
94c05305fb29 new simproc functor
paulson
parents:
diff changeset
    26
94c05305fb29 new simproc functor
paulson
parents:
diff changeset
    27
94c05305fb29 new simproc functor
paulson
parents:
diff changeset
    28
functor Abel_Cancel (Data: ABEL_CANCEL) =
94c05305fb29 new simproc functor
paulson
parents:
diff changeset
    29
struct
94c05305fb29 new simproc functor
paulson
parents:
diff changeset
    30
5728
1d1175ef2d56 eliminated 'let ... in structure ...' to make SML/NJ 0.93 happy;
wenzelm
parents: 5610
diff changeset
    31
open Data;
1d1175ef2d56 eliminated 'let ... in structure ...' to make SML/NJ 0.93 happy;
wenzelm
parents: 5610
diff changeset
    32
14756
9c8cc63714f4 modified abel_cancel.ML for polymorphic types
obua
parents: 13480
diff changeset
    33
 fun zero t = Const ("0", t);
9c8cc63714f4 modified abel_cancel.ML for polymorphic types
obua
parents: 13480
diff changeset
    34
 fun minus t = Const ("uminus", t --> t);
5589
94c05305fb29 new simproc functor
paulson
parents:
diff changeset
    35
16569
a12992c34c12 cancels completely within terms as well now.
nipkow
parents: 16423
diff changeset
    36
 fun add_terms pos (Const ("op +", _) $ x $ y, ts) =
a12992c34c12 cancels completely within terms as well now.
nipkow
parents: 16423
diff changeset
    37
         add_terms pos (x, add_terms pos (y, ts))
a12992c34c12 cancels completely within terms as well now.
nipkow
parents: 16423
diff changeset
    38
   | add_terms pos (Const ("op -", _) $ x $ y, ts) =
a12992c34c12 cancels completely within terms as well now.
nipkow
parents: 16423
diff changeset
    39
         add_terms pos (x, add_terms (not pos) (y, ts))
a12992c34c12 cancels completely within terms as well now.
nipkow
parents: 16423
diff changeset
    40
   | add_terms pos (Const ("uminus", _) $ x, ts) =
a12992c34c12 cancels completely within terms as well now.
nipkow
parents: 16423
diff changeset
    41
         add_terms (not pos) (x, ts)
a12992c34c12 cancels completely within terms as well now.
nipkow
parents: 16423
diff changeset
    42
   | add_terms pos (x, ts) = (pos,x) :: ts;
5589
94c05305fb29 new simproc functor
paulson
parents:
diff changeset
    43
16569
a12992c34c12 cancels completely within terms as well now.
nipkow
parents: 16423
diff changeset
    44
 fun terms fml = add_terms true (fml, []);
5589
94c05305fb29 new simproc functor
paulson
parents:
diff changeset
    45
16569
a12992c34c12 cancels completely within terms as well now.
nipkow
parents: 16423
diff changeset
    46
fun zero1 pt (u as (c as Const("op +",_)) $ x $ y) =
a12992c34c12 cancels completely within terms as well now.
nipkow
parents: 16423
diff changeset
    47
      (case zero1 pt x of NONE => (case zero1 pt y of NONE => NONE
a12992c34c12 cancels completely within terms as well now.
nipkow
parents: 16423
diff changeset
    48
                                   | SOME z => SOME(c $ x $ z))
a12992c34c12 cancels completely within terms as well now.
nipkow
parents: 16423
diff changeset
    49
       | SOME z => SOME(c $ z $ y))
a12992c34c12 cancels completely within terms as well now.
nipkow
parents: 16423
diff changeset
    50
  | zero1 (pos,t) (u as (c as Const("op -",_)) $ x $ y) =
a12992c34c12 cancels completely within terms as well now.
nipkow
parents: 16423
diff changeset
    51
      (case zero1 (pos,t) x of
a12992c34c12 cancels completely within terms as well now.
nipkow
parents: 16423
diff changeset
    52
         NONE => (case zero1 (not pos,t) y of NONE => NONE
a12992c34c12 cancels completely within terms as well now.
nipkow
parents: 16423
diff changeset
    53
                  | SOME z => SOME(c $ x $ z))
a12992c34c12 cancels completely within terms as well now.
nipkow
parents: 16423
diff changeset
    54
       | SOME z => SOME(c $ z $ y))
a12992c34c12 cancels completely within terms as well now.
nipkow
parents: 16423
diff changeset
    55
  | zero1 (pos,t) (u as (c as Const("uminus",_)) $ x) =
a12992c34c12 cancels completely within terms as well now.
nipkow
parents: 16423
diff changeset
    56
      (case zero1 (not pos,t) x of NONE => NONE
a12992c34c12 cancels completely within terms as well now.
nipkow
parents: 16423
diff changeset
    57
       | SOME z => SOME(c $ z))
a12992c34c12 cancels completely within terms as well now.
nipkow
parents: 16423
diff changeset
    58
  | zero1 (pos,t) u =
a12992c34c12 cancels completely within terms as well now.
nipkow
parents: 16423
diff changeset
    59
      if pos andalso (t aconv u) then SOME(zero(fastype_of t)) else NONE
5589
94c05305fb29 new simproc functor
paulson
parents:
diff changeset
    60
94c05305fb29 new simproc functor
paulson
parents:
diff changeset
    61
 exception Cancel;
94c05305fb29 new simproc functor
paulson
parents:
diff changeset
    62
94c05305fb29 new simproc functor
paulson
parents:
diff changeset
    63
 val trace = ref false;
94c05305fb29 new simproc functor
paulson
parents:
diff changeset
    64
16569
a12992c34c12 cancels completely within terms as well now.
nipkow
parents: 16423
diff changeset
    65
fun find_common _ [] _ = raise Cancel
a12992c34c12 cancels completely within terms as well now.
nipkow
parents: 16423
diff changeset
    66
  | find_common opp ((p,l)::ls) rs =
a12992c34c12 cancels completely within terms as well now.
nipkow
parents: 16423
diff changeset
    67
      let val pr = if opp then not p else p
a12992c34c12 cancels completely within terms as well now.
nipkow
parents: 16423
diff changeset
    68
      in if exists (fn (q,r) => pr = q andalso l aconv r) rs then (p,l)
a12992c34c12 cancels completely within terms as well now.
nipkow
parents: 16423
diff changeset
    69
         else find_common opp ls rs
a12992c34c12 cancels completely within terms as well now.
nipkow
parents: 16423
diff changeset
    70
      end
5589
94c05305fb29 new simproc functor
paulson
parents:
diff changeset
    71
16569
a12992c34c12 cancels completely within terms as well now.
nipkow
parents: 16423
diff changeset
    72
(* turns t1(t) OP t2(t) into t1(0) OP t2(0) where OP can be +, -, =, etc.
a12992c34c12 cancels completely within terms as well now.
nipkow
parents: 16423
diff changeset
    73
   If OP = +, it must be t2(-t) rather than t2(t)
a12992c34c12 cancels completely within terms as well now.
nipkow
parents: 16423
diff changeset
    74
*)
a12992c34c12 cancels completely within terms as well now.
nipkow
parents: 16423
diff changeset
    75
fun cancel sg t =
a12992c34c12 cancels completely within terms as well now.
nipkow
parents: 16423
diff changeset
    76
  let val _ = if !trace
a12992c34c12 cancels completely within terms as well now.
nipkow
parents: 16423
diff changeset
    77
              then tracing ("Abel cancel: " ^ string_of_cterm (cterm_of sg t))
a12992c34c12 cancels completely within terms as well now.
nipkow
parents: 16423
diff changeset
    78
              else ()
a12992c34c12 cancels completely within terms as well now.
nipkow
parents: 16423
diff changeset
    79
      val c $ lhs $ rhs = t
a12992c34c12 cancels completely within terms as well now.
nipkow
parents: 16423
diff changeset
    80
      val opp = case c of Const("op +",_) => true | _ => false;
a12992c34c12 cancels completely within terms as well now.
nipkow
parents: 16423
diff changeset
    81
      val (pos,l) = find_common opp (terms lhs) (terms rhs)
a12992c34c12 cancels completely within terms as well now.
nipkow
parents: 16423
diff changeset
    82
      val posr = if opp then not pos else pos
a12992c34c12 cancels completely within terms as well now.
nipkow
parents: 16423
diff changeset
    83
      val t' = c $ (the(zero1 (pos,l) lhs)) $ (the(zero1 (posr,l) rhs))
a12992c34c12 cancels completely within terms as well now.
nipkow
parents: 16423
diff changeset
    84
      val _ = if !trace
a12992c34c12 cancels completely within terms as well now.
nipkow
parents: 16423
diff changeset
    85
              then tracing ("cancelled: " ^ string_of_cterm (cterm_of sg t'))
a12992c34c12 cancels completely within terms as well now.
nipkow
parents: 16423
diff changeset
    86
              else ()
a12992c34c12 cancels completely within terms as well now.
nipkow
parents: 16423
diff changeset
    87
  in t' end;
a12992c34c12 cancels completely within terms as well now.
nipkow
parents: 16423
diff changeset
    88
a12992c34c12 cancels completely within terms as well now.
nipkow
parents: 16423
diff changeset
    89
(* A simproc to cancel complementary terms in arbitrary sums. *)
a12992c34c12 cancels completely within terms as well now.
nipkow
parents: 16423
diff changeset
    90
a12992c34c12 cancels completely within terms as well now.
nipkow
parents: 16423
diff changeset
    91
fun sum_proc sg _ t =
a12992c34c12 cancels completely within terms as well now.
nipkow
parents: 16423
diff changeset
    92
   let val t' = cancel sg t
a12992c34c12 cancels completely within terms as well now.
nipkow
parents: 16423
diff changeset
    93
       val thm = Tactic.prove sg [] [] (Logic.mk_equals (t, t'))
a12992c34c12 cancels completely within terms as well now.
nipkow
parents: 16423
diff changeset
    94
                   (fn _ => simp_tac cancel_ss 1)
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15183
diff changeset
    95
   in SOME thm end
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15183
diff changeset
    96
   handle Cancel => NONE;
5589
94c05305fb29 new simproc functor
paulson
parents:
diff changeset
    97
13462
56610e2ba220 sane interface for simprocs;
wenzelm
parents: 12262
diff changeset
    98
 val sum_conv =
5589
94c05305fb29 new simproc functor
paulson
parents:
diff changeset
    99
     Simplifier.mk_simproc "cancel_sums"
16423
24abe4c0e4b4 renamed sg_ref to thy_ref;
wenzelm
parents: 15796
diff changeset
   100
       (map (Drule.cterm_fun Logic.varify o Thm.read_cterm (Theory.deref thy_ref))
13462
56610e2ba220 sane interface for simprocs;
wenzelm
parents: 12262
diff changeset
   101
        [("x + y", Data.T), ("x - y", Data.T)])  (* FIXME depends on concrete syntax !???!!??! *)
5589
94c05305fb29 new simproc functor
paulson
parents:
diff changeset
   102
       sum_proc;
94c05305fb29 new simproc functor
paulson
parents:
diff changeset
   103
94c05305fb29 new simproc functor
paulson
parents:
diff changeset
   104
94c05305fb29 new simproc functor
paulson
parents:
diff changeset
   105
 (*A simproc to cancel like terms on the opposite sides of relations:
94c05305fb29 new simproc functor
paulson
parents:
diff changeset
   106
     (x + y - z < -z + x) = (y < 0)
94c05305fb29 new simproc functor
paulson
parents:
diff changeset
   107
   Works for (=) and (<=) as well as (<), if the necessary rules are supplied.
16569
a12992c34c12 cancels completely within terms as well now.
nipkow
parents: 16423
diff changeset
   108
   Reduces the problem to subtraction.
5589
94c05305fb29 new simproc functor
paulson
parents:
diff changeset
   109
 *)
94c05305fb29 new simproc functor
paulson
parents:
diff changeset
   110
16569
a12992c34c12 cancels completely within terms as well now.
nipkow
parents: 16423
diff changeset
   111
 fun rel_proc sg _ t =
a12992c34c12 cancels completely within terms as well now.
nipkow
parents: 16423
diff changeset
   112
   let val t' = cancel sg t
a12992c34c12 cancels completely within terms as well now.
nipkow
parents: 16423
diff changeset
   113
       val thm = Tactic.prove sg [] [] (Logic.mk_equals (t, t'))
13480
bb72bd43c6c3 use Tactic.prove instead of prove_goalw_cterm in internal proofs!
wenzelm
parents: 13462
diff changeset
   114
                   (fn _ => rtac eq_reflection 1 THEN
bb72bd43c6c3 use Tactic.prove instead of prove_goalw_cterm in internal proofs!
wenzelm
parents: 13462
diff changeset
   115
                            resolve_tac eqI_rules 1 THEN
16569
a12992c34c12 cancels completely within terms as well now.
nipkow
parents: 16423
diff changeset
   116
                            simp_tac cancel_ss 1)
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15183
diff changeset
   117
   in SOME thm end
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15183
diff changeset
   118
   handle Cancel => NONE;
5589
94c05305fb29 new simproc functor
paulson
parents:
diff changeset
   119
13462
56610e2ba220 sane interface for simprocs;
wenzelm
parents: 12262
diff changeset
   120
 val rel_conv =
16423
24abe4c0e4b4 renamed sg_ref to thy_ref;
wenzelm
parents: 15796
diff changeset
   121
     Simplifier.simproc_i (Theory.deref thy_ref) "cancel_relations"
13462
56610e2ba220 sane interface for simprocs;
wenzelm
parents: 12262
diff changeset
   122
       (map Data.dest_eqI eqI_rules) rel_proc;
5589
94c05305fb29 new simproc functor
paulson
parents:
diff changeset
   123
94c05305fb29 new simproc functor
paulson
parents:
diff changeset
   124
end;