src/Provers/Arith/abel_cancel.ML
author berghofe
Thu, 21 Apr 2005 19:02:54 +0200
changeset 15796 348ce23d2fc2
parent 15570 8d8c70b41bab
child 16423 24abe4c0e4b4
permissions -rw-r--r--
Moved cterm_fun from Thm to Drule.
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
5589
94c05305fb29 new simproc functor
paulson
parents:
diff changeset
    11
*)
94c05305fb29 new simproc functor
paulson
parents:
diff changeset
    12
14756
9c8cc63714f4 modified abel_cancel.ML for polymorphic types
obua
parents: 13480
diff changeset
    13
(* Modificaton in May 2004 by Steven Obua: polymorphic types work also now *) 
5589
94c05305fb29 new simproc functor
paulson
parents:
diff changeset
    14
94c05305fb29 new simproc functor
paulson
parents:
diff changeset
    15
signature ABEL_CANCEL =
94c05305fb29 new simproc functor
paulson
parents:
diff changeset
    16
sig
15104
ballarin
parents: 15027
diff changeset
    17
  val ss                : simpset       (*basic simpset of object-logic*)
13462
56610e2ba220 sane interface for simprocs;
wenzelm
parents: 12262
diff changeset
    18
  val eq_reflection     : thm           (*object-equality to meta-equality*)
5589
94c05305fb29 new simproc functor
paulson
parents:
diff changeset
    19
13462
56610e2ba220 sane interface for simprocs;
wenzelm
parents: 12262
diff changeset
    20
  val sg_ref            : Sign.sg_ref   (*signature of the theory of the group*)
56610e2ba220 sane interface for simprocs;
wenzelm
parents: 12262
diff changeset
    21
  val T                 : typ           (*the type of group elements*)
5589
94c05305fb29 new simproc functor
paulson
parents:
diff changeset
    22
7586
ae28545cd104 The restrict_to_left rule fixes some bugs
paulson
parents: 6391
diff changeset
    23
  val restrict_to_left  : thm
13462
56610e2ba220 sane interface for simprocs;
wenzelm
parents: 12262
diff changeset
    24
  val add_cancel_21     : thm
56610e2ba220 sane interface for simprocs;
wenzelm
parents: 12262
diff changeset
    25
  val add_cancel_end    : thm
56610e2ba220 sane interface for simprocs;
wenzelm
parents: 12262
diff changeset
    26
  val add_left_cancel   : thm
56610e2ba220 sane interface for simprocs;
wenzelm
parents: 12262
diff changeset
    27
  val add_assoc         : thm
56610e2ba220 sane interface for simprocs;
wenzelm
parents: 12262
diff changeset
    28
  val add_commute       : thm
56610e2ba220 sane interface for simprocs;
wenzelm
parents: 12262
diff changeset
    29
  val add_left_commute  : thm
56610e2ba220 sane interface for simprocs;
wenzelm
parents: 12262
diff changeset
    30
  val add_0             : thm
56610e2ba220 sane interface for simprocs;
wenzelm
parents: 12262
diff changeset
    31
  val add_0_right       : thm
5589
94c05305fb29 new simproc functor
paulson
parents:
diff changeset
    32
13462
56610e2ba220 sane interface for simprocs;
wenzelm
parents: 12262
diff changeset
    33
  val eq_diff_eq        : thm
56610e2ba220 sane interface for simprocs;
wenzelm
parents: 12262
diff changeset
    34
  val eqI_rules         : thm list
56610e2ba220 sane interface for simprocs;
wenzelm
parents: 12262
diff changeset
    35
  val dest_eqI          : thm -> term
5589
94c05305fb29 new simproc functor
paulson
parents:
diff changeset
    36
13462
56610e2ba220 sane interface for simprocs;
wenzelm
parents: 12262
diff changeset
    37
  val diff_def          : thm
56610e2ba220 sane interface for simprocs;
wenzelm
parents: 12262
diff changeset
    38
  val minus_add_distrib : thm
56610e2ba220 sane interface for simprocs;
wenzelm
parents: 12262
diff changeset
    39
  val minus_minus       : thm
56610e2ba220 sane interface for simprocs;
wenzelm
parents: 12262
diff changeset
    40
  val minus_0           : thm
