src/HOL/Tools/int_arith.ML
author wenzelm
Sat, 30 Nov 2024 16:01:58 +0100
changeset 81513 d11ed1bf0ad2
parent 80721 ac39d932ddfc
permissions -rw-r--r--
tuned signature: more operations;
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
70356
4a327c061870 streamlined setup for linear algebra, particularly removed redundant rule declarations
haftmann
parents: 69593
diff changeset
     3
A special simproc for the instantiation of the generic linear arithmetic package for int.
23164
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
70356
4a327c061870 streamlined setup for linear algebra, particularly removed redundant rule declarations
haftmann
parents: 69593
diff changeset
     8
  val zero_one_idom_simproc: simproc
31068
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
24266
bdb48fd8fbdd extended linear arith capabilities with code by Amine
nipkow
parents: 24196
diff changeset
    23
val zero_to_of_int_zero_simproc =
78808
64973b03b778 more standard simproc_setup using ML antiquotation;
wenzelm
parents: 70356
diff changeset
    24
  \<^simproc_setup>\<open>passive zero_to_of_int_zero ("0::'a::ring") =
64973b03b778 more standard simproc_setup using ML antiquotation;
wenzelm
parents: 70356
diff changeset
    25
    \<open>fn _ => fn _ => fn ct =>
61144
5e94dfead1c2 simplified simproc programming interfaces;
wenzelm
parents: 60801
diff changeset
    26
      let val T = Thm.ctyp_of_cterm ct in
80721
ac39d932ddfc tuned: more antiquotations;
wenzelm
parents: 80701
diff changeset
    27
        if Thm.typ_of T = \<^Type>\<open>int\<close> then NONE
ac39d932ddfc tuned: more antiquotations;
wenzelm
parents: 80701
diff changeset
    28
        else SOME \<^instantiate>\<open>'a = T in lemma "0 \<equiv> of_int 0" by simp\<close>
78808
64973b03b778 more standard simproc_setup using ML antiquotation;
wenzelm
parents: 70356
diff changeset
    29
      end\<close>\<close>;
30496
7cdcc9dd95cb vague cleanup in arith proof tools setup: deleted dead code, more proper structures, clearer arrangement
haftmann
parents: 29269
diff changeset
    30
24266
bdb48fd8fbdd extended linear arith capabilities with code by Amine
nipkow
parents: 24196
diff changeset
    31
val one_to_of_int_one_simproc =
78808
64973b03b778 more standard simproc_setup using ML antiquotation;
wenzelm
parents: 70356
diff changeset
    32
  \<^simproc_setup>\<open>passive one_to_of_int_one ("1::'a::ring_1") =
64973b03b778 more standard simproc_setup using ML antiquotation;
wenzelm
parents: 70356
diff changeset
    33
    \<open>fn _ => fn _ => fn ct =>
61144
5e94dfead1c2 simplified simproc programming interfaces;
wenzelm
parents: 60801
diff changeset
    34
      let val T = Thm.ctyp_of_cterm ct in
80721
ac39d932ddfc tuned: more antiquotations;
wenzelm
parents: 80701
diff changeset
    35
        if Thm.typ_of T = \<^Type>\<open>int\<close> then NONE
ac39d932ddfc tuned: more antiquotations;
wenzelm
parents: 80701
diff changeset
    36
        else SOME \<^instantiate>\<open>'a = T in lemma "1 \<equiv> of_int 1" by simp\<close>
78808
64973b03b778 more standard simproc_setup using ML antiquotation;
wenzelm
parents: 70356
diff changeset
    37
      end\<close>\<close>;
24266
bdb48fd8fbdd extended linear arith capabilities with code by Amine
nipkow
parents: 24196
diff changeset
    38
80721
ac39d932ddfc tuned: more antiquotations;
wenzelm
parents: 80701
diff changeset
    39
fun check \<^Const_>\<open>Groups.one \<^Type>\<open>int\<close>\<close> = false
ac39d932ddfc tuned: more antiquotations;
wenzelm
parents: 80701
diff changeset
    40
  | check \<^Const_>\<open>Groups.one _\<close> = true
ac39d932ddfc tuned: more antiquotations;
wenzelm
parents: 80701
diff changeset
    41
  | check \<^Const_>\<open>Groups.zero _\<close> = true
ac39d932ddfc tuned: more antiquotations;
wenzelm
parents: 80701
diff changeset
    42
  | check \<^Const_>\<open>HOL.eq _\<close> = true
ac39d932ddfc tuned: more antiquotations;
wenzelm
parents: 80701
diff changeset
    43
  | check \<^Const_>\<open>times _\<close> = true
