author | wenzelm |
Sat, 16 Apr 2011 18:11:20 +0200 | |
changeset 42364 | 8c674b3b8e44 |
parent 38715 | 6513ea67d95d |
child 45464 | 5a5a6e6c6789 |
permissions | -rw-r--r-- |
37884
314a88278715
discontinued pretending that abel_cancel is logic-independent; cleaned up junk
haftmann
parents:
35845
diff
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:
35845
diff
changeset
|
5 |
Simplification procedures for abelian groups: |
314a88278715
discontinued pretending that abel_cancel is logic-independent; cleaned up junk
haftmann
parents:
35845
diff
changeset
|
6 |
- Cancel complementary terms in sums. |
314a88278715
discontinued pretending that abel_cancel is logic-independent; cleaned up junk
haftmann
parents:
35845
diff
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:
35845
diff
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:
5610
diff
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 |
|
27 |
| add_atoms pos x = cons (pos, x); |
|
37884
314a88278715
discontinued pretending that abel_cancel is logic-independent; cleaned up junk
haftmann
parents:
35845
diff
changeset
|
28 |
|
37896 | 29 |
fun atoms t = add_atoms true t []; |
5589 | 30 |
|
37896 | 31 |
fun zerofy pt ((c as Const (@{const_name Groups.plus}, _)) $ x $ y) = |
32 |
(case zerofy pt x of NONE => |
|
33 |
(case zerofy pt y of NONE => NONE |
|
34 |
| SOME z => SOME (c $ x $ z)) |
|
37894 | 35 |
| SOME z => SOME (c $ z $ y)) |
37896 | 36 |
| zerofy pt ((c as Const (@{const_name Groups.minus}, _)) $ x $ y) = |
37 |
(case zerofy pt x of NONE => |
|
38 |
(case zerofy (apfst not pt) y of NONE => NONE |
|
39 |
| SOME z => SOME (c $ x $ z)) |
|
37894 | 40 |
| SOME z => SOME (c $ z $ y)) |
37896 | 41 |
| zerofy pt ((c as Const (@{const_name Groups.uminus}, _)) $ x) = |
42 |
(case zerofy (apfst not pt) x of NONE => NONE |
|
37894 | 43 |
| SOME z => SOME (c $ z)) |
37896 | 44 |
| zerofy (pos, t) u = |
37894 | 45 |
if pos andalso (t aconv u) |
46 |
then SOME (Const (@{const_name Groups.zero}, fastype_of t)) |
|
47 |
else NONE |
|
5589 | 48 |
|
20055
24c127b97ab5
simprocs: no theory argument -- use simpset context instead;
wenzelm
parents:
19798
diff
changeset
|
49 |
exception Cancel; |
5589 | 50 |
|
16569 | 51 |
fun find_common _ [] _ = raise Cancel |
37894 | 52 |
| find_common opp ((p, l) :: ls) rs = |
16569 | 53 |
let val pr = if opp then not p else p |
37894 | 54 |
in if exists (fn (q, r) => pr = q andalso l aconv r) rs then (p, l) |
16569 | 55 |
else find_common opp ls rs |
56 |
end |
|
5589 | 57 |
|
16569 | 58 |
(* turns t1(t) OP t2(t) into t1(0) OP t2(0) where OP can be +, -, =, etc. |
59 |
If OP = +, it must be t2(-t) rather than t2(t) |
|
60 |
*) |
|
37896 | 61 |
fun cancel (c $ lhs $ rhs) = |
20055
24c127b97ab5
simprocs: no theory argument -- use simpset context instead;
wenzelm
parents:
19798
diff
changeset
|
62 |
let |
37894 | 63 |
val opp = case c of Const(@{const_name Groups.plus}, _) => true | _ => false; |
64 |
val (pos, l) = find_common opp (atoms lhs) (atoms rhs); |
|
65 |
val posr = if opp then not pos else pos; |
|
37896 | 66 |
in c $ the (zerofy (pos, l) lhs) $ the (zerofy (posr, l) rhs) end; |
16569 | 67 |
|
68 |
||
37894 | 69 |
(** prove cancellation **) |
70 |
||
71 |
fun agrp_ord (Const (a, _)) = find_index (fn a' => a = a') |
|
72 |
[@{const_name Groups.zero}, @{const_name Groups.plus}, |
|
73 |
@{const_name Groups.uminus}, @{const_name Groups.minus}] |
|
74 |
| agrp_ord _ = ~1; |
|
75 |
||
76 |
fun less_agrp (a, b) = (Term_Ord.term_lpo agrp_ord (a, b) = LESS); |
|
77 |
||
78 |
fun solve (_ $ (Const (@{const_name Groups.plus}, _) $ _ $ _) $ _) = |
|
79 |
SOME @{thm add_assoc [THEN eq_reflection]} |
|
80 |
| solve (_ $ x $ (Const (@{const_name Groups.plus}, _) $ y $ _)) = |
|
81 |
if less_agrp (y, x) then |
|
82 |
SOME @{thm add_left_commute [THEN eq_reflection]} |
|
83 |
else NONE |
|
84 |
| solve (_ $ x $ y) = |
|
85 |
if less_agrp (y, x) then |
|
86 |
SOME @{thm add_commute [THEN eq_reflection]} else |
|
87 |
NONE |
|
88 |
| solve _ = NONE; |
|
89 |
||
38715
6513ea67d95d
renamed Simplifier.simproc(_i) to Simplifier.simproc_global(_i) to emphasize that this is not the real thing;
wenzelm
parents:
37896
diff
changeset
|
90 |
val simproc = Simplifier.simproc_global @{theory} |
37894 | 91 |
"add_ac_proc" ["x + y::'a::ab_semigroup_add"] ((K o K) solve); |
92 |
||
93 |
val cancel_ss = HOL_basic_ss settermless less_agrp |
|
94 |
addsimprocs [simproc] addsimps |
|
95 |
[@{thm add_0_left}, @{thm add_0_right}, @{thm diff_minus}, |
|
96 |
@{thm minus_add_distrib}, @{thm minus_minus}, @{thm minus_zero}, |
|
97 |
@{thm right_minus}, @{thm left_minus}, @{thm add_minus_cancel}, |
|
98 |
@{thm minus_add_cancel}]; |
|
99 |
||
100 |
fun cancel_simp_tac ss = simp_tac (Simplifier.inherit_context ss cancel_ss); |
|
101 |
||
102 |
||
103 |
(** simprocs **) |
|
104 |
||
105 |
(* cancel complementary terms in arbitrary sums *) |
|
106 |
||
37890 | 107 |
fun sum_proc ss ct = |
108 |
let |
|
109 |
val t = Thm.term_of ct; |
|
37896 | 110 |
val prop = Logic.mk_equals (t, cancel t); |
111 |
val thm = Goal.prove (Simplifier.the_context ss) [] [] prop |
|
112 |
(fn _ => cancel_simp_tac ss 1) |
|
37890 | 113 |
in SOME thm end handle Cancel => NONE; |
5589 | 114 |
|
115 |
||
37894 | 116 |
(* cancel like terms on the opposite sides of relations: |
117 |
(x + y - z < -z + x) = (y < 0) |
|
118 |
Works for (=) and (<=) as well as (<), if the necessary rules are supplied. |
|
119 |
Reduces the problem to subtraction. *) |
|
120 |
||
37890 | 121 |
fun rel_proc ss ct = |
20055
24c127b97ab5
simprocs: no theory argument -- use simpset context instead;
wenzelm
parents:
19798
diff
changeset
|
122 |
let |
37890 | 123 |
val t = Thm.term_of ct; |
37896 | 124 |
val prop = Logic.mk_equals (t, cancel t); |
125 |
val thm = Goal.prove (Simplifier.the_context ss) [] [] prop |
|
37884
314a88278715
discontinued pretending that abel_cancel is logic-independent; cleaned up junk
haftmann
parents:
35845
diff
changeset
|
126 |
(fn _ => rtac @{thm eq_reflection} 1 THEN |
37896 | 127 |
resolve_tac [@{thm diff_eq_diff_less}, @{thm diff_eq_diff_less_eq}, @{thm diff_eq_diff_eq}] 1 THEN |
128 |
cancel_simp_tac ss 1) |
|
37890 | 129 |
in SOME thm end handle Cancel => NONE; |
5589 | 130 |
|
131 |
end; |