5589
94c05305fb29 new simproc functor
paulson
parents:
diff changeset
    41
13462
56610e2ba220 sane interface for simprocs;
wenzelm
parents: 12262
diff changeset
    42
  val add_inverses      : thm list
56610e2ba220 sane interface for simprocs;
wenzelm
parents: 12262
diff changeset
    43
  val cancel_simps      : thm list
5589
94c05305fb29 new simproc functor
paulson
parents:
diff changeset
    44
end;
94c05305fb29 new simproc functor
paulson
parents:
diff changeset
    45
94c05305fb29 new simproc functor
paulson
parents:
diff changeset
    46
94c05305fb29 new simproc functor
paulson
parents:
diff changeset
    47
functor Abel_Cancel (Data: ABEL_CANCEL) =
94c05305fb29 new simproc functor
paulson
parents:
diff changeset
    48
struct
94c05305fb29 new simproc functor
paulson
parents:
diff changeset
    49
5728
1d1175ef2d56 eliminated 'let ... in structure ...' to make SML/NJ 0.93 happy;
wenzelm
parents: 5610
diff changeset
    50
open Data;
1d1175ef2d56 eliminated 'let ... in structure ...' to make SML/NJ 0.93 happy;
wenzelm
parents: 5610
diff changeset
    51
13462
56610e2ba220 sane interface for simprocs;
wenzelm
parents: 12262
diff changeset
    52
 val prepare_ss = Data.ss addsimps [add_assoc, diff_def,
56610e2ba220 sane interface for simprocs;
wenzelm
parents: 12262
diff changeset
    53
                                    minus_add_distrib, minus_minus,
56610e2ba220 sane interface for simprocs;
wenzelm
parents: 12262
diff changeset
    54
                                    minus_0, add_0, add_0_right];
5589
94c05305fb29 new simproc functor
paulson
parents:
diff changeset
    55
14756
9c8cc63714f4 modified abel_cancel.ML for polymorphic types
obua
parents: 13480
diff changeset
    56
 
9c8cc63714f4 modified abel_cancel.ML for polymorphic types
obua
parents: 13480
diff changeset
    57
 fun zero t = Const ("0", t);
9c8cc63714f4 modified abel_cancel.ML for polymorphic types
obua
parents: 13480
diff changeset
    58
 fun plus t = Const ("op +", [t,t] ---> t);
9c8cc63714f4 modified abel_cancel.ML for polymorphic types
obua
parents: 13480
diff changeset
    59
 fun minus t = Const ("uminus", t --> t);
5589
94c05305fb29 new simproc functor
paulson
parents:
diff changeset
    60
94c05305fb29 new simproc functor
paulson
parents:
diff changeset
    61
 (*Cancel corresponding terms on the two sides of the equation, NOT on
94c05305fb29 new simproc functor
paulson
parents:
diff changeset
    62
   the same side!*)
13462
56610e2ba220 sane interface for simprocs;
wenzelm
parents: 12262
diff changeset
    63
 val cancel_ss =
56610e2ba220 sane interface for simprocs;
wenzelm
parents: 12262
diff changeset
    64
   Data.ss addsimps [add_cancel_21, add_cancel_end, minus_minus] @
7586
ae28545cd104 The restrict_to_left rule fixes some bugs
paulson
parents: 6391
diff changeset
    65
                    (map (fn th => th RS restrict_to_left) Data.cancel_simps);
5589
94c05305fb29 new simproc functor
paulson
parents:
diff changeset
    66
7586
ae28545cd104 The restrict_to_left rule fixes some bugs
paulson
parents: 6391
diff changeset
    67
 val inverse_ss = Data.ss addsimps Data.add_inverses @ Data.cancel_simps;
5589
94c05305fb29 new simproc functor
paulson
parents:
diff changeset
    68
14756
9c8cc63714f4 modified abel_cancel.ML for polymorphic types
obua
parents: 13480
diff changeset
    69
 fun mk_sum ty []  = zero ty
9c8cc63714f4 modified abel_cancel.ML for polymorphic types
obua
parents: 13480
diff changeset
    70
   | mk_sum ty tms = foldr1 (fn (x,y) => (plus ty) $ x $ y) tms;
