| author | wenzelm | 
| Fri, 05 Aug 2016 20:26:13 +0200 | |
| changeset 63615 | d786d54efc70 | 
| parent 57514 | bdc2c6b40bf2 | 
| child 67149 | e61557884799 | 
| permissions | -rw-r--r-- | 
| 48556 
62a3fbf9d35b
replace abel_cancel simprocs with functionally equivalent, but simpler and faster ones
 huffman parents: diff
changeset | 1 | (* Title: HOL/Tools/group_cancel.ML | 
| 
62a3fbf9d35b
replace abel_cancel simprocs with functionally equivalent, but simpler and faster ones
 huffman parents: diff
changeset | 2 | Author: Brian Huffman, TU Munich | 
| 
62a3fbf9d35b
replace abel_cancel simprocs with functionally equivalent, but simpler and faster ones
 huffman parents: diff
changeset | 3 | |
| 
62a3fbf9d35b
replace abel_cancel simprocs with functionally equivalent, but simpler and faster ones
 huffman parents: diff
changeset | 4 | Simplification procedures for abelian groups: | 
| 
62a3fbf9d35b
replace abel_cancel simprocs with functionally equivalent, but simpler and faster ones
 huffman parents: diff
changeset | 5 | - Cancel complementary terms in sums. | 
| 
62a3fbf9d35b
replace abel_cancel simprocs with functionally equivalent, but simpler and faster ones
 huffman parents: diff
changeset | 6 | - Cancel like terms on opposite sides of relations. | 
| 
62a3fbf9d35b
replace abel_cancel simprocs with functionally equivalent, but simpler and faster ones
 huffman parents: diff
changeset | 7 | *) | 
| 
62a3fbf9d35b
replace abel_cancel simprocs with functionally equivalent, but simpler and faster ones
 huffman parents: diff
changeset | 8 | |
| 
62a3fbf9d35b
replace abel_cancel simprocs with functionally equivalent, but simpler and faster ones
 huffman parents: diff
changeset | 9 | signature GROUP_CANCEL = | 
| 
62a3fbf9d35b
replace abel_cancel simprocs with functionally equivalent, but simpler and faster ones
 huffman parents: diff
changeset | 10 | sig | 
| 
62a3fbf9d35b
replace abel_cancel simprocs with functionally equivalent, but simpler and faster ones
 huffman parents: diff
changeset | 11 | val cancel_diff_conv: conv | 
| 
62a3fbf9d35b
replace abel_cancel simprocs with functionally equivalent, but simpler and faster ones
 huffman parents: diff
changeset | 12 | val cancel_eq_conv: conv | 
| 
62a3fbf9d35b
replace abel_cancel simprocs with functionally equivalent, but simpler and faster ones
 huffman parents: diff
changeset | 13 | val cancel_le_conv: conv | 
| 
62a3fbf9d35b
replace abel_cancel simprocs with functionally equivalent, but simpler and faster ones
 huffman parents: diff
changeset | 14 | val cancel_less_conv: conv | 
| 
62a3fbf9d35b
replace abel_cancel simprocs with functionally equivalent, but simpler and faster ones
 huffman parents: diff
changeset | 15 | val cancel_add_conv: conv | 
| 
62a3fbf9d35b
replace abel_cancel simprocs with functionally equivalent, but simpler and faster ones
 huffman parents: diff
changeset | 16 | end | 
| 
62a3fbf9d35b
replace abel_cancel simprocs with functionally equivalent, but simpler and faster ones
 huffman parents: diff
changeset | 17 | |
| 
62a3fbf9d35b
replace abel_cancel simprocs with functionally equivalent, but simpler and faster ones
 huffman parents: diff
changeset | 18 | structure Group_Cancel: GROUP_CANCEL = | 
| 
62a3fbf9d35b
replace abel_cancel simprocs with functionally equivalent, but simpler and faster ones
 huffman parents: diff
changeset | 19 | struct | 
| 
62a3fbf9d35b
replace abel_cancel simprocs with functionally equivalent, but simpler and faster ones
 huffman parents: diff
