| author | blanchet | 
| Fri, 31 Jan 2014 18:43:16 +0100 | |
| changeset 55221 | ee90eebb8b73 | 
| parent 54249 | ce00f2fef556 | 
| child 59582 | 0fbed69ff081 | 
| 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  | 
|
| 36945 | 23  | 
val zeroth = (Thm.symmetric o 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 lhss0 = [@{cpat "0::?'a::ring"}];
 | 
| 
30496
 
7cdcc9dd95cb
vague cleanup in arith proof tools setup: deleted dead code, more proper structures, clearer arrangement
 
haftmann 
parents: 
29269 
diff
changeset
 | 
26  | 
|
| 
51717
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
47209 
diff
changeset
 | 
27  | 
fun proc0 phi ctxt ct =  | 
| 
24266
 
bdb48fd8fbdd
extended linear arith capabilities with code by Amine
 
nipkow 
parents: 
24196 
diff
changeset
 | 
28  | 
let val T = ctyp_of_term ct  | 
| 
 
bdb48fd8fbdd
extended linear arith capabilities with code by Amine
 
nipkow 
parents: 
24196 
diff
changeset
 | 
29  | 
  in if typ_of T = @{typ int} then NONE else
 | 
| 
 
bdb48fd8fbdd
extended linear arith capabilities with code by Amine
 
nipkow 
parents: 
24196 
diff
changeset
 | 
30  | 
SOME (instantiate' [SOME T] [] zeroth)  | 
| 
 
bdb48fd8fbdd
extended linear arith capabilities with code by Amine
 
nipkow 
parents: 
24196 
diff
changeset
 | 
31  | 
end;  | 
| 
30496
 
7cdcc9dd95cb
vague cleanup in arith proof tools setup: deleted dead code, more proper structures, clearer arrangement
 
haftmann 
parents: 
29269 
diff
changeset
 | 
32  | 
|
| 
24266
 
bdb48fd8fbdd
extended linear arith capabilities with code by Amine
 
nipkow 
parents: 
24196 
diff
changeset
 | 
33  | 
val zero_to_of_int_zero_simproc =  | 
| 
 
bdb48fd8fbdd
extended linear arith capabilities with code by Amine
 
nipkow 
parents: 
24196 
diff
changeset
 | 
34  | 
  make_simproc {lhss = lhss0, name = "zero_to_of_int_zero_simproc",
 | 
| 
 
bdb48fd8fbdd
extended linear arith capabilities with code by Amine
 
nipkow 
parents: 
24196 
diff
changeset
 | 
35  | 
proc = proc0, identifier = []};  | 
| 
 
bdb48fd8fbdd
extended linear arith capabilities with code by Amine
 
nipkow 
parents: 
24196 
diff
changeset
 | 
36  | 
|
| 36945 | 37  | 
val oneth = (Thm.symmetric o 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
 | 
38  | 
|
| 
24266
 
bdb48fd8fbdd
extended linear arith capabilities with code by Amine
 
nipkow 
parents: 
24196 
diff
changeset
 | 
39  | 
val lhss1 = [@{cpat "1::?'a::ring_1"}];
 | 
| 
30496
 
7cdcc9dd95cb
vague cleanup in arith proof tools setup: deleted dead code, more proper structures, clearer arrangement
 
haftmann 
parents: 
29269 
diff
changeset
 | 
40  | 
|
| 
51717
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
47209 
diff
changeset
 | 
41  | 
fun proc1 phi ctxt ct =  | 
| 
24266
 
bdb48fd8fbdd
extended linear arith capabilities with code by Amine
 
nipkow 
parents: 
24196 
diff
changeset
 | 
42  | 
let val T = ctyp_of_term ct  | 
| 
 
bdb48fd8fbdd
extended linear arith capabilities with code by Amine
 
nipkow 
parents: 
24196 
diff
changeset
 | 
43  | 
  in if typ_of T = @{typ int} then NONE else
 | 
| 
 
bdb48fd8fbdd
extended linear arith capabilities with code by Amine
 
nipkow 
parents: 
24196 
diff
changeset
 | 
44  | 
SOME (instantiate' [SOME T] [] oneth)  | 
| 
 
bdb48fd8fbdd
extended linear arith capabilities with code by Amine
 
nipkow 
parents: 
24196 
diff
changeset
 | 
45  | 
end;  | 
| 
30496
 
7cdcc9dd95cb
vague cleanup in arith proof tools setup: deleted dead code, more proper structures, clearer arrangement
 
haftmann 
parents: 
29269 
diff
changeset
 | 
46  | 
|
| 
24266
 
bdb48fd8fbdd
extended linear arith capabilities with code by Amine
 
nipkow 
parents: 
24196 
diff
changeset
 | 
47  | 
val one_to_of_int_one_simproc =  | 
| 
 
bdb48fd8fbdd
extended linear arith capabilities with code by Amine
 
nipkow 
parents: 
24196 
diff
changeset
 | 
48  | 
  make_simproc {lhss = lhss1, name = "one_to_of_int_one_simproc",
 | 
| 
 
bdb48fd8fbdd
extended linear arith capabilities with code by Amine
 
nipkow 
parents: 
24196 
diff
changeset
 | 
49  | 
proc = proc1, identifier = []};  | 
| 
 
bdb48fd8fbdd
extended linear arith capabilities with code by Amine
 
nipkow 
parents: 
24196 
diff
changeset
 | 
50  | 
|
| 
35267
 
8dfd816713c6
moved remaning class operations from Algebras.thy to Groups.thy
 
haftmann 
parents: 
35092 
diff
changeset
 | 
51  | 
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
 | 
52  | 
  | check (Const (@{const_name Groups.one}, _)) = true
 | 
| 
38864
 
4abe644fcea5
formerly unnamed infix equality now named HOL.eq
 
haftmann 
parents: 
38763 
diff
changeset
 | 
53  | 
  | 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
 | 
54  | 
      @{const_name Groups.times}, @{const_name Groups.uminus},
 | 
| 
 
8dfd816713c6
moved remaning class operations from Algebras.thy to Groups.thy
 
haftmann 
parents: 
35092 
diff
changeset
 | 
55  | 
      @{const_name Groups.minus}, @{const_name Groups.plus},
 | 
| 
 
8dfd816713c6
moved remaning class operations from Algebras.thy to Groups.thy
 
haftmann 
parents: 
35092 
diff
changeset
 | 
56  | 
      @{const_name Groups.zero},
 | 
| 
35092
 
cfe605c54e50
moved less_eq, less to Orderings.thy; moved abs, sgn to Groups.thy
 
haftmann 
parents: 
35028 
diff
changeset
 | 
57  | 
      @{const_name Orderings.less}, @{const_name Orderings.less_eq}] s
 | 
| 31100 | 58  | 
| check (a $ b) = check a andalso check b  | 
59  | 
| check _ = false;  | 
|
| 
24266
 
bdb48fd8fbdd
extended linear arith capabilities with code by Amine
 
nipkow 
parents: 
24196 
diff
changeset
 | 
60  | 
|
| 
51717
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
47209 
diff
changeset
 | 
61  | 
val conv_ss =  | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
47209 
diff
changeset
 | 
62  | 
  simpset_of (put_simpset HOL_basic_ss @{context}
 | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
47209 
diff
changeset
 | 
63  | 
addsimps  | 
| 
24266
 
bdb48fd8fbdd
extended linear arith capabilities with code by Amine
 
nipkow 
parents: 
24196 
diff
changeset
 | 
64  | 
     ((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
 | 
65  | 
             @{thm of_int_diff},  @{thm of_int_minus}])@
 | 
| 
 
bdb48fd8fbdd
extended linear arith capabilities with code by Amine
 
nipkow 
parents: 
24196 
diff
changeset
 | 
66  | 
      [@{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
 | 
67  | 
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
 | 
68  | 
|
| 
51717
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
47209 
diff
changeset
 | 
69  | 
fun sproc phi ctxt ct =  | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
47209 
diff
changeset
 | 
70  | 
if check (term_of ct) then SOME (Simplifier.rewrite (put_simpset conv_ss ctxt) ct)  | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
47209 
diff
changeset
 | 
71  | 
else NONE;  | 
| 
30496
 
7cdcc9dd95cb
vague cleanup in arith proof tools setup: deleted dead code, more proper structures, clearer arrangement
 
haftmann 
parents: 
29269 
diff
changeset
 | 
72  | 
|
| 
24266
 
bdb48fd8fbdd
extended linear arith capabilities with code by Amine
 
nipkow 
parents: 
24196 
diff
changeset
 | 
73  | 
val lhss' =  | 
| 
 
bdb48fd8fbdd
extended linear arith capabilities with code by Amine
 
nipkow 
parents: 
24196 
diff
changeset
 | 
74  | 
  [@{cpat "(?x::?'a::ring_char_0) = (?y::?'a)"},
 | 
| 
35028
 
108662d50512
more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
 
haftmann 
parents: 
34974 
diff
changeset
 | 
75  | 
   @{cpat "(?x::?'a::linordered_idom) < (?y::?'a)"},
 | 
| 
 
108662d50512
more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
 
haftmann 
parents: 
34974 
diff
changeset
 | 
76  | 
   @{cpat "(?x::?'a::linordered_idom) <= (?y::?'a)"}]
 | 
| 
30496
 
7cdcc9dd95cb
vague cleanup in arith proof tools setup: deleted dead code, more proper structures, clearer arrangement
 
haftmann 
parents: 
29269 
diff
changeset
 | 
77  | 
|
| 
24266
 
bdb48fd8fbdd
extended linear arith capabilities with code by Amine
 
nipkow 
parents: 
24196 
diff
changeset
 | 
78  | 
val zero_one_idom_simproc =  | 
| 
 
bdb48fd8fbdd
extended linear arith capabilities with code by Amine
 
nipkow 
parents: 
24196 
diff
changeset
 | 
79  | 
  make_simproc {lhss = lhss' , name = "zero_one_idom_simproc",
 | 
| 
 
bdb48fd8fbdd
extended linear arith capabilities with code by Amine
 
nipkow 
parents: 
24196 
diff
changeset
 | 
80  | 
proc = sproc, identifier = []}  | 
| 
 
bdb48fd8fbdd
extended linear arith capabilities with code by Amine
 
nipkow 
parents: 
24196 
diff
changeset
 | 
81  | 
|
| 
31510
 
e0f2bb4b0021
fast_lin_arith uses proper multiplication instead of unfolding to additions
 
boehmes 
parents: 
31101 
diff
changeset
 | 
82  | 
fun number_of thy T n =  | 
| 
47108
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
43595 
diff
changeset
 | 
83  | 
  if not (Sign.of_sort thy (T, @{sort numeral}))
 | 
| 
31510
 
e0f2bb4b0021
fast_lin_arith uses proper multiplication instead of unfolding to additions
 
boehmes 
parents: 
31101 
diff
changeset
 | 
84  | 
  then raise CTERM ("number_of", [])
 | 
| 
38763
 
283f1f9969ba
Fast_Lin_Arith.number_of: more conventional merge that prefers the left side -- note that former ordering wrt. serial numbers makes it depend on accidental load order;
 
wenzelm 
parents: 
38715 
diff
changeset
 | 
85  | 
else Numeral.mk_cnumber (Thm.ctyp_of thy T) n;  | 
| 
31510
 
e0f2bb4b0021
fast_lin_arith uses proper multiplication instead of unfolding to additions
 
boehmes 
parents: 
31101 
diff
changeset
 | 
86  | 
|
| 31100 | 87  | 
val setup =  | 
88  | 
  Lin_Arith.add_inj_thms [@{thm zle_int} RS iffD2, @{thm int_int_eq} RS iffD2]
 | 
|
89  | 
  #> Lin_Arith.add_lessD @{thm zless_imp_add1_zle}
 | 
|
| 54249 | 90  | 
  #> Lin_Arith.add_simps @{thms of_nat_simps of_int_simps}
 | 
91  | 
#> Lin_Arith.add_simps  | 
|
92  | 
      [@{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
 | 
93  | 
#> 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
 | 
94  | 
#> Lin_Arith.set_number_of number_of  | 
| 31100 | 95  | 
  #> Lin_Arith.add_inj_const (@{const_name of_nat}, HOLogic.natT --> HOLogic.intT)
 | 
96  | 
  #> Lin_Arith.add_discrete_type @{type_name Int.int}
 | 
|
97  | 
||
98  | 
end;  |