src/HOL/Tools/nat_arith.ML
author wenzelm
Thu, 06 Mar 2014 22:15:01 +0100
changeset 55965 0c2c61a87a7d
parent 54742 7a86358a3c0b
child 57514 bdc2c6b40bf2
permissions -rw-r--r--
merged
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 add1 = @{lemma "(A::'a::comm_monoid_add) == k + a ==> A + b == k + (a + b)"
e0875d956a6b replace Nat_Arith simprocs with simpler conversions that do less rearrangement of terms
huffman
parents: 48559
diff changeset
    19
      by (simp only: add_ac)}
e0875d956a6b replace Nat_Arith simprocs with simpler conversions that do less rearrangement of terms
huffman
parents: 48559
diff changeset
    20
val add2 = @{lemma "(B::'a::comm_monoid_add) == k + b ==> a + B == k + (a + b)"
e0875d956a6b replace Nat_Arith simprocs with simpler conversions that do less rearrangement of terms
huffman
parents: 48559
diff changeset
    21
      by (simp only: add_ac)}
e0875d956a6b replace Nat_Arith simprocs with simpler conversions that do less rearrangement of terms
huffman
parents: 48559
diff changeset
    22
val suc1 = @{lemma "A == k + a ==> Suc A == k + Suc a"
e0875d956a6b replace Nat_Arith simprocs with simpler conversions that do less rearrangement of terms
huffman
parents: 48559
diff changeset
    23
      by (simp only: add_Suc_right)}
e0875d956a6b replace Nat_Arith simprocs with simpler conversions that do less rearrangement of terms
huffman
parents: 48559
diff changeset
    24
val rule0 = @{lemma "(a::'a::comm_monoid_add) == a + 0"
e0875d956a6b replace Nat_Arith simprocs with simpler conversions that do less rearrangement of terms
huffman
parents: 48559
diff changeset
    25
      by (simp only: add_0_right)}
20217
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 20044
diff changeset
    26
48560
e0875d956a6b replace Nat_Arith simprocs with simpler conversions that do less rearrangement of terms
huffman
parents: 48559
diff changeset
    27
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
    28
54742
7a86358a3c0b proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents: 48571
diff changeset
    29
fun move_to_front ctxt path = Conv.every_conv
48560
e0875d956a6b replace Nat_Arith simprocs with simpler conversions that do less rearrangement of terms
huffman
parents: 48559
diff changeset
    30
    [Conv.rewr_conv (Library.foldl (op RS) (rule0, path)),
54742
7a86358a3c0b proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents: 48571
diff changeset
    31
     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
    32
48560
e0875d956a6b replace Nat_Arith simprocs with simpler conversions that do less rearrangement of terms
huffman
parents: 48559
diff changeset
    33
fun add_atoms path (Const (@{const_name Groups.plus}, _) $ x $ y) =
e0875d956a6b replace Nat_Arith simprocs with simpler conversions that do less rearrangement of terms
huffman
parents: 48559
diff changeset
    34
      add_atoms (add1::path) x #> add_atoms (add2::path) y
e0875d956a6b replace Nat_Arith simprocs with simpler conversions that do less rearrangement of terms
huffman
parents: 48559
diff changeset
    35
  | add_atoms path (Const (@{const_name Nat.Suc}, _) $ x) =
e0875d956a6b replace Nat_Arith simprocs with simpler conversions that do less rearrangement of terms
huffman
parents: 48559
diff changeset
    36
      add_atoms (suc1::path) x
e0875d956a6b replace Nat_Arith simprocs with simpler conversions that do less rearrangement of terms
huffman
parents: 48559
diff changeset
    37
  | add_atoms _ (Const (@{const_name Groups.zero}, _)) = I
e0875d956a6b replace Nat_Arith simprocs with simpler conversions that do less rearrangement of terms
huffman
parents: 48559
diff changeset
    38
  | 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
    39
e0875d956a6b replace Nat_Arith simprocs with simpler conversions that do less rearrangement of terms
huffman
parents: 48559
diff changeset
    40
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
    41
e0875d956a6b replace Nat_Arith simprocs with simpler conversions that do less rearrangement of terms
huffman
parents: 48559
diff changeset
    42
exception Cancel
9436
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
    43
48560
e0875d956a6b replace Nat_Arith simprocs with simpler conversions that do less rearrangement of terms
huffman
parents: 48559
diff changeset
    44
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
    45
  let
e0875d956a6b replace Nat_Arith simprocs with simpler conversions that do less rearrangement of terms
huffman
parents: 48559
diff changeset
    46
    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
    47
        (case ord (x, y) of
e0875d956a6b replace Nat_Arith simprocs with simpler conversions that do less rearrangement of terms
huffman
parents: 48559
diff changeset
    48
          EQUAL => (px, py)
e0875d956a6b replace Nat_Arith simprocs with simpler conversions that do less rearrangement of terms
huffman
parents: 48559
diff changeset
    49
        | LESS => find xs' ys
e0875d956a6b replace Nat_Arith simprocs with simpler conversions that do less rearrangement of terms
huffman
parents: 48559
diff changeset
    50
        | GREATER => find xs ys')
e0875d956a6b replace Nat_Arith simprocs with simpler conversions that do less rearrangement of terms
huffman
parents: 48559
diff changeset
    51
      | find _ _ = raise Cancel
e0875d956a6b replace Nat_Arith simprocs with simpler conversions that do less rearrangement of terms
huffman
parents: 48559
diff changeset
    52
    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
    53
  in
e0875d956a6b replace Nat_Arith simprocs with simpler conversions that do less rearrangement of terms
huffman
parents: 48559
diff changeset
    54
    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
    55
  end
9436
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
    56
54742
7a86358a3c0b proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents: 48571
diff changeset
    57
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
    58
  let
e0875d956a6b replace Nat_Arith simprocs with simpler conversions that do less rearrangement of terms
huffman
parents: 48559
diff changeset
    59
    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
    60
    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
    61
    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
    62
    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
    63
    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
    64
    val conv = conv1 then_conv Conv.rewr_conv rule
48571
d68b74435605 move exception handlers outside of let block
huffman
parents: 48561
diff changeset
    65
  in conv ct end
d68b74435605 move exception handlers outside of let block
huffman
parents: 48561
diff changeset
    66
    handle Cancel => raise CTERM ("no_conversion", [])
9436
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
    67
48560
e0875d956a6b replace Nat_Arith simprocs with simpler conversions that do less rearrangement of terms
huffman
parents: 48559
diff changeset
    68
val cancel_diff_conv = cancel_conv (mk_meta_eq @{thm diff_cancel})
e0875d956a6b replace Nat_Arith simprocs with simpler conversions that do less rearrangement of terms
huffman
parents: 48559
diff changeset
    69
val cancel_eq_conv = cancel_conv (mk_meta_eq @{thm add_left_cancel})
e0875d956a6b replace Nat_Arith simprocs with simpler conversions that do less rearrangement of terms
huffman
parents: 48559
diff changeset
    70
val cancel_le_conv = cancel_conv (mk_meta_eq @{thm add_le_cancel_left})
e0875d956a6b replace Nat_Arith simprocs with simpler conversions that do less rearrangement of terms
huffman
parents: 48559
diff changeset
    71
val cancel_less_conv = cancel_conv (mk_meta_eq @{thm add_less_cancel_left})
24076
ae946f751c44 arith method setup: proper context;
wenzelm
parents: 23881
diff changeset
    72
24095
785c3cd7fcb5 moved lin_arith stuff to Tools/lin_arith.ML;
wenzelm
parents: 24076
diff changeset
    73
end;