src/HOL/Tools/nat_arith.ML
author wenzelm
Sun, 13 Dec 2020 23:11:41 +0100
changeset 72907 3883f536d84d
parent 70490 c42a0a0a9a8d
permissions -rw-r--r--
unused (see 29566b6810f7);
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
30496
7cdcc9dd95cb vague cleanup in arith proof tools setup: deleted dead code, more proper structures, clearer arrangement
haftmann
parents: 29302
diff changeset
     1
(* Author: Markus Wenzel, Stefan Berghofer, and Tobias Nipkow
48560
e0875d956a6b replace Nat_Arith simprocs with simpler conversions that do less rearrangement of terms
huffman
parents: 48559
diff changeset
     2
   Author: Brian Huffman
9436
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
     3
30496
7cdcc9dd95cb vague cleanup in arith proof tools setup: deleted dead code, more proper structures, clearer arrangement
haftmann
parents: 29302
diff changeset
     4
Basic arithmetic for natural numbers.
9436
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
     5
*)
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
     6
30496
7cdcc9dd95cb vague cleanup in arith proof tools setup: deleted dead code, more proper structures, clearer arrangement
haftmann
parents: 29302
diff changeset
     7
signature NAT_ARITH =
26101
a657683e902a tuned structures in arith_data.ML
haftmann
parents: 25484
diff changeset
     8
sig
54742
7a86358a3c0b proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents: 48571
diff changeset
     9
  val cancel_diff_conv: Proof.context -> conv
7a86358a3c0b proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents: 48571
diff changeset
    10
  val cancel_eq_conv: Proof.context -> conv
7a86358a3c0b proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents: 48571
diff changeset
    11
  val cancel_le_conv: Proof.context -> conv
7a86358a3c0b proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents: 48571
diff changeset
    12
  val cancel_less_conv: Proof.context -> conv
26101
a657683e902a tuned structures in arith_data.ML
haftmann
parents: 25484
diff changeset
    13
end;
9436
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
    14
30496
7cdcc9dd95cb vague cleanup in arith proof tools setup: deleted dead code, more proper structures, clearer arrangement
haftmann
parents: 29302
diff changeset
    15
structure Nat_Arith: NAT_ARITH =
9436
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
    16
struct
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
    17
48560
e0875d956a6b replace Nat_Arith simprocs with simpler conversions that do less rearrangement of terms
huffman
parents: 48559
diff changeset
    18
val norm_rules = map mk_meta_eq @{thms add_0_left add_0_right}
9436
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
    19
54742
7a86358a3c0b proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents: 48571
diff changeset
    20
fun move_to_front ctxt path = Conv.every_conv
70490
c42a0a0a9a8d prefer named lemmas -- more compact proofterms;
wenzelm
parents: 67149
diff changeset
    21
    [Conv.rewr_conv (Library.foldl (op RS) (@{thm nat_arith.rule0}, path)),
54742
7a86358a3c0b proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents: 48571
diff changeset
    22
     Conv.arg_conv (Raw_Simplifier.rewrite ctxt false norm_rules)]
9436
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
    23
67149
e61557884799 prefer control symbol antiquotations;
wenzelm
parents: 62365
diff changeset
    24
fun add_atoms path (Const (\<^const_name>\<open>Groups.plus\<close>, _) $ x $ y) =
70490
c42a0a0a9a8d prefer named lemmas -- more compact proofterms;
wenzelm
parents: 67149
diff changeset
    25
      add_atoms (@{thm nat_arith.add1}::path) x #>
c42a0a0a9a8d prefer named lemmas -- more compact proofterms;
wenzelm
parents: 67149
diff changeset
    26
      add_atoms (@{thm nat_arith.add2}::path) y
67149
e61557884799 prefer control symbol antiquotations;
wenzelm
parents: 62365
diff changeset
    27
  | add_atoms path (Const (\<^const_name>\<open>Nat.Suc\<close>, _) $ x) =
70490
c42a0a0a9a8d prefer named lemmas -- more compact proofterms;
wenzelm
parents: 67149
diff changeset
    28
      add_atoms (@{thm nat_arith.suc1}::path) x
67149
e61557884799 prefer control symbol antiquotations;
wenzelm
parents: 62365
diff changeset
    29
  | add_atoms _ (Const (\<^const_name>\<open>Groups.zero\<close>, _)) = I
48560
e0875d956a6b replace Nat_Arith simprocs with simpler conversions that do less rearrangement of terms
huffman
parents: 48559
diff changeset
    30
  | add_atoms path x = cons (x, path)
e0875d956a6b replace Nat_Arith simprocs with simpler conversions that do less rearrangement of terms
huffman
parents: 48559
diff changeset
    31
e0875d956a6b replace Nat_Arith simprocs with simpler conversions that do less rearrangement of terms
huffman
parents: 48559
diff changeset
    32
fun atoms t = add_atoms [] t []
e0875d956a6b replace Nat_Arith simprocs with simpler conversions that do less rearrangement of terms
huffman
parents: 48559
diff changeset
    33
e0875d956a6b replace Nat_Arith simprocs with simpler conversions that do less rearrangement of terms
huffman
parents: 48559
diff changeset
    34
exception Cancel
9436
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
    35
