src/HOL/Tools/int_arith.ML
author haftmann
Fri, 08 May 2009 09:48:07 +0200
changeset 31068 f591144b0f17
parent 31024 0fdf666e08bf
child 31082 54a442b2d727
permissions -rw-r--r--
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
31068
f591144b0f17 modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents: 31024
diff changeset
     1
(* Author: Tobias Nipkow
30802
f9e9e800d27e simplify theorem references
huffman
parents: 30732
diff changeset
     2
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
     3
Instantiation of the generic linear arithmetic package for int.
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
     4
*)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
     5
31068
f591144b0f17 modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents: 31024
diff changeset
     6
signature INT_ARITH =
f591144b0f17 modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents: 31024
diff changeset
     7
sig
f591144b0f17 modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents: 31024
diff changeset
     8
  val fast_int_arith_simproc: simproc
f591144b0f17 modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents: 31024
diff changeset
     9
  val setup: Context.generic -> Context.generic
f591144b0f17 modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents: 31024
diff changeset
    10
end
f591144b0f17 modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents: 31024
diff changeset
    11
f591144b0f17 modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents: 31024
diff changeset
    12
structure Int_Arith : INT_ARITH =
30496
7cdcc9dd95cb vague cleanup in arith proof tools setup: deleted dead code, more proper structures, clearer arrangement
haftmann
parents: 29269
diff changeset
    13
struct
7cdcc9dd95cb vague cleanup in arith proof tools setup: deleted dead code, more proper structures, clearer arrangement
haftmann
parents: 29269
diff changeset
    14
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    15
(* Update parameters of arithmetic prover *)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    16
24266
bdb48fd8fbdd extended linear arith capabilities with code by Amine
nipkow
parents: 24196
diff changeset
    17
(* reduce contradictory =/</<= to False *)
bdb48fd8fbdd extended linear arith capabilities with code by Amine
nipkow
parents: 24196
diff changeset
    18
bdb48fd8fbdd extended linear arith capabilities with code by Amine
nipkow
parents: 24196
diff changeset
    19
(* Evaluation of terms of the form "m R n" where R is one of "=", "<=" or "<",
bdb48fd8fbdd extended linear arith capabilities with code by Amine
nipkow
parents: 24196
diff changeset
    20
   and m and n are ground terms over rings (roughly speaking).
bdb48fd8fbdd extended linear arith capabilities with code by Amine
nipkow
parents: 24196
diff changeset
    21
   That is, m and n consist only of 1s combined with "+", "-" and "*".
bdb48fd8fbdd extended linear arith capabilities with code by Amine
nipkow
parents: 24196
diff changeset
    22
*)
30496
7cdcc9dd95cb vague cleanup in arith proof tools setup: deleted dead code, more proper structures, clearer arrangement
haftmann
parents: 29269
diff changeset
    23
24266
bdb48fd8fbdd extended linear arith capabilities with code by Amine
nipkow
parents: 24196
diff changeset
    24
val zeroth = (symmetric o mk_meta_eq) @{thm of_int_0};
30496
7cdcc9dd95cb vague cleanup in arith proof tools setup: deleted dead code, more proper structures, clearer arrangement
haftmann
parents: 29269
diff changeset
    25
24266
bdb48fd8fbdd extended linear arith capabilities with code by Amine
nipkow
parents: 24196
diff changeset
    26
val lhss0 = [@{cpat "0::?'a::ring"}];
30496
7cdcc9dd95cb vague cleanup in arith proof tools setup: deleted dead code, more proper structures, clearer arrangement
haftmann
parents: 29269
diff changeset
    27
24266
bdb48fd8fbdd extended linear arith capabilities with code by Amine
nipkow
parents: 24196
diff changeset
    28
fun proc0 phi ss ct =
bdb48fd8fbdd extended linear arith capabilities with code by Amine
nipkow
parents: 24196
diff changeset
    29
  let val T = ctyp_of_term ct
bdb48fd8fbdd extended linear arith capabilities with code by Amine
nipkow
parents: 24196
diff changeset
    30
  in if typ_of T = @{typ int} then NONE else