5589
94c05305fb29 new simproc functor
paulson
parents:
diff changeset
    71
94c05305fb29 new simproc functor
paulson
parents:
diff changeset
    72
 (*We map -t to t and (in other cases) t to -t.  No need to check the type of
94c05305fb29 new simproc functor
paulson
parents:
diff changeset
    73
   uminus, since the simproc is only called on sums of type T.*)
94c05305fb29 new simproc functor
paulson
parents:
diff changeset
    74
 fun negate (Const("uminus",_) $ t) = t
14756
9c8cc63714f4 modified abel_cancel.ML for polymorphic types
obua
parents: 13480
diff changeset
    75
   | negate t                       = (minus (fastype_of t)) $ t;
5589
94c05305fb29 new simproc functor
paulson
parents:
diff changeset
    76
94c05305fb29 new simproc functor
paulson
parents:
diff changeset
    77
 (*Flatten a formula built from +, binary - and unary -.
94c05305fb29 new simproc functor
paulson
parents:
diff changeset
    78
   No need to check types PROVIDED they are checked upon entry!*)
94c05305fb29 new simproc functor
paulson
parents:
diff changeset
    79
 fun add_terms neg (Const ("op +", _) $ x $ y, ts) =
13462
56610e2ba220 sane interface for simprocs;
wenzelm
parents: 12262
diff changeset
    80
         add_terms neg (x, add_terms neg (y, ts))
5589
94c05305fb29 new simproc functor
paulson
parents:
diff changeset
    81
   | add_terms neg (Const ("op -", _) $ x $ y, ts) =
13462
56610e2ba220 sane interface for simprocs;
wenzelm
parents: 12262
diff changeset
    82
         add_terms neg (x, add_terms (not neg) (y, ts))
56610e2ba220 sane interface for simprocs;
wenzelm
parents: 12262
diff changeset
    83
   | add_terms neg (Const ("uminus", _) $ x, ts) =
56610e2ba220 sane interface for simprocs;
wenzelm
parents: 12262
diff changeset
    84
         add_terms (not neg) (x, ts)
56610e2ba220 sane interface for simprocs;
wenzelm
parents: 12262
diff changeset
    85
   | add_terms neg (x, ts) =
56610e2ba220 sane interface for simprocs;
wenzelm
parents: 12262
diff changeset
    86
         (if neg then negate x else x) :: ts;
5589
94c05305fb29 new simproc functor
paulson
parents:
diff changeset
    87
94c05305fb29 new simproc functor
paulson
parents:
diff changeset
    88
 fun terms fml = add_terms false (fml, []);
94c05305fb29 new simproc functor
paulson
parents:
diff changeset
    89
94c05305fb29 new simproc functor
paulson
parents:
diff changeset
    90
 exception Cancel;
94c05305fb29 new simproc functor
paulson
parents:
diff changeset
    91
94c05305fb29 new simproc functor
paulson
parents:
diff changeset
    92
 (*Cancels just the first occurrence of u, leaving the rest unchanged*)
94c05305fb29 new simproc functor
paulson
parents:
diff changeset
    93
 fun cancelled (u, t::ts) = if u aconv t then ts else t :: cancelled(u,ts)
94c05305fb29 new simproc functor
paulson
parents:
diff changeset
    94
   | cancelled _          = raise Cancel;
94c05305fb29 new simproc functor
paulson
parents:
diff changeset
    95
94c05305fb29 new simproc functor
paulson
parents:
diff changeset
    96
94c05305fb29 new simproc functor
paulson
parents:
diff changeset
    97
 val trace = ref false;
94c05305fb29 new simproc functor
paulson
parents:
diff changeset
    98
94c05305fb29 new simproc functor
paulson
parents:
diff changeset
    99
 (*Make a simproc to cancel complementary terms in sums.  Examples:
94c05305fb29 new simproc functor
paulson
parents:
diff changeset
   100
    x-x = 0    x+(y-x) = y   -x+(y+(x+z)) = y+z
13462
56610e2ba220 sane interface for simprocs;
wenzelm
parents: 12262
diff changeset
   101
   It will unfold the definition of diff and associate to the right if
5589
94c05305fb29 new simproc functor
paulson
parents:
diff changeset
   102
   necessary.  Rewriting is faster if the formula is already
94c05305fb29 new simproc functor
paulson
parents:
diff changeset
   103
   in that form.
94c05305fb29 new simproc functor
paulson
parents:
diff changeset
   104
 *)