48560
e0875d956a6b replace Nat_Arith simprocs with simpler conversions that do less rearrangement of terms
huffman
parents: 48559
diff changeset
    36
fun find_common ord xs ys =
e0875d956a6b replace Nat_Arith simprocs with simpler conversions that do less rearrangement of terms
huffman
parents: 48559
diff changeset
    37
  let
e0875d956a6b replace Nat_Arith simprocs with simpler conversions that do less rearrangement of terms
huffman
parents: 48559
diff changeset
    38
    fun find (xs as (x, px)::xs') (ys as (y, py)::ys') =
e0875d956a6b replace Nat_Arith simprocs with simpler conversions that do less rearrangement of terms
huffman
parents: 48559
diff changeset
    39
        (case ord (x, y) of
e0875d956a6b replace Nat_Arith simprocs with simpler conversions that do less rearrangement of terms
huffman
parents: 48559
diff changeset
    40
          EQUAL => (px, py)
e0875d956a6b replace Nat_Arith simprocs with simpler conversions that do less rearrangement of terms
huffman
parents: 48559
diff changeset
    41
        | LESS => find xs' ys
e0875d956a6b replace Nat_Arith simprocs with simpler conversions that do less rearrangement of terms
huffman
parents: 48559
diff changeset
    42
        | GREATER => find xs ys')
e0875d956a6b replace Nat_Arith simprocs with simpler conversions that do less rearrangement of terms
huffman
parents: 48559
diff changeset
    43
      | find _ _ = raise Cancel
e0875d956a6b replace Nat_Arith simprocs with simpler conversions that do less rearrangement of terms
huffman
parents: 48559
diff changeset
    44
    fun ord' ((x, _), (y, _)) = ord (x, y)
e0875d956a6b replace Nat_Arith simprocs with simpler conversions that do less rearrangement of terms
huffman
parents: 48559
diff changeset
    45
  in
e0875d956a6b replace Nat_Arith simprocs with simpler conversions that do less rearrangement of terms
huffman
parents: 48559
diff changeset
    46
    find (sort ord' xs) (sort ord' ys)
e0875d956a6b replace Nat_Arith simprocs with simpler conversions that do less rearrangement of terms
huffman
parents: 48559
diff changeset
    47
  end
9436
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
    48
54742
7a86358a3c0b proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents: 48571
diff changeset
    49
fun cancel_conv rule ctxt ct =
48560
e0875d956a6b replace Nat_Arith simprocs with simpler conversions that do less rearrangement of terms
huffman
parents: 48559
diff changeset
    50
  let
e0875d956a6b replace Nat_Arith simprocs with simpler conversions that do less rearrangement of terms
huffman
parents: 48559
diff changeset
    51
    val ((_, lhs), rhs) = (apfst dest_comb o dest_comb) (Thm.term_of ct)
e0875d956a6b replace Nat_Arith simprocs with simpler conversions that do less rearrangement of terms
huffman
parents: 48559
diff changeset
    52
    val (lpath, rpath) = find_common Term_Ord.term_ord (atoms lhs) (atoms rhs)
54742
7a86358a3c0b proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents: 48571
diff changeset
    53
    val lconv = move_to_front ctxt lpath
7a86358a3c0b proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents: 48571
diff changeset
    54
    val rconv = move_to_front ctxt rpath
48560
e0875d956a6b replace Nat_Arith simprocs with simpler conversions that do less rearrangement of terms
huffman
parents: 48559
diff changeset
    55
    val conv1 = Conv.combination_conv (Conv.arg_conv lconv) rconv
e0875d956a6b replace Nat_Arith simprocs with simpler conversions that do less rearrangement of terms
huffman
parents: 48559
diff changeset
    56
    val conv = conv1 then_conv Conv.rewr_conv rule
48571
d68b74435605 move exception handlers outside of let block
huffman
parents: 48561
diff changeset
    57
  in conv ct end
d68b74435605 move exception handlers outside of let block
huffman
parents: 48561
diff changeset
    58
    handle Cancel => raise CTERM ("no_conversion", [])
9436
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
    59
62365
8a105c235b1f sorted out some duplicate fact bindings
haftmann
parents: 57514
diff changeset
    60
val cancel_diff_conv = cancel_conv (mk_meta_eq @{thm add_diff_cancel_left [where ?'a = nat]})
8a105c235b1f sorted out some duplicate fact bindings
haftmann
parents: 57514
diff changeset
    61
val cancel_eq_conv = cancel_conv (mk_meta_eq @{thm add_left_cancel [where ?'a = nat]})
8a105c235b1f sorted out some duplicate fact bindings
haftmann
parents: 57514
diff changeset
    62
val cancel_le_conv = cancel_conv (mk_meta_eq @{thm add_le_cancel_left [where ?'a = nat]})
8a105c235b1f sorted out some duplicate fact bindings
haftmann
parents: 57514
diff changeset
    63
val cancel_less_conv = cancel_conv (mk_meta_eq @{thm add_less_cancel_left [where ?'a = nat]})
24076
ae946f751c44 arith method setup: proper context;
wenzelm
parents: 23881
diff changeset
    64
24095
785c3cd7fcb5 moved lin_arith stuff to Tools/lin_arith.ML;
wenzelm
parents: 24076
diff changeset
    65
end;