src/HOL/Tools/group_cancel.ML
author wenzelm
Sun, 17 Dec 2023 21:34:44 +0100
changeset 79272 899f37f6d218
parent 70490 c42a0a0a9a8d
child 80722 b7d051e25d9d
permissions -rw-r--r--
proper beta_norm after instantiation (amending 90c5aadcc4b2);
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
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 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
    22
62a3fbf9d35b replace abel_cancel simprocs with functionally equivalent, but simpler and faster ones
huffman
parents:
diff changeset
    23
fun move_to_front path = Conv.every_conv
70490
c42a0a0a9a8d prefer named lemmas -- more compact proofterms;
wenzelm
parents: 67149
diff changeset
    24
    [Conv.rewr_conv (Library.foldl (op RS) (@{thm group_cancel.rule0}, path)),
48556
62a3fbf9d35b replace abel_cancel simprocs with functionally equivalent, but simpler and faster ones
huffman
parents:
diff changeset
    25
     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
    26
67149
e61557884799 prefer control symbol antiquotations;
wenzelm
parents: 57514
diff changeset
    27
fun add_atoms pos path (Const (\<^const_name>\<open>Groups.plus\<close>, _) $ x $ y) =
70490
c42a0a0a9a8d prefer named lemmas -- more compact proofterms;
wenzelm
parents: 67149
diff changeset
    28
      add_atoms pos (@{thm group_cancel.add1}::path) x #>
c42a0a0a9a8d prefer named lemmas -- more compact proofterms;
wenzelm
parents: 67149
diff changeset
    29
      add_atoms pos (@{thm group_cancel.add2}::path) y
67149
e61557884799 prefer control symbol antiquotations;
wenzelm
parents: 57514
diff changeset
    30
  | add_atoms pos path (Const (\<^const_name>\<open>Groups.minus\<close>, _) $ x $ y) =
70490
c42a0a0a9a8d prefer named lemmas -- more compact proofterms;
wenzelm
parents: 67149
diff changeset
    31
      add_atoms pos (@{thm group_cancel.sub1}::path) x #>
c42a0a0a9a8d prefer named lemmas -- more compact proofterms;
wenzelm
parents: 67149
diff changeset
    32
      add_atoms (not pos) (@{thm group_cancel.sub2}::path) y
67149
e61557884799 prefer control symbol antiquotations;
wenzelm
parents: 57514
diff changeset
    33
  | add_atoms pos path (Const (\<^const_name>\<open>Groups.uminus\<close>, _) $ x) =
70490
c42a0a0a9a8d prefer named lemmas -- more compact proofterms;
wenzelm
parents: 67149
diff changeset
    34
      add_atoms (not pos) (@{thm group_cancel.neg1}::path) x
67149
e61557884799 prefer control symbol antiquotations;
wenzelm
parents: 57514
diff changeset
    35
  | add_atoms _ _ (Const (\<^const_name>\<open>Groups.zero\<close>, _)) = I
48556
62a3fbf9d35b replace abel_cancel simprocs with functionally equivalent, but simpler and faster ones
huffman
parents:
diff changeset
    36
  | 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
    37
62a3fbf9d35b replace abel_cancel simprocs with functionally equivalent, but simpler and faster ones
huffman
parents:
diff changeset
    38
fun atoms t = add_atoms true [] t []
62a3fbf9d35b replace abel_cancel simprocs with functionally equivalent, but simpler and faster ones
huffman
parents:
diff changeset
    39
62a3fbf9d35b replace abel_cancel simprocs with functionally equivalent, but simpler and faster ones
huffman
parents:
diff changeset
    40
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
    41
48694
188cbbce880e modify group_cancel simprocs so that they can cancel multiple terms at once
huffman
parents: 48571
diff changeset
    42
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
    43
  let
62a3fbf9d35b replace abel_cancel simprocs with functionally equivalent, but simpler and faster ones
huffman
parents:
diff changeset
    44
    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
    45
        (case ord (x, y) of
48694
188cbbce880e modify group_cancel simprocs so that they can cancel multiple terms at once
huffman
parents: 48571
diff changeset
    46
          EQUAL => (px, py) :: find xs' ys'
48556
62a3fbf9d35b replace abel_cancel simprocs with functionally equivalent, but simpler and faster ones
huffman
parents:
diff changeset
    47
        | LESS => find xs' ys
62a3fbf9d35b replace abel_cancel simprocs with functionally equivalent, but simpler and faster ones
huffman
parents:
diff changeset
    48
        | GREATER => find xs ys')
48694
188cbbce880e modify group_cancel simprocs so that they can cancel multiple terms at once
huffman
parents: 48571
diff changeset
    49
      | find _ _ = []
48556
62a3fbf9d35b replace abel_cancel simprocs with functionally equivalent, but simpler and faster ones
huffman
parents:
diff changeset
    50
    fun ord' ((x, _), (y, _)) = ord (x, y)
62a3fbf9d35b replace abel_cancel simprocs with functionally equivalent, but simpler and faster ones
huffman
parents:
diff changeset
    51
  in
62a3fbf9d35b replace abel_cancel simprocs with functionally equivalent, but simpler and faster ones
huffman
parents:
diff changeset
    52
    find (sort ord' xs) (sort ord' ys)
62a3fbf9d35b replace abel_cancel simprocs with functionally equivalent, but simpler and faster ones
huffman
parents:
diff changeset
    53
  end
62a3fbf9d35b replace abel_cancel simprocs with functionally equivalent, but simpler and faster ones
huffman
parents:
diff changeset
    54
62a3fbf9d35b replace abel_cancel simprocs with functionally equivalent, but simpler and faster ones
huffman
parents:
diff changeset
    55
fun cancel_conv rule ct =
62a3fbf9d35b replace abel_cancel simprocs with functionally equivalent, but simpler and faster ones
huffman
parents:
diff changeset
    56
  let
48694
188cbbce880e modify group_cancel simprocs so that they can cancel multiple terms at once
huffman
parents: 48571
diff changeset
    57
    fun cancel1_conv (lpath, rpath) =
188cbbce880e modify group_cancel simprocs so that they can cancel multiple terms at once
huffman
parents: 48571
diff changeset
    58
      let
188cbbce880e modify group_cancel simprocs so that they can cancel multiple terms at once
huffman
parents: 48571
diff changeset
    59
        val lconv = move_to_front lpath
188cbbce880e modify group_cancel simprocs so that they can cancel multiple terms at once
huffman
parents: 48571
diff changeset
    60
        val rconv = move_to_front rpath
188cbbce880e modify group_cancel simprocs so that they can cancel multiple terms at once
huffman
parents: 48571
diff changeset
    61
        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: 48571
diff changeset
    62
      in
188cbbce880e modify group_cancel simprocs so that they can cancel multiple terms at once
huffman
parents: 48571
diff changeset
    63
        conv1 then_conv Conv.rewr_conv rule
188cbbce880e modify group_cancel simprocs so that they can cancel multiple terms at once
huffman
parents: 48571
diff changeset
    64
      end
48556
62a3fbf9d35b replace abel_cancel simprocs with functionally equivalent, but simpler and faster ones
huffman
parents:
diff changeset
    65
    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: 48571
diff changeset
    66
    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: 48571
diff changeset
    67
    val conv =
188cbbce880e modify group_cancel simprocs so that they can cancel multiple terms at once
huffman
parents: 48571
diff changeset
    68
      if null common then Conv.no_conv
188cbbce880e modify group_cancel simprocs so that they can cancel multiple terms at once
huffman
parents: 48571
diff changeset
    69
      else Conv.every_conv (map cancel1_conv common)
48571
d68b74435605 move exception handlers outside of let block
huffman
parents: 48556
diff changeset
    70
  in conv ct end
48556
62a3fbf9d35b replace abel_cancel simprocs with functionally equivalent, but simpler and faster ones
huffman
parents:
diff changeset
    71
62a3fbf9d35b replace abel_cancel simprocs with functionally equivalent, but simpler and faster ones
huffman
parents:
diff changeset
    72
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
    73
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
    74
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
    75
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
    76
62a3fbf9d35b replace abel_cancel simprocs with functionally equivalent, but simpler and faster ones
huffman
parents:
diff changeset
    77
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
    78
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
    79
val cancel_add_conv = Conv.every_conv
62a3fbf9d35b replace abel_cancel simprocs with functionally equivalent, but simpler and faster ones
huffman
parents:
diff changeset
    80
  [Conv.rewr_conv add_eq_diff_minus,
62a3fbf9d35b replace abel_cancel simprocs with functionally equivalent, but simpler and faster ones
huffman
parents:
diff changeset
    81
   cancel_diff_conv,
62a3fbf9d35b replace abel_cancel simprocs with functionally equivalent, but simpler and faster ones
huffman
parents:
diff changeset
    82
   Conv.rewr_conv diff_minus_eq_add]
62a3fbf9d35b replace abel_cancel simprocs with functionally equivalent, but simpler and faster ones
huffman
parents:
diff changeset
    83
62a3fbf9d35b replace abel_cancel simprocs with functionally equivalent, but simpler and faster ones
huffman
parents:
diff changeset
    84
end