File ‹Tools/int_arith.ML›
signature INT_ARITH =
sig
val zero_one_idom_simproc: simproc
end
structure Int_Arith : INT_ARITH =
struct
val zeroth = Thm.symmetric (mk_meta_eq @{thm of_int_0});
val zero_to_of_int_zero_simproc =
\<^simproc_setup>‹passive zero_to_of_int_zero ("0::'a::ring") =
‹fn _ => fn _ => fn ct =>
let val T = Thm.ctyp_of_cterm ct in
if Thm.typ_of T = \<^typ>‹int› then NONE
else SOME (Thm.instantiate' [SOME T] [] zeroth)
end››;
val oneth = Thm.symmetric (mk_meta_eq @{thm of_int_1});
val one_to_of_int_one_simproc =
\<^simproc_setup>‹passive one_to_of_int_one ("1::'a::ring_1") =
‹fn _ => fn _ => fn ct =>
let val T = Thm.ctyp_of_cterm ct in
if Thm.typ_of T = \<^typ>‹int› then NONE
else SOME (Thm.instantiate' [SOME T] [] oneth)
end››;
fun check (Const (\<^const_name>‹Groups.one›, \<^typ>‹int›)) = false
| check (Const (\<^const_name>‹Groups.one›, _)) = true
| check (Const (s, _)) = member (op =) [\<^const_name>‹HOL.eq›,
\<^const_name>‹Groups.times›, \<^const_name>‹Groups.uminus›,
\<^const_name>‹Groups.minus›, \<^const_name>‹Groups.plus›,
\<^const_name>‹Groups.zero›,
\<^const_name>‹Orderings.less›, \<^const_name>‹Orderings.less_eq›] s
| check (a $ b) = check a andalso check b
| check _ = false;
val conv_ss =
\<^context>
|> put_simpset HOL_basic_ss
|> fold (Simplifier.add_simp o (fn th => th RS sym))
@{thms of_int_add of_int_mult
of_int_diff of_int_minus of_int_less_iff
of_int_le_iff of_int_eq_iff}
|> (fn ss => ss addsimprocs [zero_to_of_int_zero_simproc, one_to_of_int_one_simproc])
|> simpset_of;
val zero_one_idom_simproc =
\<^simproc_setup>‹passive zero_one_idom
("(x::'a::ring_char_0) = y" | "(u::'b::linordered_idom) < v" | "(u::'b::linordered_idom) ≤ v") =
‹fn _ => fn ctxt => fn ct =>
if check (Thm.term_of ct)
then SOME (Simplifier.rewrite (put_simpset conv_ss ctxt) ct)
else NONE››
end;