author | wenzelm |
Sat, 01 Jun 2019 11:29:59 +0200 | |
changeset 70299 | 83774d669b51 |
parent 69593 | 3dda49e08b9d |
child 70356 | 4a327c061870 |
permissions | -rw-r--r-- |
31068
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
haftmann
parents:
31024
diff
changeset
|
1 |
(* Author: Tobias Nipkow |
30802 | 2 |
|
23164 | 3 |
Instantiation of the generic linear arithmetic package for int. |
4 |
*) |
|
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 | 14 |
(* Update parameters of arithmetic prover *) |
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 | 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 = |
69593 | 26 |
Simplifier.make_simproc \<^context> "zero_to_of_int_zero_simproc" |
67149 | 27 |
{lhss = [\<^term>\<open>0::'a::ring\<close>], |
61144 | 28 |
proc = fn _ => fn ctxt => fn ct => |
29 |
let val T = Thm.ctyp_of_cterm ct in |
|
67149 | 30 |
if Thm.typ_of T = \<^typ>\<open>int\<close> then NONE |
61144 | 31 |
else SOME (Thm.instantiate' [SOME T] [] zeroth) |
62913 | 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 |
|
61144 | 34 |
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
|
35 |
|
24266
bdb48fd8fbdd
extended linear arith capabilities with code by Amine
nipkow
parents:
24196
diff
changeset
|
36 |
val one_to_of_int_one_simproc = |
69593 | 37 |
Simplifier.make_simproc \<^context> "one_to_of_int_one_simproc" |
67149 | 38 |
{lhss = [\<^term>\<open>1::'a::ring_1\<close>], |
61144 | 39 |
proc = fn _ => fn ctxt => fn ct => |
40 |
let val T = Thm.ctyp_of_cterm ct in |
|
67149 | 41 |
if Thm.typ_of T = \<^typ>\<open>int\<close> then NONE |
61144 | 42 |
else SOME (Thm.instantiate' [SOME T] [] oneth) |
62913 | 43 |
end}; |
24266
bdb48fd8fbdd
extended linear arith capabilities with code by Amine
nipkow
parents:
24196
diff
changeset
|
44 |
|
67149 | 45 |
fun check (Const (\<^const_name>\<open>Groups.one\<close>, \<^typ>\<open>int\<close>)) = false |
46 |
| check (Const (\<^const_name>\<open>Groups.one\<close>, _)) = true |
|
47 |
| check (Const (s, _)) = member (op =) [\<^const_name>\<open>HOL.eq\<close>, |
|
48 |
\<^const_name>\<open>Groups.times\<close>, \<^const_name>\<open>Groups.uminus\<close>, |
|
49 |
\<^const_name>\<open>Groups.minus\<close>, \<^const_name>\<open>Groups.plus\<close>, |
|
50 |
\<^const_name>\<open>Groups.zero\<close>, |
|
51 |
\<^const_name>\<open>Orderings.less\<close>, \<^const_name>\<open>Orderings.less_eq\<close>] s |
|
31100 | 52 |
| check (a $ b) = check a andalso check b |
53 |
| check _ = false; |
|
24266
bdb48fd8fbdd
extended linear arith capabilities with code by Amine
nipkow
parents:
24196
diff
changeset
|
54 |
|
51717
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
47209
diff
changeset
|
55 |
val conv_ss = |
69593 | 56 |
simpset_of (put_simpset HOL_basic_ss \<^context> |
51717
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
47209
diff
changeset
|
57 |
addsimps |
24266
bdb48fd8fbdd
extended linear arith capabilities with code by Amine
nipkow
parents:
24196
diff
changeset
|
58 |
((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
|
59 |
@{thm of_int_diff}, @{thm of_int_minus}])@ |
bdb48fd8fbdd
extended linear arith capabilities with code by Amine
nipkow
parents:
24196
diff
changeset
|
60 |
[@{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
|
61 |
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
|
62 |
|
61144 | 63 |
val zero_one_idom_simproc = |
69593 | 64 |
Simplifier.make_simproc \<^context> "zero_one_idom_simproc" |
61144 | 65 |
{lhss = |
67149 | 66 |
[\<^term>\<open>(x::'a::ring_char_0) = y\<close>, |
67 |
\<^term>\<open>(x::'a::linordered_idom) < y\<close>, |
|
68 |
\<^term>\<open>(x::'a::linordered_idom) \<le> y\<close>], |
|
61144 | 69 |
proc = fn _ => fn ctxt => fn ct => |
70 |
if check (Thm.term_of ct) |
|
71 |
then SOME (Simplifier.rewrite (put_simpset conv_ss ctxt) ct) |
|
62913 | 72 |
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 |
|
59996 | 75 |
fun number_of ctxt T n = |
67149 | 76 |
if not (Sign.of_sort (Proof_Context.theory_of ctxt) (T, \<^sort>\<open>numeral\<close>)) |
31510
e0f2bb4b0021
fast_lin_arith uses proper multiplication instead of unfolding to additions
boehmes
parents:
31101
diff
changeset
|
77 |
then raise CTERM ("number_of", []) |
59996 | 78 |
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
|
79 |
|
31100 | 80 |
val setup = |
67116 | 81 |
Lin_Arith.add_inj_thms [@{thm of_nat_le_iff} RS iffD2, @{thm of_nat_eq_iff} RS iffD2] |
31100 | 82 |
#> Lin_Arith.add_lessD @{thm zless_imp_add1_zle} |
54249 | 83 |
#> Lin_Arith.add_simps @{thms of_nat_simps of_int_simps} |
84 |
#> Lin_Arith.add_simps |
|
85 |
[@{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
|
86 |
#> 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
|
87 |
#> Lin_Arith.set_number_of number_of |
67149 | 88 |
#> Lin_Arith.add_inj_const (\<^const_name>\<open>of_nat\<close>, HOLogic.natT --> HOLogic.intT) |
89 |
#> Lin_Arith.add_discrete_type \<^type_name>\<open>Int.int\<close> |
|
31100 | 90 |
|
91 |
end; |