src/HOL/Tools/int_arith.ML
author wenzelm
Sun, 18 Oct 2015 21:30:01 +0200
changeset 61476 1884c40f1539
parent 61144 5e94dfead1c2
child 62913 13252110a6fe
permissions -rw-r--r--
tuned signature;
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 setup: Context.generic -> Context.generic
f591144b0f17 modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents: 31024
diff changeset
     9
end
f591144b0f17 modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents: 31024
diff changeset
    10
f591144b0f17 modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents: 31024
diff changeset
    11
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
    12
struct
7cdcc9dd95cb vague cleanup in arith proof tools setup: deleted dead code, more proper structures, clearer arrangement
haftmann
parents: 29269
diff changeset
    13
23164
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    14
(* Update parameters of arithmetic prover *)
69e55066dbca moved Integ files to canonical place;
wenzelm
parents:
diff changeset
    15
24266
bdb48fd8fbdd extended linear arith capabilities with code by Amine
nipkow
parents: 24196
diff changeset
    16
(* reduce contradictory =/</<= to False *)
bdb48fd8fbdd extended linear arith capabilities with code by Amine
nipkow
parents: 24196
diff changeset
    17
bdb48fd8fbdd extended linear arith capabilities with code by Amine
nipkow
parents: 24196
diff changeset
    18
(* 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
    19
   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
    20
   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
    21
*)
30496
7cdcc9dd95cb vague cleanup in arith proof tools setup: deleted dead code, more proper structures, clearer arrangement
haftmann
parents: 29269
diff changeset
    22
61144
5e94dfead1c2 simplified simproc programming interfaces;
wenzelm
parents: 60801
diff changeset
    23
