| author | wenzelm | 
| Wed, 07 Aug 2024 15:11:54 +0200 | |
| changeset 80662 | ad9647592a81 | 
| parent 78808 | 64973b03b778 | 
| child 80701 | 39cd50407f79 | 
| permissions | -rw-r--r-- | 
| 31068 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31024diff
changeset | 1 | (* Author: Tobias Nipkow | 
| 30802 | 2 | |
| 70356 
4a327c061870
streamlined setup for linear algebra, particularly removed redundant rule declarations
 haftmann parents: 
69593diff
changeset | 3 | A special simproc for the instantiation of the generic linear arithmetic package for int. | 
| 23164 | 4 | *) | 
| 5 | ||
| 31068 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31024diff
changeset | 6 | signature INT_ARITH = | 
| 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31024diff
changeset | 7 | sig | 
| 70356 
4a327c061870
streamlined setup for linear algebra, particularly removed redundant rule declarations
 haftmann parents: 
69593diff
changeset | 8 | val zero_one_idom_simproc: simproc | 
| 31068 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31024diff
changeset | 9 | end | 
| 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31024diff
changeset | 10 | |
| 
f591144b0f17
modules numeral_simprocs, nat_numeral_simprocs; proper structures for numeral simprocs
 haftmann parents: 
31024diff
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: 
29269diff
changeset | 12 | struct | 
| 
7cdcc9dd95cb
vague cleanup in arith proof tools setup: deleted dead code, more proper structures, clearer arrangement
 haftmann parents: 
29269diff
changeset | 13 | |
| 23164 | 14 | (* Update parameters of arithmetic prover *) | 
| 15 | ||
| 24266 
bdb48fd8fbdd
extended linear arith capabilities with code by Amine
 nipkow parents: 
24196diff
changeset | 16 | (* reduce contradictory =/</<= to False *) | 
| 
bdb48fd8fbdd
extended linear arith capabilities with code by Amine
 nipkow parents: 
24196diff
changeset | 17 | |
| 
bdb48fd8fbdd
extended linear arith capabilities with code by Amine
 nipkow parents: 
24196diff
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: 
24196diff
changeset | 19 | and m and n are ground terms over rings (roughly speaking). | 
| 
bdb48fd8fbdd
extended linear arith capabilities with code by Amine
 nipkow parents: 
24196diff
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: 
24196diff
changeset | 21 | *) | 
| 30496 
7cdcc9dd95cb
vague cleanup in arith proof tools setup: deleted dead code, more proper structures, clearer arrangement
 haftmann parents: 
29269diff
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: 
29269diff
changeset | 24 | |
| 24266 
bdb48fd8fbdd
extended linear arith capabilities with code by Amine
 nipkow parents: 
24196diff
changeset | 25 | val zero_to_of_int_zero_simproc = | 
| 78808 
64973b03b778
more standard simproc_setup using ML antiquotation;
 wenzelm parents: 
70356diff
changeset | 26 |   \<^simproc_setup>\<open>passive zero_to_of_int_zero ("0::'a::ring") =
 | 
| 
64973b03b778
more standard simproc_setup using ML antiquotation;
 wenzelm parents: 