ac39d932ddfc tuned: more antiquotations;
wenzelm
parents: 80701
diff changeset
    44
  | check \<^Const_>\<open>uminus _\<close> = true
ac39d932ddfc tuned: more antiquotations;
wenzelm
parents: 80701
diff changeset
    45
  | check \<^Const_>\<open>minus _\<close> = true
ac39d932ddfc tuned: more antiquotations;
wenzelm
parents: 80701
diff changeset
    46
  | check \<^Const_>\<open>plus _\<close> = true
ac39d932ddfc tuned: more antiquotations;
wenzelm
parents: 80701
diff changeset
    47
  | check \<^Const_>\<open>less _\<close> = true
ac39d932ddfc tuned: more antiquotations;
wenzelm
parents: 80701
diff changeset
    48
  | check \<^Const_>\<open>less_eq _\<close> = true
31100
6a2e67fe4488 tuned interface of Lin_Arith
haftmann
parents: 31082
diff changeset
    49
  | check (a $ b) = check a andalso check b
6a2e67fe4488 tuned interface of Lin_Arith
haftmann
parents: 31082
diff changeset
    50
  | check _ = false;
24266
bdb48fd8fbdd extended linear arith capabilities with code by Amine
nipkow
parents: 24196
diff changeset
    51
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 47209
diff changeset
    52
val conv_ss =
70356
4a327c061870 streamlined setup for linear algebra, particularly removed redundant rule declarations
haftmann
parents: 69593
diff changeset
    53
  \<^context>
4a327c061870 streamlined setup for linear algebra, particularly removed redundant rule declarations
haftmann
parents: 69593
diff changeset
    54
  |> put_simpset HOL_basic_ss
4a327c061870 streamlined setup for linear algebra, particularly removed redundant rule declarations
haftmann
parents: 69593
diff changeset
    55
  |> fold (Simplifier.add_simp o (fn th => th RS sym))
4a327c061870 streamlined setup for linear algebra, particularly removed redundant rule declarations
haftmann
parents: 69593
diff changeset
    56
       @{thms of_int_add of_int_mult
4a327c061870 streamlined setup for linear algebra, particularly removed redundant rule declarations
haftmann
parents: 69593
diff changeset
    57
         of_int_diff of_int_minus of_int_less_iff
4a327c061870 streamlined setup for linear algebra, particularly removed redundant rule declarations
haftmann
parents: 69593
diff changeset
    58
         of_int_le_iff of_int_eq_iff}
80701
39cd50407f79 tuned: prefer canonical argument order;
wenzelm
parents: 78808
diff changeset
    59
  |> Simplifier.add_proc zero_to_of_int_zero_simproc
39cd50407f79 tuned: prefer canonical argument order;
wenzelm
parents: 78808
diff changeset
    60
  |> Simplifier.add_proc one_to_of_int_one_simproc
70356
4a327c061870 streamlined setup for linear algebra, particularly removed redundant rule declarations
haftmann
parents: 69593
diff changeset
    61
  |> simpset_of;
24266
bdb48fd8fbdd extended linear arith capabilities with code by Amine
nipkow
parents: 24196
diff changeset
    62
61144
5e94dfead1c2 simplified simproc programming interfaces;
wenzelm
parents: 60801
diff changeset
    63
val zero_one_idom_simproc =
78808
64973b03b778 more standard simproc_setup using ML antiquotation;
wenzelm
parents: 70356
diff changeset
    64
  \<^simproc_setup>\<open>passive zero_one_idom
64973b03b778 more standard simproc_setup using ML antiquotation;
wenzelm
parents: 70356
diff changeset
    65
    ("(x::'a::ring_char_0) = y" | "(u::'b::linordered_idom) < v" | "(u::'b::linordered_idom) \<le> v") =
64973b03b778 more standard simproc_setup using ML antiquotation;
wenzelm
parents: 70356
diff changeset
    66
    \<open>fn _ => fn ctxt => fn ct =>
61144
5e94dfead1c2 simplified simproc programming interfaces;
wenzelm
parents: 60801
diff changeset
    67
      if check (Thm.term_of ct)
5e94dfead1c2 simplified simproc programming interfaces;
wenzelm
parents: 60801
diff changeset
    68
      then SOME (Simplifier.rewrite (put_simpset conv_ss ctxt) ct)
78808
64973b03b778 more standard simproc_setup using ML antiquotation;
wenzelm
parents: 70356
diff changeset
    69
      else NONE\<close>\<close>
30496
7cdcc9dd95cb vague cleanup in arith proof tools setup: deleted dead code, more proper structures, clearer arrangement
haftmann
parents: 29269
diff changeset
    70
31100
6a2e67fe4488 tuned interface of Lin_Arith
haftmann
parents: 31082
diff changeset
    71
end;