author | paulson |
Thu, 15 Mar 2012 16:35:02 +0000 | |
changeset 46953 | 2b6e55924af3 |
parent 46821 | ff6b0c1087f2 |
child 48891 | c0eafbd55de3 |
permissions | -rw-r--r-- |
23146 | 1 |
|
2 |
theory IntArith imports Bin |
|
27237 | 3 |
uses ("int_arith.ML") |
4 |
begin |
|
5 |
||
6 |
||
7 |
(** To simplify inequalities involving integer negation and literals, |
|
8 |
such as -x = #3 |
|
9 |
**) |
|
10 |
||
11 |
lemmas [simp] = |
|
45602 | 12 |
zminus_equation [where y = "integ_of(w)"] |
13 |
equation_zminus [where x = "integ_of(w)"] |
|
14 |
for w |
|
27237 | 15 |
|
16 |
lemmas [iff] = |
|
45602 | 17 |
zminus_zless [where y = "integ_of(w)"] |
18 |
zless_zminus [where x = "integ_of(w)"] |
|
19 |
for w |
|
27237 | 20 |
|
21 |
lemmas [iff] = |
|
45602 | 22 |
zminus_zle [where y = "integ_of(w)"] |
23 |
zle_zminus [where x = "integ_of(w)"] |
|
24 |
for w |
|
27237 | 25 |
|
26 |
lemmas [simp] = |
|
45602 | 27 |
Let_def [where s = "integ_of(w)"] for w |
27237 | 28 |
|
29 |
||
30 |
(*** Simprocs for numeric literals ***) |
|
31 |
||
32 |
(** Combining of literal coefficients in sums of products **) |
|
33 |
||
46821
ff6b0c1087f2
Using mathematical notation for <-> and cardinal arithmetic
paulson
parents:
45602
diff
changeset
|
34 |
lemma zless_iff_zdiff_zless_0: "(x $< y) \<longleftrightarrow> (x$-y $< #0)" |
27237 | 35 |
by (simp add: zcompare_rls) |
36 |
||
46953 | 37 |
lemma eq_iff_zdiff_eq_0: "[| x \<in> int; y \<in> int |] ==> (x = y) \<longleftrightarrow> (x$-y = #0)" |
27237 | 38 |
by (simp add: zcompare_rls) |
39 |
||
46821
ff6b0c1087f2
Using mathematical notation for <-> and cardinal arithmetic
paulson
parents:
45602
diff
changeset
|
40 |
lemma zle_iff_zdiff_zle_0: "(x $<= y) \<longleftrightarrow> (x$-y $<= #0)" |
27237 | 41 |
by (simp add: zcompare_rls) |
42 |
||
43 |
||
44 |
(** For combine_numerals **) |
|
45 |
||
46 |
lemma left_zadd_zmult_distrib: "i$*u $+ (j$*u $+ k) = (i$+j)$*u $+ k" |
|
47 |
by (simp add: zadd_zmult_distrib zadd_ac) |
|
48 |
||
49 |
||
50 |
(** For cancel_numerals **) |
|
51 |
||
52 |
lemmas rel_iff_rel_0_rls = |
|
45602 | 53 |
zless_iff_zdiff_zless_0 [where y = "u $+ v"] |
54 |
eq_iff_zdiff_eq_0 [where y = "u $+ v"] |
|
55 |
zle_iff_zdiff_zle_0 [where y = "u $+ v"] |
|
27237 | 56 |
zless_iff_zdiff_zless_0 [where y = n] |
57 |
eq_iff_zdiff_eq_0 [where y = n] |
|
58 |
zle_iff_zdiff_zle_0 [where y = n] |
|
45602 | 59 |
for u v (* FIXME n (!?) *) |
27237 | 60 |
|
46821
ff6b0c1087f2
Using mathematical notation for <-> and cardinal arithmetic
paulson
parents:
45602
diff
changeset
|
61 |
lemma eq_add_iff1: "(i$*u $+ m = j$*u $+ n) \<longleftrightarrow> ((i$-j)$*u $+ m = intify(n))" |
27237 | 62 |
apply (simp add: zdiff_def zadd_zmult_distrib) |
63 |
apply (simp add: zcompare_rls) |
|
64 |
apply (simp add: zadd_ac) |
|
65 |
done |
|
66 |
||
46821
ff6b0c1087f2
Using mathematical notation for <-> and cardinal arithmetic
paulson
parents:
45602
diff
changeset
|
67 |
lemma eq_add_iff2: "(i$*u $+ m = j$*u $+ n) \<longleftrightarrow> (intify(m) = (j$-i)$*u $+ n)" |
27237 | 68 |
apply (simp add: zdiff_def zadd_zmult_distrib) |
69 |
apply (simp add: zcompare_rls) |
|
70 |
apply (simp add: zadd_ac) |
|
71 |
done |
|
72 |
||
46821
ff6b0c1087f2
Using mathematical notation for <-> and cardinal arithmetic
paulson
parents:
45602
diff
changeset
|
73 |
lemma less_add_iff1: "(i$*u $+ m $< j$*u $+ n) \<longleftrightarrow> ((i$-j)$*u $+ m $< n)" |
27237 | 74 |
apply (simp add: zdiff_def zadd_zmult_distrib zadd_ac rel_iff_rel_0_rls) |
75 |
done |
|
76 |
||
46821
ff6b0c1087f2
Using mathematical notation for <-> and cardinal arithmetic
paulson
parents:
45602
diff
changeset
|
77 |
lemma less_add_iff2: "(i$*u $+ m $< j$*u $+ n) \<longleftrightarrow> (m $< (j$-i)$*u $+ n)" |
27237 | 78 |
apply (simp add: zdiff_def zadd_zmult_distrib zadd_ac rel_iff_rel_0_rls) |
79 |
done |
|
80 |
||
46821
ff6b0c1087f2
Using mathematical notation for <-> and cardinal arithmetic
paulson
parents:
45602
diff
changeset
|
81 |
lemma le_add_iff1: "(i$*u $+ m $<= j$*u $+ n) \<longleftrightarrow> ((i$-j)$*u $+ m $<= n)" |
27237 | 82 |
apply (simp add: zdiff_def zadd_zmult_distrib) |
83 |
apply (simp add: zcompare_rls) |
|
84 |
apply (simp add: zadd_ac) |
|
85 |
done |
|
86 |
||
46821
ff6b0c1087f2
Using mathematical notation for <-> and cardinal arithmetic
paulson
parents:
45602
diff
changeset
|
87 |
lemma le_add_iff2: "(i$*u $+ m $<= j$*u $+ n) \<longleftrightarrow> (m $<= (j$-i)$*u $+ n)" |
27237 | 88 |
apply (simp add: zdiff_def zadd_zmult_distrib) |
89 |
apply (simp add: zcompare_rls) |
|
90 |
apply (simp add: zadd_ac) |
|
91 |
done |
|
92 |
||
93 |
use "int_arith.ML" |
|
23146 | 94 |
|
95 |
end |