src/Provers/Arith/abel_cancel.ML
author haftmann
Sat, 15 Sep 2007 19:27:35 +0200
changeset 24584 01e83ffa6c54
parent 22997 d4f3b015b50b
child 27250 7eef2b183032
permissions -rw-r--r--
fixed title
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
24584
01e83ffa6c54 fixed title
haftmann
parents: 22997
diff changeset
     1
(*  Title:      Provers/Arith/abel_cancel.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
20055
24c127b97ab5 simprocs: no theory argument -- use simpset context instead;
wenzelm
parents: 19798
diff changeset
     6
Simplification procedures for abelian groups (e.g. integers, reals,
24c127b97ab5 simprocs: no theory argument -- use simpset context instead;
wenzelm
parents: 19798
diff changeset
     7
polymorphic types).
5589
94c05305fb29 new simproc functor
paulson
parents:
diff changeset
     8
13462
56610e2ba220 sane interface for simprocs;
wenzelm
parents: 12262
diff changeset
     9
- Cancel complementary terms in sums
5589
94c05305fb29 new simproc functor
paulson
parents:
diff changeset
    10
- Cancel like terms on opposite sides of relations
94c05305fb29 new simproc functor
paulson
parents:
diff changeset
    11
*)
94c05305fb29 new simproc functor
paulson
parents:
diff changeset
    12
94c05305fb29 new simproc functor
paulson
parents:
diff changeset
    13
signature ABEL_CANCEL =
94c05305fb29 new simproc functor
paulson
parents:
diff changeset
    14
sig
16569
a12992c34c12 cancels completely within terms as well now.
nipkow
parents: 16423
diff changeset
    15
  val cancel_ss       : simpset       (*abelian group cancel simpset*)
a12992c34c12 cancels completely within terms as well now.
nipkow
parents: 16423
diff changeset
    16
  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
    17
  val T               : typ           (*the type of group elements*)
a12992c34c12 cancels completely within terms as well now.
nipkow
parents: 16423
diff changeset
    18
  val eq_reflection   : thm           (*object-equality to meta-equality*)
a12992c34c12 cancels completely within terms as well now.
nipkow
parents: 16423
diff changeset
    19
  val eqI_rules       : thm list
a12992c34c12 cancels completely within terms as well now.
nipkow
parents: 16423
diff changeset
    20
  val dest_eqI        : thm -> term
5589
94c05305fb29 new simproc functor
paulson
parents:
diff changeset
    21
end;
94c05305fb29 new simproc functor
paulson
parents:
diff changeset
    22
94c05305fb29 new simproc functor
paulson
parents:
diff changeset
    23
94c05305fb29 new simproc functor
paulson
parents:
diff changeset
    24
functor Abel_Cancel (Data: ABEL_CANCEL) =
94c05305fb29 new simproc functor
paulson
parents:
diff changeset
    25
struct
94c05305fb29 new simproc functor
paulson
parents:
diff changeset
    26
5728
1d1175ef2d56 eliminated 'let ... in structure ...' to make SML/NJ 0.93 happy;
wenzelm
parents: 5610
diff changeset
    27
open Data;
1d1175ef2d56 eliminated 'let ... in structure ...' to make SML/NJ 0.93 happy;
wenzelm
parents: 5610
diff changeset
    28
20055
24c127b97ab5 simprocs: no theory argument -- use simpset context instead;
wenzelm
parents: 19798
diff changeset
    29
(* FIXME dependent on abstract syntax *)
24c127b97ab5 simprocs: no theory argument -- use simpset context instead;
wenzelm
parents: 19798
diff changeset
    30
22997
d4f3b015b50b canonical prefixing of class constants
haftmann
parents: 20713
diff changeset
    31
fun zero t = Const (@{const_name HOL.zero}, t);
d4f3b015b50b canonical prefixing of class constants
haftmann
parents: 20713
diff changeset
    32
fun minus t = Const (@{const_name HOL.uminus}, t --> t);
5589
94c05305fb29 new simproc functor
paulson
parents:
diff changeset
    33
22997
d4f3b015b50b canonical prefixing of class constants
haftmann
parents: 20713
diff changeset
    34