val zeroth = Thm.symmetric (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
    24
24266
bdb48fd8fbdd extended linear arith capabilities with code by Amine
nipkow
parents: 24196
diff changeset
    25
val zero_to_of_int_zero_simproc =
61144
5e94dfead1c2 simplified simproc programming interfaces;
wenzelm
parents: 60801
diff changeset
    26
  Simplifier.make_simproc @{context} "zero_to_of_int_zero_simproc"
5e94dfead1c2 simplified simproc programming interfaces;
wenzelm
parents: 60801
diff changeset
    27
   {lhss = [@{term "0::'a::ring"}],
5e94dfead1c2 simplified simproc programming interfaces;
wenzelm
parents: 60801
diff changeset
    28
    proc = fn _ => fn ctxt => fn ct =>
5e94dfead1c2 simplified simproc programming interfaces;
wenzelm
parents: 60801
diff changeset
    29
      let val T = Thm.ctyp_of_cterm ct in
5e94dfead1c2 simplified simproc programming interfaces;
wenzelm
parents: 60801
diff changeset
    30
        if Thm.typ_of T = @{typ int} then NONE
5e94dfead1c2 simplified simproc programming interfaces;
wenzelm
parents: 60801
diff changeset
    31
        else SOME (Thm.instantiate' [SOME T] [] zeroth)
5e94dfead1c2 simplified simproc programming interfaces;
wenzelm
parents: 60801
diff changeset
    32
      end,
5e94dfead1c2 simplified simproc programming interfaces;
wenzelm
parents: 60801
diff changeset
    33
     identifier = []};
30496
7cdcc9dd95cb vague cleanup in arith proof tools setup: deleted dead code, more proper structures, clearer arrangement
haftmann
parents: 29269
diff changeset
    34
61144
5e94dfead1c2 simplified simproc programming interfaces;
wenzelm
parents: 60801
diff changeset
    35
val oneth = Thm.symmetric (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
    36
24266
bdb48fd8fbdd extended linear arith capabilities with code by Amine
nipkow
parents: 24196
diff changeset
    37
val one_to_of_int_one_simproc =
61144
5e94dfead1c2 simplified simproc programming interfaces;
wenzelm
parents: 60801
diff changeset
    38
  Simplifier.make_simproc @{context} "one_to_of_int_one_simproc"
5e94dfead1c2 simplified simproc programming interfaces;
wenzelm
parents: 60801
diff changeset
    39
   {lhss = [@{term "1::'a::ring_1"}],
5e94dfead1c2 simplified simproc programming interfaces;
wenzelm
parents: 60801
diff changeset
    40
    proc = fn _ => fn ctxt => fn ct =>
5e94dfead1c2 simplified simproc programming interfaces;
wenzelm
parents: 60801
diff changeset
    41
      let val T = Thm.ctyp_of_cterm ct in
5e94dfead1c2 simplified simproc programming interfaces;
wenzelm
parents: 60801
diff changeset
    42
        if Thm.typ_of T = @{typ int} then NONE
5e94dfead1c2 simplified simproc programming interfaces;
wenzelm
parents: 60801
diff changeset
    43
        else SOME (Thm.instantiate' [SOME T] [] oneth)
5e94dfead1c2 simplified simproc programming interfaces;
wenzelm
parents: 60801
diff changeset
    44
      end,
5e94dfead1c2 simplified simproc programming interfaces;
wenzelm
parents: 60801
diff changeset
    45
    identifier = []};
24266
bdb48fd8fbdd extended linear arith capabilities with code by Amine
nipkow
parents: 24196
diff changeset
    46
35267
8dfd816713c6 moved remaning class operations from Algebras.thy to Groups.thy
haftmann
parents: 35092
diff changeset
    47
fun check (Const (@{const_name Groups.one}, @{typ int})) = false
8dfd816713c6 moved remaning class operations from Algebras.thy to Groups.thy
haftmann
parents: 35092
diff changeset
    48
  | check (Const (@{const_name Groups.one}, _)) = true
38864
4abe644fcea5 formerly unnamed infix equality now named HOL.eq
haftmann
parents: 38763
diff changeset
    49
  | check (Const (s, _)) = member (op =) [@{const_name HOL.eq},
35267
8dfd816713c6 moved remaning class operations from Algebras.thy to Groups.thy
haftmann
parents: 35092
diff changeset
    50
      @{const_name Groups.times}, @{const_name Groups.uminus},
8dfd816713c6 moved remaning class operations from Algebras.thy to Groups.thy
haftmann
parents: 35092
diff changeset
    51
      @{const_name Groups.minus}, @{const_name Groups.plus},
8dfd816713c6 moved remaning class operations from Algebras.thy to Groups.thy
haftmann
parents: 35092
diff changeset
    52
      @{const_name Groups.zero},
35092
cfe605c54e50 moved less_eq, less to Orderings.thy; moved abs, sgn to Groups.thy
haftmann
parents: 35028
diff changeset
    53
      @{const_name Orderings.less}, @{const_name Orderings.less_eq}] s
31100
6a2e67fe4488 tuned interface of Lin_Arith
haftmann
parents: 31082
diff changeset
    54
  | check (a $ b) = check a andalso check b
6a2e67fe4488 tuned interface of Lin_Arith
haftmann
parents: 31082
diff changeset
    55
  | check _ = false;
24266
bdb48fd8fbdd extended linear arith capabilities with code by Amine
nipkow
parents: 24196
diff changeset
    56
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 47209
diff changeset
    57
val conv_ss =
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 47209
diff changeset
    58
  simpset_of (put_simpset HOL_basic_ss @{context}
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 47209
diff changeset
    59
    addsimps
24266
bdb48fd8fbdd extended linear arith capabilities with code by Amine
nipkow
parents: 24196
diff changeset
    60
     ((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
    61
             @{thm of_int_diff},  @{thm of_int_minus}])@
bdb48fd8fbdd extended linear arith capabilities with code by Amine
nipkow
parents: 24196
diff changeset
    62
      [@{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
    63
     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
    64
61144
5e94dfead1c2 simplified simproc programming interfaces;
wenzelm
parents: 60801
diff changeset
    65
val zero_one_idom_simproc =
5e94dfead1c2 simplified simproc programming interfaces;
wenzelm
parents: 60801
diff changeset
    66
  Simplifier.make_simproc @{context} "zero_one_idom_simproc"
5e94dfead1c2 simplified simproc programming interfaces;
wenzelm
parents: 60801
diff changeset
    67
   {lhss =
5e94dfead1c2 simplified simproc programming interfaces;
wenzelm
parents: 60801
diff changeset
    68
      [@{term "(x::'a::ring_char_0) = y"},
5e94dfead1c2 simplified simproc programming interfaces;
wenzelm
parents: 60801
diff changeset
    69
       @{term "(x::'a::linordered_idom) < y"},
5e94dfead1c2 simplified simproc programming interfaces;
wenzelm
parents: 60801
diff changeset
    70
       @{term "(x::'a::linordered_idom) \<le> y"}],
5e94dfead1c2 simplified simproc programming interfaces;
wenzelm
parents: 60801
diff changeset
    71
    proc = fn _ => fn ctxt => fn ct =>
5e94dfead1c2 simplified simproc programming interfaces;
wenzelm
parents: 60801
diff changeset
    72
      if check (Thm.term_of ct)
5e94dfead1c2 simplified simproc programming interfaces;
wenzelm
parents: 60801
diff changeset
    73
      then SOME (Simplifier.rewrite (put_simpset conv_ss ctxt) ct)
5e94dfead1c2 simplified simproc programming interfaces;
wenzelm
parents: 60801
diff changeset
    74
      else NONE,
5e94dfead1c2 simplified simproc programming interfaces;
wenzelm
parents: 60801
diff changeset
    75
    identifier = []};
30496
7cdcc9dd95cb vague cleanup in arith proof tools setup: deleted dead code, more proper structures, clearer arrangement
haftmann
parents: 29269
diff changeset
    76
24266
bdb48fd8fbdd extended linear arith capabilities with code by Amine
nipkow
parents: 24196
diff changeset
    77
59996
4dca48557921 tuned signature;
wenzelm
parents: 59621
diff changeset
    78
fun number_of ctxt T n =
4dca48557921 tuned signature;
wenzelm
parents: 59621
diff changeset
    79
  if not (Sign.of_sort (Proof_Context.theory_of ctxt) (T, @{sort numeral}))
31510
e0f2bb4b0021 fast_lin_arith uses proper multiplication instead of unfolding to additions
boehmes
parents: 31101
diff changeset
    80
  then raise CTERM ("number_of", [])
59996
4dca48557921 tuned signature;
wenzelm
parents: 59621
diff changeset
    81
  else Numeral.mk_cnumber (Thm.ctyp_of ctxt T) n;
31510
e0f2bb4b0021 fast_lin_arith uses proper multiplication instead of unfolding to additions
boehmes
parents: 31101
diff changeset
    82
31100
6a2e67fe4488 tuned interface of Lin_Arith
haftmann
parents: 31082
diff changeset
    83
val setup =
6a2e67fe4488 tuned interface of Lin_Arith
haftmann
parents: 31082
diff changeset
    84
  Lin_Arith.add_inj_thms [@{thm zle_int} RS iffD2, @{thm int_int_eq} RS iffD2]
6a2e67fe4488 tuned interface of Lin_Arith
haftmann
parents: 31082
diff changeset
    85
  #> Lin_Arith.add_lessD @{thm zless_imp_add1_zle}
54249
ce00f2fef556 streamlined setup of linear arithmetic
haftmann
parents: 51717
diff changeset
    86
  #> Lin_Arith.add_simps @{thms of_nat_simps of_int_simps}
ce00f2fef556 streamlined setup of linear arithmetic
haftmann
parents: 51717
diff changeset
    87
  #> Lin_Arith.add_simps
ce00f2fef556 streamlined setup of linear arithmetic
haftmann
parents: 51717
diff changeset
    88
      [@{thm of_int_numeral}, @{thm nat_0}, @{thm nat_1}, @{thm diff_nat_numeral}, @{thm nat_numeral}]
33296
a3924d1069e5 moved theory Divides after theory Nat_Numeral; tuned some proof texts
haftmann
parents: 32603
diff changeset
    89
  #> Lin_Arith.add_simprocs [zero_one_idom_simproc]
31510
e0f2bb4b0021 fast_lin_arith uses proper multiplication instead of unfolding to additions
boehmes
parents: 31101
diff changeset
    90
  #> Lin_Arith.set_number_of number_of
31100
6a2e67fe4488 tuned interface of Lin_Arith
haftmann
parents: 31082
diff changeset
    91
  #> Lin_Arith.add_inj_const (@{const_name of_nat}, HOLogic.natT --> HOLogic.intT)
6a2e67fe4488 tuned interface of Lin_Arith
haftmann
parents: 31082
diff changeset
    92
  #> Lin_Arith.add_discrete_type @{type_name Int.int}
6a2e67fe4488 tuned interface of Lin_Arith
haftmann
parents: 31082
diff changeset
    93
6a2e67fe4488 tuned interface of Lin_Arith
haftmann
parents: 31082
diff changeset
    94
end;