| author | blanchet | 
| Mon, 23 Jul 2012 15:32:30 +0200 | |
| changeset 48442 | 3c9890c19e90 | 
| parent 45625 | 750c5a47400b | 
| permissions | -rw-r--r-- | 
| 37884 
314a88278715
discontinued pretending that abel_cancel is logic-independent; cleaned up junk
 haftmann parents: 
35845diff
changeset | 1 | (* Title: HOL/Tools/abel_cancel.ML | 
| 5589 | 2 | Author: Lawrence C Paulson, Cambridge University Computer Laboratory | 
| 3 | Copyright 1998 University of Cambridge | |
| 4 | ||
| 37884 
314a88278715
discontinued pretending that abel_cancel is logic-independent; cleaned up junk
 haftmann parents: 
35845diff
changeset | 5 | Simplification procedures for abelian groups: | 
| 
314a88278715
discontinued pretending that abel_cancel is logic-independent; cleaned up junk
 haftmann parents: 
35845diff
changeset | 6 | - Cancel complementary terms in sums. | 
| 
314a88278715
discontinued pretending that abel_cancel is logic-independent; cleaned up junk
 haftmann parents: 
35845diff
changeset | 7 | - Cancel like terms on opposite sides of relations. | 
| 5589 | 8 | *) | 
| 9 | ||
| 10 | signature ABEL_CANCEL = | |
| 11 | sig | |
| 37890 | 12 | val sum_proc: simpset -> cterm -> thm option | 
| 13 | val rel_proc: simpset -> cterm -> thm option | |
| 5589 | 14 | end; | 
| 15 | ||
| 37884 
314a88278715
discontinued pretending that abel_cancel is logic-independent; cleaned up junk
 haftmann parents: 
35845diff
changeset | 16 | structure Abel_Cancel: ABEL_CANCEL = | 
| 5589 | 17 | struct | 
| 18 | ||
| 37894 | 19 | (** compute cancellation **) | 
| 5728 
1d1175ef2d56
eliminated 'let ... in structure ...' to make SML/NJ 0.93 happy;
 wenzelm parents: 
5610diff
changeset | 20 | |
| 37894 | 21 | fun add_atoms pos (Const (@{const_name Groups.plus}, _) $ x $ y) =
 | 
| 22 | add_atoms pos x #> add_atoms pos y | |
| 23 |   | add_atoms pos (Const (@{const_name Groups.minus}, _) $ x $ y) =
 | |
| 24 | add_atoms pos x #> add_atoms (not pos) y | |
| 25 |   | add_atoms pos (Const (@{const_name Groups.uminus}, _) $ x) =
 | |
| 26 | add_atoms (not pos) x | |
| 45464 
5a5a6e6c6789
abel_cancel.ML: avoid canceling zeros, which would cause the simproc to return a trivial rewrite;
 huffman parents: 
38715diff
changeset | 27 |   | add_atoms pos (Const (@{const_name Groups.zero}, _)) = I
 | 
| 37894 | 28 | | add_atoms pos x = cons (pos, x); | 
| 37884 
314a88278715
discontinued pretending that abel_cancel is logic-independent; cleaned up junk
 haftmann parents: 
35845diff
changeset | 29 | |
| 37896 | 30 | fun atoms t = add_atoms true t []; | 
| 5589 | 31 | |
| 37896 | 32 | fun zerofy pt ((c as Const (@{const_name Groups.plus}, _)) $ x $ y) =
 | 
| 33 | (case zerofy pt x of NONE => | |
| 34 | (case zerofy pt y of NONE => NONE | |
| 35 | | SOME z => SOME (c $ x $ z)) | |
| 37894 | 36 | | SOME z => SOME (c $ z $ y)) | 
| 37896 | 37 |   | zerofy pt ((c as Const (@{const_name Groups.minus}, _)) $ x $ y) =
 | 
| 38 | (case zerofy pt x of NONE => | |
| 39 | (case zerofy (apfst not pt) y of NONE => NONE | |
| 40 | | SOME z => SOME (c $ x $ z)) | |
| 37894 | 41 | | SOME z => SOME (c $ z $ y)) | 
| 37896 | 42 |   | zerofy pt ((c as Const (@{const_name Groups.uminus}, _)) $ x) =
 | 
| 43 | (case zerofy (apfst not pt) x of NONE => NONE | |
| 37894 | 44 | | SOME z => SOME (c $ z)) | 
| 37896 | 45 | | zerofy (pos, t) u = | 
| 37894 | 46 | if pos andalso (t aconv u) | 
| 47 |         then SOME (Const (@{const_name Groups.zero}, fastype_of t))
 | |
| 48 | else NONE | |
| 5589 | 49 | |
| 20055 
24c127b97ab5
simprocs: no theory argument -- use simpset context instead;
 wenzelm parents: 
19798diff
changeset | 50 | exception Cancel; | 
| 5589 | 51 | |
| 16569 | 52 | fun find_common _ [] _ = raise Cancel | 
| 37894 | 53 | | find_common opp ((p, l) :: ls) rs = | 
| 16569 | 54 | let val pr = if opp then not p else p | 
| 37894 | 55 | in if exists (fn (q, r) => pr = q andalso l aconv r) rs then (p, l) | 
| 16569 | 56 | else find_common opp ls rs | 
| 57 | end | |
| 5589 | 58 | |
| 16569 | 59 | (* turns t1(t) OP t2(t) into t1(0) OP t2(0) where OP can be +, -, =, etc. | 
| 60 | If OP = +, it must be t2(-t) rather than t2(t) | |
| 61 | *) | |
| 37896 | 62 | fun cancel (c $ lhs $ rhs) = | 
| 20055 
24c127b97ab5
simprocs: no theory argument -- use simpset context instead;
 wenzelm parents: 