bdb48fd8fbdd extended linear arith capabilities with code by Amine
nipkow
parents: 24196
diff changeset
    31
     SOME (instantiate' [SOME T] [] zeroth)
bdb48fd8fbdd extended linear arith capabilities with code by Amine
nipkow
parents: 24196
diff changeset
    32
  end;
30496
7cdcc9dd95cb vague cleanup in arith proof tools setup: deleted dead code, more proper structures, clearer arrangement
haftmann
parents: 29269
diff changeset
    33
24266
bdb48fd8fbdd extended linear arith capabilities with code by Amine
nipkow
parents: 24196
diff changeset
    34
val zero_to_of_int_zero_simproc =
bdb48fd8fbdd extended linear arith capabilities with code by Amine
nipkow
parents: 24196
diff changeset
    35
  make_simproc {lhss = lhss0, name = "zero_to_of_int_zero_simproc",
bdb48fd8fbdd extended linear arith capabilities with code by Amine
nipkow
parents: 24196
diff changeset
    36
  proc = proc0, identifier = []};
bdb48fd8fbdd extended linear arith capabilities with code by Amine
nipkow
parents: 24196
diff changeset
    37
bdb48fd8fbdd extended linear arith capabilities with code by Amine
nipkow
parents: 24196
diff changeset
    38
val oneth = (symmetric o mk_meta_eq) @{thm of_int_1};
30496
7cdcc9dd95cb vague cleanup in arith proof tools setup: deleted dead code, more proper structures, clearer arrangement
haftmann
parents: 29269
diff changeset
    39
24266
bdb48fd8fbdd extended linear arith capabilities with code by Amine
nipkow
parents: 24196
diff changeset
    40
val lhss1 = [@{cpat "1::?'a::ring_1"}];
30496
7cdcc9dd95cb vague cleanup in arith proof tools setup: deleted dead code, more proper structures, clearer arrangement
haftmann
parents: 29269
diff changeset
    41
24266
bdb48fd8fbdd extended linear arith capabilities with code by Amine
nipkow
parents: 24196
diff changeset
    42
fun proc1 phi ss ct =
bdb48fd8fbdd extended linear arith capabilities with code by Amine
nipkow
parents: 24196
diff changeset
    43
  let val T = ctyp_of_term ct
bdb48fd8fbdd extended linear arith capabilities with code by Amine
nipkow
parents: 24196
diff changeset
    44
  in if typ_of T = @{typ int} then NONE else
bdb48fd8fbdd extended linear arith capabilities with code by Amine
nipkow
parents: 24196
diff changeset
    45
     SOME (instantiate' [SOME T] [] oneth)
bdb48fd8fbdd extended linear arith capabilities with code by Amine
nipkow
parents: 24196
diff changeset
    46
  end;
30496
7cdcc9dd95cb vague cleanup in arith proof tools setup: deleted dead code, more proper structures, clearer arrangement
haftmann
parents: 29269
diff changeset
    47
24266
bdb48fd8fbdd extended linear arith capabilities with code by Amine
nipkow
parents: 24196
diff changeset
    48
val one_to_of_int_one_simproc =
bdb48fd8fbdd extended linear arith capabilities with code by Amine
nipkow
parents: 24196
diff changeset
    49
  make_simproc {lhss = lhss1, name = "one_to_of_int_one_simproc",
bdb48fd8fbdd extended linear arith capabilities with code by Amine
nipkow
parents: 24196
diff changeset
    50
  proc = proc1, identifier = []};
bdb48fd8fbdd extended linear arith capabilities with code by Amine
nipkow
parents: 24196
diff changeset
    51
bdb48fd8fbdd extended linear arith capabilities with code by Amine
nipkow
parents: 24196
diff changeset
    52
val allowed_consts =
bdb48fd8fbdd extended linear arith capabilities with code by Amine
nipkow
parents: 24196
diff changeset
    53
  [@{const_name "op ="}, @{const_name "HOL.times"}, @{const_name "HOL.uminus"},
bdb48fd8fbdd extended linear arith capabilities with code by Amine
nipkow
parents: 24196
diff changeset
    54
   @{const_name "HOL.minus"}, @{const_name "HOL.plus"},
bdb48fd8fbdd extended linear arith capabilities with code by Amine
nipkow
parents: 24196
diff changeset
    55
   @{const_name "HOL.zero"}, @{const_name "HOL.one"}, @{const_name "HOL.less"},
bdb48fd8fbdd extended linear arith capabilities with code by Amine
nipkow
parents: 24196
diff changeset
    56
   @{const_name "HOL.less_eq"}];
bdb48fd8fbdd extended linear arith capabilities with code by Amine
nipkow
parents: 24196
diff changeset
    57
bdb48fd8fbdd extended linear arith capabilities with code by Amine
nipkow
parents: 24196
diff changeset
    58
fun check t = case t of
bdb48fd8fbdd extended linear arith capabilities with code by Amine
nipkow
parents: 24196
diff changeset
    59
   Const(s,t) => if s = @{const_name "HOL.one"} then not (t = @{typ int})
bdb48fd8fbdd extended linear arith capabilities with code by Amine
nipkow
parents: 24196
diff changeset
    60
                else s mem_string allowed_consts
bdb48fd8fbdd extended linear arith capabilities with code by Amine
nipkow
parents: 24196
diff changeset
    61
 | a$b => check a andalso check b
bdb48fd8fbdd extended linear arith capabilities with code by Amine
nipkow
parents: 24196
diff changeset
    62
 | _ => false;
bdb48fd8fbdd extended linear arith capabilities with code by Amine
nipkow
parents: 24196
diff changeset
    63
bdb48fd8fbdd extended linear arith capabilities with code by Amine
nipkow
parents: 24196
diff changeset
    64
val conv =
bdb48fd8fbdd extended linear arith capabilities with code by Amine
nipkow
parents: 24196
diff changeset
    65
  Simplifier.rewrite
bdb48fd8fbdd extended linear arith capabilities with code by Amine
nipkow
parents: 24196
diff changeset
    66
   (HOL_basic_ss addsimps
bdb48fd8fbdd extended linear arith capabilities with code by Amine
nipkow
parents: 24196
diff changeset
    67
     ((map (fn th => th RS sym) [@{thm of_int_add}, @{thm of_int_mult},
bdb48fd8fbdd extended linear arith capabilities with code by Amine
nipkow
parents: 24196
diff changeset
    68
             @{thm of_int_diff},  @{thm of_int_minus}])@
bdb48fd8fbdd extended linear arith capabilities with code by Amine
nipkow
parents: 24196
diff changeset
    69
      [@{thm of_int_less_iff}, @{thm of_int_le_iff}, @{thm of_int_eq_iff}])
bdb48fd8fbdd extended linear arith capabilities with code by Amine
nipkow
parents: 24196
diff changeset
    70
     addsimprocs [zero_to_of_int_zero_simproc,one_to_of_int_one_simproc]);
bdb48fd8fbdd extended linear arith capabilities with code by Amine
nipkow
parents: 24196
diff changeset
    71
bdb48fd8fbdd extended linear arith capabilities with code by Amine
nipkow
parents: 24196
diff changeset
    72
fun sproc phi ss ct = if check (term_of ct) then SOME (conv ct) else NONE
30496
7cdcc9dd95cb vague cleanup in arith proof tools setup: deleted dead code, more proper structures, clearer arrangement
haftmann
parents: 29269
diff changeset
    73
24266
bdb48fd8fbdd extended linear arith capabilities with code by Amine
nipkow
parents: 24196
diff changeset
    74
val lhss' =
bdb48fd8fbdd extended linear arith capabilities with code by Amine
nipkow
parents: 24196
diff changeset
    75
  [@{cpat "(?x::?'a::ring_char_0) = (?y::?'a)"},
bdb48fd8fbdd extended linear arith capabilities with code by Amine
nipkow
parents: 24196
diff changeset
    76
   @{cpat "(?x::?'a::ordered_idom) < (?y::?'a)"},
bdb48fd8fbdd extended linear arith capabilities with code by Amine
nipkow
parents: 24196
diff changeset
    77
   @{cpat "(?x::?'a::ordered_idom) <= (?y::?'a)"}]
30496
7cdcc9dd95cb vague cleanup in arith proof tools setup: deleted dead code, more proper structures, clearer arrangement
haftmann
parents: 29269
diff changeset
    78
24266
bdb48fd8fbdd extended linear arith capabilities with code by Amine
nipkow
parents: 24196
diff changeset
    79
val zero_one_idom_simproc =
bdb48fd8fbdd extended linear arith capabilities with code by Amine
nipkow
parents: 24196
diff changeset
    80
  make_simproc {lhss = lhss' , name = "zero_one_idom_simproc",
bdb48fd8fbdd extended linear arith capabilities with code by Amine
nipkow
parents: 24196
diff changeset
    81
  proc = sproc, identifier = []}
bdb48fd8fbdd extended linear arith capabilities with code by Amine
nipkow
parents: 24196
diff changeset
    82
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    83
val add_rules =
25481
aa16cd919dcc dropped legacy ml bindings
haftmann
parents: 24630
diff changeset
    84
    simp_thms @ @{thms arith_simps} @ @{thms rel_simps} @ @{thms arith_special} @
30802
f9e9e800d27e simplify theorem references
huffman
parents: 30732
diff changeset
    85
    @{thms int_arith_rules}
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    86
23365
f31794033ae1 removed constant int :: nat => int;
huffman
parents: 23164
diff changeset
    87
val nat_inj_thms = [@{thm zle_int} RS iffD2, @{thm int_int_eq} RS iffD2]
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    88
31068
f591144b0f17 modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents: 31024
diff changeset
    89
val numeral_base_simprocs = Numeral_Simprocs.assoc_fold_simproc :: zero_one_idom_simproc
f591144b0f17 modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents: 31024
diff changeset
    90
  :: Numeral_Simprocs.combine_numerals
f591144b0f17 modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents: 31024
diff changeset
    91
  :: Numeral_Simprocs.cancel_numerals;
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    92
30496
7cdcc9dd95cb vague cleanup in arith proof tools setup: deleted dead code, more proper structures, clearer arrangement
haftmann
parents: 29269
diff changeset
    93
val setup =
30685
dd5fe091ff04 structure LinArith now named Lin_Arith
haftmann
parents: 30518
diff changeset
    94
  Lin_Arith.map_data (fn {add_mono_thms, mult_mono_thms, inj_thms, lessD, neqE, simpset} =>
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    95
   {add_mono_thms = add_mono_thms,
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    96
    mult_mono_thms = @{thm mult_strict_left_mono} :: @{thm mult_left_mono} :: mult_mono_thms,
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    97
    inj_thms = nat_inj_thms @ inj_thms,
25481
aa16cd919dcc dropped legacy ml bindings
haftmann
parents: 24630
diff changeset
    98
    lessD = lessD @ [@{thm zless_imp_add1_zle}],
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    99
    neqE = neqE,
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   100
    simpset = simpset addsimps add_rules
31068
f591144b0f17 modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents: 31024
diff changeset
   101
                      addsimprocs numeral_base_simprocs
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   102
                      addcongs [if_weak_cong]}) #>
24196
f1dbfd7e3223 localized of_nat
haftmann
parents: 24093
diff changeset
   103
  arith_inj_const (@{const_name of_nat}, HOLogic.natT --> HOLogic.intT) #>
25919
8b1c0d434824 joined theories IntDef, Numeral, IntArith to theory Int
haftmann
parents: 25481
diff changeset
   104
  arith_discrete @{type_name Int.int}
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   105
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   106
val fast_int_arith_simproc =
28262
aa7ca36d67fd back to dynamic the_context(), because static @{theory} is invalidated if ML environment changes within the same code block;
wenzelm
parents: 26086
diff changeset
   107
  Simplifier.simproc (the_context ())
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   108
  "fast_int_arith" 
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   109
     ["(m::'a::{ordered_idom,number_ring}) < n",
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   110
      "(m::'a::{ordered_idom,number_ring}) <= n",
30685
dd5fe091ff04 structure LinArith now named Lin_Arith
haftmann
parents: 30518
diff changeset
   111
      "(m::'a::{ordered_idom,number_ring}) = n"] (K Lin_Arith.lin_arith_simproc);
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
   112
30496
7cdcc9dd95cb vague cleanup in arith proof tools setup: deleted dead code, more proper structures, clearer arrangement
haftmann
parents: 29269
diff changeset
   113
end;
7cdcc9dd95cb vague cleanup in arith proof tools setup: deleted dead code, more proper structures, clearer arrangement
haftmann
parents: 29269
diff changeset
   114
7cdcc9dd95cb vague cleanup in arith proof tools setup: deleted dead code, more proper structures, clearer arrangement
haftmann
parents: 29269
diff changeset
   115
Addsimprocs [Int_Arith.fast_int_arith_simproc];