fun add_terms pos (Const (@{const_name HOL.plus}, _) $ x $ y, ts) =
20055
24c127b97ab5 simprocs: no theory argument -- use simpset context instead;
wenzelm
parents: 19798
diff changeset
    35
      add_terms pos (x, add_terms pos (y, ts))
22997
d4f3b015b50b canonical prefixing of class constants
haftmann
parents: 20713
diff changeset
    36
  | add_terms pos (Const (@{const_name HOL.minus}, _) $ x $ y, ts) =
20055
24c127b97ab5 simprocs: no theory argument -- use simpset context instead;
wenzelm
parents: 19798
diff changeset
    37
      add_terms pos (x, add_terms (not pos) (y, ts))
22997
d4f3b015b50b canonical prefixing of class constants
haftmann
parents: 20713
diff changeset
    38
  | add_terms pos (Const (@{const_name HOL.uminus}, _) $ x, ts) =
20055
24c127b97ab5 simprocs: no theory argument -- use simpset context instead;
wenzelm
parents: 19798
diff changeset
    39
      add_terms (not pos) (x, ts)
24c127b97ab5 simprocs: no theory argument -- use simpset context instead;
wenzelm
parents: 19798
diff changeset
    40
  | add_terms pos (x, ts) = (pos,x) :: ts;
5589
94c05305fb29 new simproc functor
paulson
parents:
diff changeset
    41
20055
24c127b97ab5 simprocs: no theory argument -- use simpset context instead;
wenzelm
parents: 19798
diff changeset
    42
fun terms fml = add_terms true (fml, []);
5589
94c05305fb29 new simproc functor
paulson
parents:
diff changeset
    43
22997
d4f3b015b50b canonical prefixing of class constants
haftmann
parents: 20713
diff changeset
    44
fun zero1 pt (u as (c as Const(@{const_name HOL.plus},_)) $ x $ y) =
16569
a12992c34c12 cancels completely within terms as well now.
nipkow
parents: 16423
diff changeset
    45
      (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
    46
                                   | SOME z => SOME(c $ x $ z))
a12992c34c12 cancels completely within terms as well now.
nipkow
parents: 16423
diff changeset
    47
       | SOME z => SOME(c $ z $ y))
22997
d4f3b015b50b canonical prefixing of class constants
haftmann
parents: 20713
diff changeset
    48
  | zero1 (pos,t) (u as (c as Const(@{const_name HOL.minus},_)) $ x $ y) =
16569
a12992c34c12 cancels completely within terms as well now.
nipkow
parents: 16423
diff changeset
    49
      (case zero1 (pos,t) x of
a12992c34c12 cancels completely within terms as well now.
nipkow
parents: 16423
diff changeset
    50
         NONE => (case zero1 (not pos,t) y of NONE => NONE
a12992c34c12 cancels completely within terms as well now.
nipkow
parents: 16423
diff changeset
    51
                  | SOME z => SOME(c $ x $ z))
a12992c34c12 cancels completely within terms as well now.
nipkow
parents: 16423
diff changeset
    52
       | SOME z => SOME(c $ z $ y))
22997
d4f3b015b50b canonical prefixing of class constants
haftmann
parents: 20713
diff changeset
    53
  | zero1 (pos,t) (u as (c as Const(@{const_name HOL.uminus},_)) $ x) =
16569
a12992c34c12 cancels completely within terms as well now.
nipkow
parents: 16423
diff changeset
    54
      (case zero1 (not pos,t) x of NONE => NONE
a12992c34c12 cancels completely within terms as well now.
nipkow
parents: 16423
diff changeset
    55
       | SOME z => SOME(c $ z))
a12992c34c12 cancels completely within terms as well now.
nipkow
parents: 16423
diff changeset
    56
  | zero1 (pos,t) u =
a12992c34c12 cancels completely within terms as well now.
nipkow
parents: 16423
diff changeset
    57
      if pos andalso (t aconv u) then SOME(zero(fastype_of t)) else NONE
5589
94c05305fb29 new simproc functor
paulson
parents:
diff changeset
    58
20055
24c127b97ab5 simprocs: no theory argument -- use simpset context instead;
wenzelm
parents: 19798
diff changeset
    59
exception Cancel;
5589
94c05305fb29 new simproc functor
paulson
parents:
diff changeset
    60