19798diff
changeset | 63 | let | 
| 37894 | 64 |     val opp = case c of Const(@{const_name Groups.plus}, _) => true | _ => false;
 | 
| 65 | val (pos, l) = find_common opp (atoms lhs) (atoms rhs); | |
| 66 | val posr = if opp then not pos else pos; | |
| 37896 | 67 | in c $ the (zerofy (pos, l) lhs) $ the (zerofy (posr, l) rhs) end; | 
| 16569 | 68 | |
| 69 | ||
| 37894 | 70 | (** prove cancellation **) | 
| 71 | ||
| 72 | fun agrp_ord (Const (a, _)) = find_index (fn a' => a = a') | |
| 73 |       [@{const_name Groups.zero}, @{const_name Groups.plus},
 | |
| 74 |         @{const_name Groups.uminus}, @{const_name Groups.minus}]
 | |
| 75 | | agrp_ord _ = ~1; | |
| 76 | ||
| 77 | fun less_agrp (a, b) = (Term_Ord.term_lpo agrp_ord (a, b) = LESS); | |
| 78 | ||
| 79 | fun solve (_ $ (Const (@{const_name Groups.plus}, _) $ _ $ _) $ _) =
 | |
| 80 |       SOME @{thm add_assoc [THEN eq_reflection]}
 | |
| 81 |   | solve (_ $ x $ (Const (@{const_name Groups.plus}, _) $ y $ _)) =
 | |
| 82 | if less_agrp (y, x) then | |
| 83 |         SOME @{thm add_left_commute [THEN eq_reflection]}
 | |
| 84 | else NONE | |
| 85 | | solve (_ $ x $ y) = | |
| 86 | if less_agrp (y, x) then | |
| 87 |         SOME @{thm add_commute [THEN eq_reflection]} else
 | |
| 88 | NONE | |
| 89 | | solve _ = NONE; | |
| 90 | ||
| 38715 
6513ea67d95d
renamed Simplifier.simproc(_i) to Simplifier.simproc_global(_i) to emphasize that this is not the real thing;
 wenzelm parents: 
37896diff
changeset | 91 | val simproc = Simplifier.simproc_global @{theory}
 | 
| 37894 | 92 | "add_ac_proc" ["x + y::'a::ab_semigroup_add"] ((K o K) solve); | 
| 93 | ||
| 45625 
750c5a47400b
modernized some old-style infix operations, which were left over from the time of ML proof scripts;
 wenzelm parents: 
45464diff
changeset | 94 | val cancel_ss = (HOL_basic_ss |> Simplifier.set_termless less_agrp) | 
| 37894 | 95 | addsimprocs [simproc] addsimps | 
| 96 |   [@{thm add_0_left}, @{thm add_0_right}, @{thm diff_minus},
 | |
| 97 |    @{thm minus_add_distrib}, @{thm minus_minus}, @{thm minus_zero},
 | |
| 98 |    @{thm right_minus}, @{thm left_minus}, @{thm add_minus_cancel},
 | |
| 99 |    @{thm minus_add_cancel}];
 | |
| 100 | ||
| 101 | fun cancel_simp_tac ss = simp_tac (Simplifier.inherit_context ss cancel_ss); | |
| 102 | ||
| 103 | ||
| 104 | (** simprocs **) | |
| 105 | ||
| 106 | (* cancel complementary terms in arbitrary sums *) | |
| 107 | ||
| 37890 | 108 | fun sum_proc ss ct = | 
| 109 | let | |
| 110 | val t = Thm.term_of ct; | |
| 37896 | 111 | val prop = Logic.mk_equals (t, cancel t); | 
| 112 | val thm = Goal.prove (Simplifier.the_context ss) [] [] prop | |
| 113 | (fn _ => cancel_simp_tac ss 1) | |
| 37890 | 114 | in SOME thm end handle Cancel => NONE; | 
| 5589 | 115 | |
| 116 | ||
| 37894 | 117 | (* cancel like terms on the opposite sides of relations: | 
| 118 | (x + y - z < -z + x) = (y < 0) | |
| 119 | Works for (=) and (<=) as well as (<), if the necessary rules are supplied. | |
| 120 | Reduces the problem to subtraction. *) | |
| 121 | ||
| 37890 | 122 | fun rel_proc ss ct = | 
| 20055 
24c127b97ab5
simprocs: no theory argument -- use simpset context instead;
 wenzelm parents: 
19798diff
changeset | 123 | let | 
| 37890 | 124 | val t = Thm.term_of ct; | 
| 37896 | 125 | val prop = Logic.mk_equals (t, cancel t); | 
| 126 | val thm = Goal.prove (Simplifier.the_context ss) [] [] prop | |
| 37884 
314a88278715
discontinued pretending that abel_cancel is logic-independent; cleaned up junk
 haftmann parents: 
35845diff
changeset | 127 |       (fn _ => rtac @{thm eq_reflection} 1 THEN
 | 
| 37896 | 128 |         resolve_tac [@{thm diff_eq_diff_less}, @{thm diff_eq_diff_less_eq}, @{thm diff_eq_diff_eq}] 1 THEN
 | 
| 129 | cancel_simp_tac ss 1) | |
| 37890 | 130 | in SOME thm end handle Cancel => NONE; | 
| 5589 | 131 | |
| 132 | end; |