changeset | 20 | |
| 
62a3fbf9d35b
replace abel_cancel simprocs with functionally equivalent, but simpler and faster ones
 huffman parents: diff
changeset | 21 | val add1 = @{lemma "(A::'a::comm_monoid_add) == k + a ==> A + b == k + (a + b)"
 | 
| 57514 
bdc2c6b40bf2
prefer ac_simps collections over separate name bindings for add and mult
 haftmann parents: 
54230diff
changeset | 22 | by (simp only: ac_simps)} | 
| 48556 
62a3fbf9d35b
replace abel_cancel simprocs with functionally equivalent, but simpler and faster ones
 huffman parents: diff
changeset | 23 | val add2 = @{lemma "(B::'a::comm_monoid_add) == k + b ==> a + B == k + (a + b)"
 | 
| 57514 
bdc2c6b40bf2
prefer ac_simps collections over separate name bindings for add and mult
 haftmann parents: 
54230diff
changeset | 24 | by (simp only: ac_simps)} | 
| 48556 
62a3fbf9d35b
replace abel_cancel simprocs with functionally equivalent, but simpler and faster ones
 huffman parents: diff
changeset | 25 | val sub1 = @{lemma "(A::'a::ab_group_add) == k + a ==> A - b == k + (a - b)"
 | 
| 
62a3fbf9d35b
replace abel_cancel simprocs with functionally equivalent, but simpler and faster ones
 huffman parents: diff
changeset | 26 | by (simp only: add_diff_eq)} | 
| 
62a3fbf9d35b
replace abel_cancel simprocs with functionally equivalent, but simpler and faster ones
 huffman parents: diff
changeset | 27 | val sub2 = @{lemma "(B::'a::ab_group_add) == k + b ==> a - B == - k + (a - b)"
 | 
| 57514 
bdc2c6b40bf2
prefer ac_simps collections over separate name bindings for add and mult
 haftmann parents: 
54230diff
changeset | 28 | by (simp only: minus_add diff_conv_add_uminus ac_simps)} | 
| 48556 
62a3fbf9d35b
replace abel_cancel simprocs with functionally equivalent, but simpler and faster ones
 huffman parents: diff
changeset | 29 | val neg1 = @{lemma "(A::'a::ab_group_add) == k + a ==> - A == - k + - a"
 | 
| 
62a3fbf9d35b
replace abel_cancel simprocs with functionally equivalent, but simpler and faster ones
 huffman parents: diff
changeset | 30 | by (simp only: minus_add_distrib)} | 
| 
62a3fbf9d35b
replace abel_cancel simprocs with functionally equivalent, but simpler and faster ones
 huffman parents: diff
changeset | 31 | val rule0 = @{lemma "(a::'a::comm_monoid_add) == a + 0"
 | 
| 
62a3fbf9d35b
replace abel_cancel simprocs with functionally equivalent, but simpler and faster ones
 huffman parents: diff
changeset | 32 | by (simp only: add_0_right)} | 
| 
62a3fbf9d35b
replace abel_cancel simprocs with functionally equivalent, but simpler and faster ones
 huffman parents: diff
changeset | 33 | val minus_minus = mk_meta_eq @{thm minus_minus}
 | 
| 
62a3fbf9d35b
replace abel_cancel simprocs with functionally equivalent, but simpler and faster ones
 huffman parents: diff
changeset | 34 | |
| 
62a3fbf9d35b
replace abel_cancel simprocs with functionally equivalent, but simpler and faster ones
 huffman parents: diff
changeset | 35 | fun move_to_front path = Conv.every_conv | 
| 
62a3fbf9d35b
replace abel_cancel simprocs with functionally equivalent, but simpler and faster ones
 huffman parents: diff