16569
a12992c34c12 cancels completely within terms as well now.
nipkow
parents: 16423
diff changeset
    61
fun find_common _ [] _ = raise Cancel
a12992c34c12 cancels completely within terms as well now.
nipkow
parents: 16423
diff changeset
    62
  | find_common opp ((p,l)::ls) rs =
a12992c34c12 cancels completely within terms as well now.
nipkow
parents: 16423
diff changeset
    63
      let val pr = if opp then not p else p
a12992c34c12 cancels completely within terms as well now.
nipkow
parents: 16423
diff changeset
    64
      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
    65
         else find_common opp ls rs
a12992c34c12 cancels completely within terms as well now.
nipkow
parents: 16423
diff changeset
    66
      end
5589
94c05305fb29 new simproc functor
paulson
parents:
diff changeset
    67
16569
a12992c34c12 cancels completely within terms as well now.
nipkow
parents: 16423
diff changeset
    68
(* 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
    69
   If OP = +, it must be t2(-t) rather than t2(t)
a12992c34c12 cancels completely within terms as well now.
nipkow
parents: 16423
diff changeset
    70
*)
20055
24c127b97ab5 simprocs: no theory argument -- use simpset context instead;
wenzelm
parents: 19798
diff changeset
    71
fun cancel t =
24c127b97ab5 simprocs: no theory argument -- use simpset context instead;
wenzelm
parents: 19798
diff changeset
    72
  let
24c127b97ab5 simprocs: no theory argument -- use simpset context instead;
wenzelm
parents: 19798
diff changeset
    73
    val c $ lhs $ rhs = t
22997
d4f3b015b50b canonical prefixing of class constants
haftmann
parents: 20713
diff changeset
    74
    val opp = case c of Const(@{const_name HOL.plus},_) => true | _ => false;
20055
24c127b97ab5 simprocs: no theory argument -- use simpset context instead;
wenzelm
parents: 19798
diff changeset
    75
    val (pos,l) = find_common opp (terms lhs) (terms rhs)
24c127b97ab5 simprocs: no theory argument -- use simpset context instead;
wenzelm
parents: 19798
diff changeset
    76
    val posr = if opp then not pos else pos
24c127b97ab5 simprocs: no theory argument -- use simpset context instead;
wenzelm
parents: 19798
diff changeset
    77
    val t' = c $ (the(zero1 (pos,l) lhs)) $ (the(zero1 (posr,l) rhs))
16569
a12992c34c12 cancels completely within terms as well now.
nipkow
parents: 16423
diff changeset
    78
  in t' end;
a12992c34c12 cancels completely within terms as well now.
nipkow
parents: 16423
diff changeset
    79
a12992c34c12 cancels completely within terms as well now.
nipkow
parents: 16423
diff changeset
    80
20055
24c127b97ab5 simprocs: no theory argument -- use simpset context instead;
wenzelm
parents: 19798
diff changeset
    81
(*A simproc to cancel complementary terms in arbitrary sums.*)
24c127b97ab5 simprocs: no theory argument -- use simpset context instead;
wenzelm
parents: 19798
diff changeset
    82
fun sum_proc ss t =
24c127b97ab5 simprocs: no theory argument -- use simpset context instead;
wenzelm
parents: 19798
diff changeset
    83
   let
24c127b97ab5 simprocs: no theory argument -- use simpset context instead;
wenzelm
parents: 19798
diff changeset
    84
     val t' = cancel t
24c127b97ab5 simprocs: no theory argument -- use simpset context instead;
wenzelm
parents: 19798
diff changeset
    85
     val thm =