94c05305fb29 new simproc functor
paulson
parents:
diff changeset
   105
94c05305fb29 new simproc functor
paulson
parents:
diff changeset
   106
 fun sum_proc sg _ lhs =
13462
56610e2ba220 sane interface for simprocs;
wenzelm
parents: 12262
diff changeset
   107
   let val _ = if !trace then tracing ("cancel_sums: LHS = " ^
56610e2ba220 sane interface for simprocs;
wenzelm
parents: 12262
diff changeset
   108
                                       string_of_cterm (cterm_of sg lhs))
56610e2ba220 sane interface for simprocs;
wenzelm
parents: 12262
diff changeset
   109
               else ()
5589
94c05305fb29 new simproc functor
paulson
parents:
diff changeset
   110
       val (head::tail) = terms lhs
94c05305fb29 new simproc functor
paulson
parents:
diff changeset
   111
       val head' = negate head
14756
9c8cc63714f4 modified abel_cancel.ML for polymorphic types
obua
parents: 13480
diff changeset
   112
       val rhs = mk_sum (fastype_of head) (cancelled (head',tail))
5589
94c05305fb29 new simproc functor
paulson
parents:
diff changeset
   113
       and chead' = Thm.cterm_of sg head'
13462
56610e2ba220 sane interface for simprocs;
wenzelm
parents: 12262
diff changeset
   114
       val _ = if !trace then
56610e2ba220 sane interface for simprocs;
wenzelm
parents: 12262
diff changeset
   115
                 tracing ("RHS = " ^ string_of_cterm (Thm.cterm_of sg rhs))
56610e2ba220 sane interface for simprocs;
wenzelm
parents: 12262
diff changeset
   116
               else ()
13480
bb72bd43c6c3 use Tactic.prove instead of prove_goalw_cterm in internal proofs!
wenzelm
parents: 13462
diff changeset
   117
       val thm = Tactic.prove sg [] [] (Logic.mk_equals (lhs, rhs))
bb72bd43c6c3 use Tactic.prove instead of prove_goalw_cterm in internal proofs!
wenzelm
parents: 13462
diff changeset
   118
                   (fn _ => rtac eq_reflection 1 THEN
bb72bd43c6c3 use Tactic.prove instead of prove_goalw_cterm in internal proofs!
wenzelm
parents: 13462
diff changeset
   119
                            simp_tac prepare_ss 1 THEN
bb72bd43c6c3 use Tactic.prove instead of prove_goalw_cterm in internal proofs!
wenzelm
parents: 13462
diff changeset
   120
                            IF_UNSOLVED (simp_tac cancel_ss 1) THEN
bb72bd43c6c3 use Tactic.prove instead of prove_goalw_cterm in internal proofs!
wenzelm
parents: 13462
diff changeset
   121
                            IF_UNSOLVED (simp_tac inverse_ss 1))
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15183
diff changeset
   122
   in SOME thm end
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15183
diff changeset
   123
   handle Cancel => NONE;
5589
94c05305fb29 new simproc functor
paulson
parents:
diff changeset
   124
94c05305fb29 new simproc functor
paulson
parents:
diff changeset
   125
13462
56610e2ba220 sane interface for simprocs;
wenzelm
parents: 12262
diff changeset
   126
 val sum_conv =
5589
94c05305fb29 new simproc functor
paulson
parents:
diff changeset
   127
     Simplifier.mk_simproc "cancel_sums"
15796
348ce23d2fc2 Moved cterm_fun from Thm to Drule.
berghofe
parents: 15570
diff changeset
   128
       (map (Drule.cterm_fun Logic.varify o Thm.read_cterm (Sign.deref sg_ref))
13462
56610e2ba220 sane interface for simprocs;
wenzelm
parents: 12262
diff changeset
   129
        [("x + y", Data.T), ("x - y", Data.T)])  (* FIXME depends on concrete syntax !???!!??! *)
5589
94c05305fb29 new simproc functor
paulson
parents:
diff changeset
   130
       sum_proc;
94c05305fb29 new simproc functor
paulson
parents:
diff changeset
   131
94c05305fb29 new simproc functor
paulson
parents:
diff changeset
   132
94c05305fb29 new simproc functor
paulson
parents:
diff changeset
   133
 (*A simproc to cancel like terms on the opposite sides of relations:
94c05305fb29 new simproc functor
paulson
parents:
diff changeset
   134
     (x + y - z < -z + x) = (y < 0)
94c05305fb29 new simproc functor
paulson
parents:
diff changeset
   135
   Works for (=) and (<=) as well as (<), if the necessary rules are supplied.
94c05305fb29 new simproc functor
paulson
parents:
diff changeset
   136
   Reduces the problem to subtraction and calls the previous simproc.
94c05305fb29 new simproc functor
paulson
parents:
diff changeset
   137
 *)
94c05305fb29 new simproc functor
paulson
parents:
diff changeset
   138
94c05305fb29 new simproc functor
paulson
parents:
diff changeset
   139
 (*Cancel the FIRST occurrence of a term.  If it's repeated, then repeated
94c05305fb29 new simproc functor
paulson
parents:
diff changeset
   140
   calls to the simproc will be needed.*)
94c05305fb29 new simproc functor
paulson
parents:
diff changeset
   141
 fun cancel1 ([], u)    = raise Match (*impossible: it's a common term*)
94c05305fb29 new simproc functor
paulson
parents:
diff changeset
   142
   | cancel1 (t::ts, u) = if t aconv u then ts
13462
56610e2ba220 sane interface for simprocs;
wenzelm
parents: 12262
diff changeset
   143
                          else t :: cancel1 (ts,u);
5589
94c05305fb29 new simproc functor
paulson
parents:
diff changeset
   144
94c05305fb29 new simproc functor
paulson
parents:
diff changeset
   145
94c05305fb29 new simproc functor
paulson
parents:
diff changeset
   146
 val sum_cancel_ss = Data.ss addsimprocs [sum_conv]
13462
56610e2ba220 sane interface for simprocs;
wenzelm
parents: 12262
diff changeset
   147
                             addsimps    [add_0, add_0_right];
5589
94c05305fb29 new simproc functor
paulson
parents:
diff changeset
   148
94c05305fb29 new simproc functor
paulson
parents:
diff changeset
   149
 val add_ac_ss = Data.ss addsimps [add_assoc,add_commute,add_left_commute];
94c05305fb29 new simproc functor
paulson
parents:
diff changeset
   150
94c05305fb29 new simproc functor
paulson
parents:
diff changeset
   151
 fun rel_proc sg _ (lhs as (rel$lt$rt)) =
13462
56610e2ba220 sane interface for simprocs;
wenzelm
parents: 12262
diff changeset
   152
   let val _ = if !trace then tracing ("cancel_relations: LHS = " ^
56610e2ba220 sane interface for simprocs;
wenzelm
parents: 12262
diff changeset
   153
                                       string_of_cterm (cterm_of sg lhs))
56610e2ba220 sane interface for simprocs;
wenzelm
parents: 12262
diff changeset
   154
               else ()
5589
94c05305fb29 new simproc functor
paulson
parents:
diff changeset
   155
       val ltms = terms lt
94c05305fb29 new simproc functor
paulson
parents:
diff changeset
   156
       and rtms = terms rt
14756
9c8cc63714f4 modified abel_cancel.ML for polymorphic types
obua
parents: 13480
diff changeset
   157
       val ty = fastype_of lt 
5589
94c05305fb29 new simproc functor
paulson
parents:
diff changeset
   158
       val common = (*inter_term miscounts repetitions, so squash them*)
13462
56610e2ba220 sane interface for simprocs;
wenzelm
parents: 12262
diff changeset
   159
                    gen_distinct (op aconv) (inter_term (ltms, rtms))
5589
94c05305fb29 new simproc functor
paulson
parents:
diff changeset
   160
       val _ = if null common then raise Cancel  (*nothing to do*)
13462
56610e2ba220 sane interface for simprocs;
wenzelm
parents: 12262
diff changeset
   161
                                   else ()
5589
94c05305fb29 new simproc functor
paulson
parents:
diff changeset
   162
15570
8d8c70b41bab Move towards standard functions.
skalberg
parents: 15531
diff changeset
   163
       fun cancelled tms = mk_sum ty (Library.foldl cancel1 (tms, common))
5589
94c05305fb29 new simproc functor
paulson
parents:
diff changeset
   164
94c05305fb29 new simproc functor
paulson
parents:
diff changeset
   165
       val lt' = cancelled ltms
94c05305fb29 new simproc functor
paulson
parents:
diff changeset
   166
       and rt' = cancelled rtms
94c05305fb29 new simproc functor
paulson
parents:
diff changeset
   167
94c05305fb29 new simproc functor
paulson
parents:
diff changeset
   168
       val rhs = rel$lt'$rt'
15183
66da80cad4a2 now rejects degenerate (looping) cases
paulson
parents: 15104
diff changeset
   169
       val _ = if lhs aconv rhs then raise Cancel  (*nothing to do*)
66da80cad4a2 now rejects degenerate (looping) cases
paulson
parents: 15104
diff changeset
   170
                                     else ()
13462
56610e2ba220 sane interface for simprocs;
wenzelm
parents: 12262
diff changeset
   171
       val _ = if !trace then
56610e2ba220 sane interface for simprocs;
wenzelm
parents: 12262
diff changeset
   172
                 tracing ("RHS = " ^ string_of_cterm (Thm.cterm_of sg rhs))
56610e2ba220 sane interface for simprocs;
wenzelm
parents: 12262
diff changeset
   173
               else ()
5589
94c05305fb29 new simproc functor
paulson
parents:
diff changeset
   174
13480
bb72bd43c6c3 use Tactic.prove instead of prove_goalw_cterm in internal proofs!
wenzelm
parents: 13462
diff changeset
   175
       val thm = Tactic.prove sg [] [] (Logic.mk_equals (lhs, rhs))
bb72bd43c6c3 use Tactic.prove instead of prove_goalw_cterm in internal proofs!
wenzelm
parents: 13462
diff changeset
   176
                   (fn _ => rtac eq_reflection 1 THEN
bb72bd43c6c3 use Tactic.prove instead of prove_goalw_cterm in internal proofs!
wenzelm
parents: 13462
diff changeset
   177
                            resolve_tac eqI_rules 1 THEN
bb72bd43c6c3 use Tactic.prove instead of prove_goalw_cterm in internal proofs!
wenzelm
parents: 13462
diff changeset
   178
                            simp_tac prepare_ss 1 THEN
bb72bd43c6c3 use Tactic.prove instead of prove_goalw_cterm in internal proofs!
wenzelm
parents: 13462
diff changeset
   179
                            simp_tac sum_cancel_ss 1 THEN
bb72bd43c6c3 use Tactic.prove instead of prove_goalw_cterm in internal proofs!
wenzelm
parents: 13462
diff changeset
   180
                            IF_UNSOLVED (simp_tac add_ac_ss 1))
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15183
diff changeset
   181
   in SOME thm end
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15183
diff changeset
   182
   handle Cancel => NONE;
5589
94c05305fb29 new simproc functor
paulson
parents:
diff changeset
   183
13462
56610e2ba220 sane interface for simprocs;
wenzelm
parents: 12262
diff changeset
   184
 val rel_conv =
56610e2ba220 sane interface for simprocs;
wenzelm
parents: 12262
diff changeset
   185
     Simplifier.simproc_i (Sign.deref sg_ref) "cancel_relations"
56610e2ba220 sane interface for simprocs;
wenzelm
parents: 12262
diff changeset
   186
       (map Data.dest_eqI eqI_rules) rel_proc;
5589
94c05305fb29 new simproc functor
paulson
parents:
diff changeset
   187
94c05305fb29 new simproc functor
paulson
parents:
diff changeset
   188
end;
14756
9c8cc63714f4 modified abel_cancel.ML for polymorphic types
obua
parents: 13480
diff changeset
   189
9c8cc63714f4 modified abel_cancel.ML for polymorphic types
obua
parents: 13480
diff changeset
   190