70356diff
changeset | 27 | \<open>fn _ => fn _ => fn ct => | 
| 61144 | 28 | let val T = Thm.ctyp_of_cterm ct in | 
| 67149 | 29 | if Thm.typ_of T = \<^typ>\<open>int\<close> then NONE | 
| 61144 | 30 | else SOME (Thm.instantiate' [SOME T] [] zeroth) | 
| 78808 
64973b03b778
more standard simproc_setup using ML antiquotation;
 wenzelm parents: 
70356diff
changeset | 31 | end\<close>\<close>; | 
| 30496 
7cdcc9dd95cb
vague cleanup in arith proof tools setup: deleted dead code, more proper structures, clearer arrangement
 haftmann parents: 
29269diff
changeset | 32 | |
| 61144 | 33 | 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: 
29269diff
changeset | 34 | |
| 24266 
bdb48fd8fbdd
extended linear arith capabilities with code by Amine
 nipkow parents: 
24196diff
changeset | 35 | val one_to_of_int_one_simproc = | 
| 78808 
64973b03b778
more standard simproc_setup using ML antiquotation;
 wenzelm parents: 
70356diff
changeset | 36 |   \<^simproc_setup>\<open>passive one_to_of_int_one ("1::'a::ring_1") =
 | 
| 
64973b03b778
more standard simproc_setup using ML antiquotation;
 wenzelm parents: 
70356diff
changeset | 37 | \<open>fn _ => fn _ => fn ct => | 
| 61144 | 38 | let val T = Thm.ctyp_of_cterm ct in | 
| 67149 | 39 | if Thm.typ_of T = \<^typ>\<open>int\<close> then NONE | 
| 61144 | 40 | else SOME (Thm.instantiate' [SOME T] [] oneth) | 
| 78808 
64973b03b778
more standard simproc_setup using ML antiquotation;
 wenzelm parents: 
70356diff
changeset | 41 | end\<close>\<close>; | 
| 24266 
bdb48fd8fbdd
extended linear arith capabilities with code by Amine
 nipkow parents: 
24196diff
changeset | 42 | |
| 67149 | 43 | fun check (Const (\<^const_name>\<open>Groups.one\<close>, \<^typ>\<open>int\<close>)) = false | 
| 44 | | check (Const (\<^const_name>\<open>Groups.one\<close>, _)) = true | |
| 45 | | check (Const (s, _)) = member (op =) [\<^const_name>\<open>HOL.eq\<close>, | |
| 46 | \<^const_name>\<open>Groups.times\<close>, \<^const_name>\<open>Groups.uminus\<close>, | |
| 47 | \<^const_name>\<open>Groups.minus\<close>, \<^const_name>\<open>Groups.plus\<close>, | |
| 48 | \<^const_name>\<open>Groups.zero\<close>, | |
| 49 | \<^const_name>\<open>Orderings.less\<close>, \<^const_name>\<open>Orderings.less_eq\<close>] s | |
| 31100 | 50 | | check (a $ b) = check a andalso check b | 
| 51 | | check _ = false; | |
| 24266 
bdb48fd8fbdd
extended linear arith capabilities with code by Amine
 nipkow parents: 
24196diff
changeset | 52 | |
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
47209diff
changeset | 53 | val conv_ss = | 
| 70356 
4a327c061870
streamlined setup for linear algebra, particularly removed redundant rule declarations
 haftmann parents: 
69593diff
changeset | 54 | \<^context> | 
| 
4a327c061870
streamlined setup for linear algebra, particularly removed redundant rule declarations
 haftmann parents: 
69593diff
changeset | 55 | |> put_simpset HOL_basic_ss | 
| 
4a327c061870
streamlined setup for linear algebra, particularly removed redundant rule declarations
 haftmann parents: 
69593diff
changeset | 56 | |> fold (Simplifier.add_simp o (fn th => th RS sym)) | 
| 
4a327c061870
streamlined setup for linear algebra, particularly removed redundant rule declarations
 haftmann parents: 
69593diff
changeset | 57 |        @{thms of_int_add of_int_mult
 | 
| 
4a327c061870
streamlined setup for linear algebra, particularly removed redundant rule declarations
 haftmann parents: 
69593diff
changeset | 58 | of_int_diff of_int_minus of_int_less_iff | 
| 
4a327c061870
streamlined setup for linear algebra, particularly removed redundant rule declarations
 haftmann parents: 
69593diff
changeset | 59 | of_int_le_iff of_int_eq_iff} | 
| 
4a327c061870
streamlined setup for linear algebra, particularly removed redundant rule declarations
 haftmann parents: 
69593diff
changeset | 60 | |> (fn ss => ss addsimprocs [zero_to_of_int_zero_simproc, one_to_of_int_one_simproc]) | 
| 
4a327c061870
streamlined setup for linear algebra, particularly removed redundant rule declarations
 haftmann parents: 
69593diff
changeset | 61 | |> simpset_of; | 
| 24266 
bdb48fd8fbdd
extended linear arith capabilities with code by Amine
 nipkow parents: 
24196diff
changeset | 62 | |
| 61144 | 63 | val zero_one_idom_simproc = | 
| 78808 
64973b03b778
more standard simproc_setup using ML antiquotation;
 wenzelm parents: 
70356diff
changeset | 64 | \<^simproc_setup>\<open>passive zero_one_idom | 
| 
64973b03b778
more standard simproc_setup using ML antiquotation;
 wenzelm parents: 
70356diff
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: 
70356diff
changeset | 66 | \<open>fn _ => fn ctxt => fn ct => | 
| 61144 | 67 | if check (Thm.term_of ct) | 
| 68 | then SOME (Simplifier.rewrite (put_simpset conv_ss ctxt) ct) | |
| 78808 
64973b03b778
more standard simproc_setup using ML antiquotation;
 wenzelm parents: 
70356diff
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: 
29269diff
changeset | 70 | |
| 31100 | 71 | end; |