24c127b97ab5 simprocs: no theory argument -- use simpset context instead;
wenzelm
parents: 19798
diff changeset
    86
       Goal.prove (Simplifier.the_context ss) [] [] (Logic.mk_equals (t, t'))
24c127b97ab5 simprocs: no theory argument -- use simpset context instead;
wenzelm
parents: 19798
diff changeset
    87
         (fn _ => simp_tac (Simplifier.inherit_context ss cancel_ss) 1)
24c127b97ab5 simprocs: no theory argument -- use simpset context instead;
wenzelm
parents: 19798
diff changeset
    88
   in SOME thm end handle Cancel => NONE;
5589
94c05305fb29 new simproc functor
paulson
parents:
diff changeset
    89
20055
24c127b97ab5 simprocs: no theory argument -- use simpset context instead;
wenzelm
parents: 19798
diff changeset
    90
val sum_conv =
24c127b97ab5 simprocs: no theory argument -- use simpset context instead;
wenzelm
parents: 19798
diff changeset
    91
  Simplifier.mk_simproc "cancel_sums"
24c127b97ab5 simprocs: no theory argument -- use simpset context instead;
wenzelm
parents: 19798
diff changeset
    92
    (map (Drule.cterm_fun Logic.varify o Thm.read_cterm (Theory.deref thy_ref))
24c127b97ab5 simprocs: no theory argument -- use simpset context instead;
wenzelm
parents: 19798
diff changeset
    93
      [("x + y", Data.T), ("x - y", Data.T)])  (* FIXME depends on concrete syntax *)
24c127b97ab5 simprocs: no theory argument -- use simpset context instead;
wenzelm
parents: 19798
diff changeset
    94
    (K sum_proc);
5589
94c05305fb29 new simproc functor
paulson
parents:
diff changeset
    95
94c05305fb29 new simproc functor
paulson
parents:
diff changeset
    96
20055
24c127b97ab5 simprocs: no theory argument -- use simpset context instead;
wenzelm
parents: 19798
diff changeset
    97
(*A simproc to cancel like terms on the opposite sides of relations:
24c127b97ab5 simprocs: no theory argument -- use simpset context instead;
wenzelm
parents: 19798
diff changeset
    98
   (x + y - z < -z + x) = (y < 0)
24c127b97ab5 simprocs: no theory argument -- use simpset context instead;
wenzelm
parents: 19798
diff changeset
    99
  Works for (=) and (<=) as well as (<), if the necessary rules are supplied.
24c127b97ab5 simprocs: no theory argument -- use simpset context instead;
wenzelm
parents: 19798
diff changeset
   100
  Reduces the problem to subtraction.*)
24c127b97ab5 simprocs: no theory argument -- use simpset context instead;
wenzelm
parents: 19798
diff changeset
   101
fun rel_proc ss t =
24c127b97ab5 simprocs: no theory argument -- use simpset context instead;
wenzelm
parents: 19798
diff changeset
   102
  let
24c127b97ab5 simprocs: no theory argument -- use simpset context instead;
wenzelm
parents: 19798
diff changeset
   103
    val t' = cancel t
24c127b97ab5 simprocs: no theory argument -- use simpset context instead;
wenzelm
parents: 19798
diff changeset
   104
    val thm = Goal.prove (Simplifier.the_context ss) [] [] (Logic.mk_equals (t, t'))
24c127b97ab5 simprocs: no theory argument -- use simpset context instead;
wenzelm
parents: 19798
diff changeset
   105
      (fn _ => rtac eq_reflection 1 THEN
24c127b97ab5 simprocs: no theory argument -- use simpset context instead;
wenzelm
parents: 19798
diff changeset
   106
                    resolve_tac eqI_rules 1 THEN
24c127b97ab5 simprocs: no theory argument -- use simpset context instead;
wenzelm
parents: 19798
diff changeset
   107
                    simp_tac (Simplifier.inherit_context ss cancel_ss) 1)
24c127b97ab5 simprocs: no theory argument -- use simpset context instead;
wenzelm
parents: 19798
diff changeset
   108
   in SOME thm end handle Cancel => NONE;
5589
94c05305fb29 new simproc functor
paulson
parents:
diff changeset
   109
20055
24c127b97ab5 simprocs: no theory argument -- use simpset context instead;
wenzelm
parents: 19798
diff changeset
   110
val rel_conv =
24c127b97ab5 simprocs: no theory argument -- use simpset context instead;
wenzelm
parents: 19798
diff changeset
   111
  Simplifier.mk_simproc "cancel_relations"
24c127b97ab5 simprocs: no theory argument -- use simpset context instead;
wenzelm
parents: 19798
diff changeset
   112
    (map (Thm.cterm_of (Theory.deref thy_ref) o Data.dest_eqI) eqI_rules) (K rel_proc);
5589
94c05305fb29 new simproc functor
paulson
parents:
diff changeset
   113
94c05305fb29 new simproc functor
paulson
parents:
diff changeset
   114
end;