changeset | 36 | [Conv.rewr_conv (Library.foldl (op RS) (rule0, path)), | 
| 
62a3fbf9d35b
replace abel_cancel simprocs with functionally equivalent, but simpler and faster ones
 huffman parents: diff
changeset | 37 | Conv.arg1_conv (Conv.repeat_conv (Conv.rewr_conv minus_minus))] | 
| 
62a3fbf9d35b
replace abel_cancel simprocs with functionally equivalent, but simpler and faster ones
 huffman parents: diff
changeset | 38 | |
| 
62a3fbf9d35b
replace abel_cancel simprocs with functionally equivalent, but simpler and faster ones
 huffman parents: diff
changeset | 39 | fun add_atoms pos path (Const (@{const_name Groups.plus}, _) $ x $ y) =
 | 
| 
62a3fbf9d35b
replace abel_cancel simprocs with functionally equivalent, but simpler and faster ones
 huffman parents: diff
changeset | 40 | add_atoms pos (add1::path) x #> add_atoms pos (add2::path) y | 
| 
62a3fbf9d35b
replace abel_cancel simprocs with functionally equivalent, but simpler and faster ones
 huffman parents: diff
changeset | 41 |   | add_atoms pos path (Const (@{const_name Groups.minus}, _) $ x $ y) =
 | 
| 
62a3fbf9d35b
replace abel_cancel simprocs with functionally equivalent, but simpler and faster ones
 huffman parents: diff
changeset | 42 | add_atoms pos (sub1::path) x #> add_atoms (not pos) (sub2::path) y | 
| 
62a3fbf9d35b
replace abel_cancel simprocs with functionally equivalent, but simpler and faster ones
 huffman parents: diff
changeset | 43 |   | add_atoms pos path (Const (@{const_name Groups.uminus}, _) $ x) =
 | 
| 
62a3fbf9d35b
replace abel_cancel simprocs with functionally equivalent, but simpler and faster ones
 huffman parents: diff
changeset | 44 | add_atoms (not pos) (neg1::path) x | 
| 
62a3fbf9d35b
replace abel_cancel simprocs with functionally equivalent, but simpler and faster ones
 huffman parents: diff
changeset | 45 |   | add_atoms _ _ (Const (@{const_name Groups.zero}, _)) = I
 | 
| 
62a3fbf9d35b
replace abel_cancel simprocs with functionally equivalent, but simpler and faster ones
 huffman parents: diff
changeset | 46 | | add_atoms pos path x = cons ((pos, x), path) | 
| 
62a3fbf9d35b
replace abel_cancel simprocs with functionally equivalent, but simpler and faster ones
 huffman parents: diff
changeset | 47 | |
| 
62a3fbf9d35b
replace abel_cancel simprocs with functionally equivalent, but simpler and faster ones
 huffman parents: diff
changeset | 48 | fun atoms t = add_atoms true [] t [] | 
| 
62a3fbf9d35b
replace abel_cancel simprocs with functionally equivalent, but simpler and faster ones
 huffman parents: diff
changeset | 49 | |
| 
62a3fbf9d35b
replace abel_cancel simprocs with functionally equivalent, but simpler and faster ones
 huffman parents: diff
changeset | 50 | val coeff_ord = prod_ord bool_ord Term_Ord.term_ord | 
| 
62a3fbf9d35b
replace abel_cancel simprocs with functionally equivalent, but simpler and faster ones
 huffman parents: diff
changeset | 51 | |
| 48694 
188cbbce880e
modify group_cancel simprocs so that they can cancel multiple terms at once
 huffman parents: 
48571diff
changeset | 52 | fun find_all_common ord xs ys = | 
| 48556 
62a3fbf9d35b
replace abel_cancel simprocs with functionally equivalent, but simpler and faster ones
 huffman parents: diff
changeset | 53 | let | 
| 
62a3fbf9d35b
replace abel_cancel simprocs with functionally equivalent, but simpler and faster ones
 huffman parents: diff
changeset | 54 | fun find (xs as (x, px)::xs') (ys as (y, py)::ys') = | 
| 
62a3fbf9d35b
replace abel_cancel simprocs with functionally equivalent, but simpler and faster ones
 huffman parents: diff
changeset | 55 | (case ord (x, y) of | 
| 48694 
188cbbce880e
modify group_cancel simprocs so that they can cancel multiple terms at once
 huffman parents: 
48571diff
changeset | 56 | EQUAL => (px, py) :: find xs' ys' | 
| 48556 
62a3fbf9d35b
replace abel_cancel simprocs with functionally equivalent, but simpler and faster ones
 huffman parents: diff
changeset | 57 | | LESS => find xs' ys | 
| 
62a3fbf9d35b
replace abel_cancel simprocs with functionally equivalent, but simpler and faster ones
 huffman parents: diff
changeset | 58 | | GREATER => find xs ys') | 
| 48694 
188cbbce880e
modify group_cancel simprocs so that they can cancel multiple terms at once
 huffman parents: 
48571diff
changeset | 59 | | find _ _ = [] | 
| 48556 
62a3fbf9d35b
replace abel_cancel simprocs with functionally equivalent, but simpler and faster ones
 huffman parents: diff
changeset | 60 | fun ord' ((x, _), (y, _)) = ord (x, y) | 
| 
62a3fbf9d35b
replace abel_cancel simprocs with functionally equivalent, but simpler and faster ones
 huffman parents: diff
changeset | 61 | in | 
| 
62a3fbf9d35b
replace abel_cancel simprocs with functionally equivalent, but simpler and faster ones
 huffman parents: diff
changeset | 62 | find (sort ord' xs) (sort ord' ys) | 
| 
62a3fbf9d35b
replace abel_cancel simprocs with functionally equivalent, but simpler and faster ones
 huffman parents: diff
changeset | 63 | end | 
| 
62a3fbf9d35b
replace abel_cancel simprocs with functionally equivalent, but simpler and faster ones
 huffman parents: diff
changeset | 64 | |
| 
62a3fbf9d35b
replace abel_cancel simprocs with functionally equivalent, but simpler and faster ones
 huffman parents: diff
changeset | 65 | fun cancel_conv rule ct = | 
| 
62a3fbf9d35b
replace abel_cancel simprocs with functionally equivalent, but simpler and faster ones
 huffman parents: diff
changeset | 66 | let | 
| 48694 
188cbbce880e
modify group_cancel simprocs so that they can cancel multiple terms at once
 huffman parents: 
48571diff
changeset | 67 | fun cancel1_conv (lpath, rpath) = | 
| 
188cbbce880e
modify group_cancel simprocs so that they can cancel multiple terms at once
 huffman parents: 
48571diff
changeset | 68 | let | 
| 
188cbbce880e
modify group_cancel simprocs so that they can cancel multiple terms at once
 huffman parents: 
48571diff
changeset | 69 | val lconv = move_to_front lpath | 
| 
188cbbce880e
modify group_cancel simprocs so that they can cancel multiple terms at once
 huffman parents: 
48571diff
changeset | 70 | val rconv = move_to_front rpath | 
| 
188cbbce880e
modify group_cancel simprocs so that they can cancel multiple terms at once
 huffman parents: 
48571diff
changeset | 71 | val conv1 = Conv.combination_conv (Conv.arg_conv lconv) rconv | 
| 
188cbbce880e
modify group_cancel simprocs so that they can cancel multiple terms at once
 huffman parents: 
48571diff
changeset | 72 | in | 
| 
188cbbce880e
modify group_cancel simprocs so that they can cancel multiple terms at once
 huffman parents: 
48571diff
changeset | 73 | conv1 then_conv Conv.rewr_conv rule | 
| 
188cbbce880e
modify group_cancel simprocs so that they can cancel multiple terms at once
 huffman parents: 
48571diff
changeset | 74 | end | 
| 48556 
62a3fbf9d35b
replace abel_cancel simprocs with functionally equivalent, but simpler and faster ones
 huffman parents: diff
changeset | 75 | val ((_, lhs), rhs) = (apfst dest_comb o dest_comb) (Thm.term_of ct) | 
| 48694 
188cbbce880e
modify group_cancel simprocs so that they can cancel multiple terms at once
 huffman parents: 
48571diff
changeset | 76 | val common = find_all_common coeff_ord (atoms lhs) (atoms rhs) | 
| 
188cbbce880e
modify group_cancel simprocs so that they can cancel multiple terms at once
 huffman parents: 
48571diff
changeset | 77 | val conv = | 
| 
188cbbce880e
modify group_cancel simprocs so that they can cancel multiple terms at once
 huffman parents: 
48571diff
changeset | 78 | if null common then Conv.no_conv | 
| 
188cbbce880e
modify group_cancel simprocs so that they can cancel multiple terms at once
 huffman parents: 
48571diff
changeset | 79 | else Conv.every_conv (map cancel1_conv common) | 
| 48571 | 80 | in conv ct end | 
| 48556 
62a3fbf9d35b
replace abel_cancel simprocs with functionally equivalent, but simpler and faster ones
 huffman parents: diff
changeset | 81 | |
| 
62a3fbf9d35b
replace abel_cancel simprocs with functionally equivalent, but simpler and faster ones
 huffman parents: diff
changeset | 82 | val cancel_diff_conv = cancel_conv (mk_meta_eq @{thm add_diff_cancel_left})
 | 
| 
62a3fbf9d35b
replace abel_cancel simprocs with functionally equivalent, but simpler and faster ones
 huffman parents: diff
changeset | 83 | val cancel_eq_conv = cancel_conv (mk_meta_eq @{thm add_left_cancel})
 | 
| 
62a3fbf9d35b
replace abel_cancel simprocs with functionally equivalent, but simpler and faster ones
 huffman parents: diff
changeset | 84 | val cancel_le_conv = cancel_conv (mk_meta_eq @{thm add_le_cancel_left})
 | 
| 
62a3fbf9d35b
replace abel_cancel simprocs with functionally equivalent, but simpler and faster ones
 huffman parents: diff
changeset | 85 | val cancel_less_conv = cancel_conv (mk_meta_eq @{thm add_less_cancel_left})
 | 
| 
62a3fbf9d35b
replace abel_cancel simprocs with functionally equivalent, but simpler and faster ones
 huffman parents: diff
changeset | 86 | |
| 
62a3fbf9d35b
replace abel_cancel simprocs with functionally equivalent, but simpler and faster ones
 huffman parents: diff
changeset | 87 | val diff_minus_eq_add = mk_meta_eq @{thm diff_minus_eq_add}
 | 
| 
62a3fbf9d35b
replace abel_cancel simprocs with functionally equivalent, but simpler and faster ones
 huffman parents: diff
changeset | 88 | val add_eq_diff_minus = Thm.symmetric diff_minus_eq_add | 
| 
62a3fbf9d35b
replace abel_cancel simprocs with functionally equivalent, but simpler and faster ones
 huffman parents: diff
changeset | 89 | val cancel_add_conv = Conv.every_conv | 
| 
62a3fbf9d35b
replace abel_cancel simprocs with functionally equivalent, but simpler and faster ones
 huffman parents: diff
changeset | 90 | [Conv.rewr_conv add_eq_diff_minus, | 
| 
62a3fbf9d35b
replace abel_cancel simprocs with functionally equivalent, but simpler and faster ones
 huffman parents: diff
changeset | 91 | cancel_diff_conv, | 
| 
62a3fbf9d35b
replace abel_cancel simprocs with functionally equivalent, but simpler and faster ones
 huffman parents: diff
changeset | 92 | Conv.rewr_conv diff_minus_eq_add] | 
| 
62a3fbf9d35b
replace abel_cancel simprocs with functionally equivalent, but simpler and faster ones
 huffman parents: diff
changeset | 93 | |
| 
62a3fbf9d35b
replace abel_cancel simprocs with functionally equivalent, but simpler and faster ones
 huffman parents: diff
changeset | 94 | end |