author | wenzelm |
Mon, 13 Jun 2022 11:10:39 +0200 | |
changeset 75556 | 1f6fc2416a48 |
parent 74592 | 3c587b7c3d5c |
child 75875 | 48d032035744 |
permissions | -rw-r--r-- |
72281
beeadb35e357
more thorough treatment of division, particularly signed division on int and word
haftmann
parents:
diff
changeset
|
1 |
(* Author: Stefan Berghofer et al. |
beeadb35e357
more thorough treatment of division, particularly signed division on int and word
haftmann
parents:
diff
changeset
|
2 |
*) |
beeadb35e357
more thorough treatment of division, particularly signed division on int and word
haftmann
parents:
diff
changeset
|
3 |
|
beeadb35e357
more thorough treatment of division, particularly signed division on int and word
haftmann
parents:
diff
changeset
|
4 |
subsection \<open>Signed division: negative results rounded towards zero rather than minus infinity.\<close> |
beeadb35e357
more thorough treatment of division, particularly signed division on int and word
haftmann
parents:
diff
changeset
|
5 |
|
beeadb35e357
more thorough treatment of division, particularly signed division on int and word
haftmann
parents:
diff
changeset
|
6 |
theory Signed_Division |
beeadb35e357
more thorough treatment of division, particularly signed division on int and word
haftmann
parents:
diff
changeset
|
7 |
imports Main |
beeadb35e357
more thorough treatment of division, particularly signed division on int and word
haftmann
parents:
diff
changeset
|
8 |
begin |
beeadb35e357
more thorough treatment of division, particularly signed division on int and word
haftmann
parents:
diff
changeset
|
9 |
|
beeadb35e357
more thorough treatment of division, particularly signed division on int and word
haftmann
parents:
diff
changeset
|
10 |
class signed_division = |
beeadb35e357
more thorough treatment of division, particularly signed division on int and word
haftmann
parents:
diff
changeset
|
11 |
fixes signed_divide :: \<open>'a \<Rightarrow> 'a \<Rightarrow> 'a\<close> (infixl "sdiv" 70) |
beeadb35e357
more thorough treatment of division, particularly signed division on int and word
haftmann
parents:
diff
changeset
|
12 |
and signed_modulo :: \<open>'a \<Rightarrow> 'a \<Rightarrow> 'a\<close> (infixl "smod" 70) |
beeadb35e357
more thorough treatment of division, particularly signed division on int and word
haftmann
parents:
diff
changeset
|
13 |
|
beeadb35e357
more thorough treatment of division, particularly signed division on int and word
haftmann
parents:
diff
changeset
|
14 |
instantiation int :: signed_division |
beeadb35e357
more thorough treatment of division, particularly signed division on int and word
haftmann
parents:
diff
changeset
|
15 |
begin |
beeadb35e357
more thorough treatment of division, particularly signed division on int and word
haftmann
parents:
diff
changeset
|
16 |
|
beeadb35e357
more thorough treatment of division, particularly signed division on int and word
haftmann
parents:
diff
changeset
|
17 |
definition signed_divide_int :: \<open>int \<Rightarrow> int \<Rightarrow> int\<close> |
beeadb35e357
more thorough treatment of division, particularly signed division on int and word
haftmann
parents:
diff
changeset
|
18 |
where \<open>k sdiv l = sgn k * sgn l * (\<bar>k\<bar> div \<bar>l\<bar>)\<close> for k l :: int |
beeadb35e357
more thorough treatment of division, particularly signed division on int and word
haftmann
parents:
diff
changeset
|
19 |
|
beeadb35e357
more thorough treatment of division, particularly signed division on int and word
haftmann
parents:
diff
changeset
|
20 |
definition signed_modulo_int :: \<open>int \<Rightarrow> int \<Rightarrow> int\<close> |
beeadb35e357
more thorough treatment of division, particularly signed division on int and word
haftmann
parents:
diff
changeset
|
21 |
where \<open>k smod l = k - (k sdiv l) * l\<close> for k l :: int |
beeadb35e357
more thorough treatment of division, particularly signed division on int and word
haftmann
parents:
diff
changeset
|
22 |
|
beeadb35e357
more thorough treatment of division, particularly signed division on int and word
haftmann
parents:
diff
changeset
|
23 |
instance .. |
beeadb35e357
more thorough treatment of division, particularly signed division on int and word
haftmann
parents:
diff
changeset
|
24 |
|
beeadb35e357
more thorough treatment of division, particularly signed division on int and word
haftmann
parents:
diff
changeset
|
25 |
end |
beeadb35e357
more thorough treatment of division, particularly signed division on int and word
haftmann
parents:
diff
changeset
|
26 |
|
72768 | 27 |
lemma int_sdiv_simps [simp]: |
28 |
"(a :: int) sdiv 1 = a" |
|
29 |
"(a :: int) sdiv 0 = 0" |
|
30 |
"(a :: int) sdiv -1 = -a" |
|
31 |
apply (auto simp: signed_divide_int_def sgn_if) |
|
32 |
done |
|
33 |
||
34 |
lemma sgn_div_eq_sgn_mult: |
|
35 |
"a div b \<noteq> 0 \<Longrightarrow> sgn ((a :: int) div b) = sgn (a * b)" |
|
36 |
apply (clarsimp simp: sgn_if zero_le_mult_iff neg_imp_zdiv_nonneg_iff not_less) |
|
37 |
apply (metis less_le mult_le_0_iff neg_imp_zdiv_neg_iff not_less pos_imp_zdiv_neg_iff zdiv_eq_0_iff) |
|
38 |
done |
|
39 |
||
40 |
lemma sgn_sdiv_eq_sgn_mult: |
|
41 |
"a sdiv b \<noteq> 0 \<Longrightarrow> sgn ((a :: int) sdiv b) = sgn (a * b)" |
|
42 |
by (auto simp: signed_divide_int_def sgn_div_eq_sgn_mult sgn_mult) |
|
43 |
||
44 |
lemma int_sdiv_same_is_1 [simp]: |
|
45 |
"a \<noteq> 0 \<Longrightarrow> ((a :: int) sdiv b = a) = (b = 1)" |
|
46 |
apply (rule iffI) |
|
47 |
apply (clarsimp simp: signed_divide_int_def) |
|
48 |
apply (subgoal_tac "b > 0") |
|
49 |
apply (case_tac "a > 0") |
|
50 |
apply (clarsimp simp: sgn_if) |
|
51 |
apply (simp_all add: not_less algebra_split_simps sgn_if split: if_splits) |
|
52 |
using int_div_less_self [of a b] apply linarith |
|
53 |
apply (metis add.commute add.inverse_inverse group_cancel.rule0 int_div_less_self linorder_neqE_linordered_idom neg_0_le_iff_le not_less verit_comp_simplify1(1) zless_imp_add1_zle) |
|
54 |
apply (metis div_minus_right neg_imp_zdiv_neg_iff neg_le_0_iff_le not_less order.not_eq_order_implies_strict) |
|
55 |
apply (metis abs_le_zero_iff abs_of_nonneg neg_imp_zdiv_nonneg_iff order.not_eq_order_implies_strict) |
|
56 |
done |
|
57 |
||
58 |
lemma int_sdiv_negated_is_minus1 [simp]: |
|
59 |
"a \<noteq> 0 \<Longrightarrow> ((a :: int) sdiv b = - a) = (b = -1)" |
|
60 |
apply (clarsimp simp: signed_divide_int_def) |
|
61 |
apply (rule iffI) |
|
62 |
apply (subgoal_tac "b < 0") |
|
63 |
apply (case_tac "a > 0") |
|
64 |
apply (clarsimp simp: sgn_if algebra_split_simps not_less) |
|
65 |
apply (case_tac "sgn (a * b) = -1") |
|
66 |
apply (simp_all add: not_less algebra_split_simps sgn_if split: if_splits) |
|
67 |
apply (metis add.inverse_inverse int_div_less_self int_one_le_iff_zero_less less_le neg_0_less_iff_less) |
|
68 |
apply (metis add.inverse_inverse div_minus_right int_div_less_self int_one_le_iff_zero_less less_le neg_0_less_iff_less) |
|
69 |
apply (metis less_le neg_less_0_iff_less not_less pos_imp_zdiv_neg_iff) |
|
70 |
apply (metis div_minus_right dual_order.eq_iff neg_imp_zdiv_nonneg_iff neg_less_0_iff_less) |
|
71 |
done |
|
72 |
||
73 |
lemma sdiv_int_range: |
|
74 |
"(a :: int) sdiv b \<in> { - (abs a) .. (abs a) }" |
|
75 |
apply (unfold signed_divide_int_def) |
|
76 |
apply (subgoal_tac "(abs a) div (abs b) \<le> (abs a)") |
|
77 |
apply (auto simp add: sgn_if not_less) |
|
78 |
apply (metis le_less le_less_trans neg_equal_0_iff_equal neg_less_iff_less not_le pos_imp_zdiv_neg_iff) |
|
79 |
apply (metis add.inverse_neutral div_int_pos_iff le_less neg_le_iff_le order_trans) |
|
80 |
apply (metis div_minus_right le_less_trans neg_imp_zdiv_neg_iff neg_less_0_iff_less not_le) |
|
81 |
using div_int_pos_iff apply fastforce |
|
82 |
apply (auto simp add: abs_if not_less) |
|
83 |
apply (metis add.inverse_inverse add_0_left div_by_1 div_minus_right less_le neg_0_le_iff_le not_le not_one_le_zero zdiv_mono2 zless_imp_add1_zle) |
|
84 |
apply (metis div_by_1 neg_0_less_iff_less pos_imp_zdiv_pos_iff zdiv_mono2 zero_less_one) |
|
85 |
apply (metis add.inverse_neutral div_by_0 div_by_1 int_div_less_self int_one_le_iff_zero_less less_le less_minus_iff order_refl) |
|
86 |
apply (metis div_by_1 divide_int_def int_div_less_self less_le linorder_neqE_linordered_idom order_refl unique_euclidean_semiring_numeral_class.div_less) |
|
87 |
done |
|
88 |
||
89 |
lemma sdiv_int_div_0 [simp]: |
|
90 |
"(x :: int) sdiv 0 = 0" |
|
91 |
by (clarsimp simp: signed_divide_int_def) |
|
92 |
||
93 |
lemma sdiv_int_0_div [simp]: |
|
94 |
"0 sdiv (x :: int) = 0" |
|
95 |
by (clarsimp simp: signed_divide_int_def) |
|
96 |
||
97 |
lemma smod_int_alt_def: |
|
98 |
"(a::int) smod b = sgn (a) * (abs a mod abs b)" |
|
99 |
apply (clarsimp simp: signed_modulo_int_def signed_divide_int_def) |
|
100 |
apply (clarsimp simp: minus_div_mult_eq_mod [symmetric] abs_sgn sgn_mult sgn_if algebra_split_simps) |
|
101 |
done |
|
102 |
||
103 |
lemma smod_int_range: |
|
104 |
"b \<noteq> 0 \<Longrightarrow> (a::int) smod b \<in> { - abs b + 1 .. abs b - 1 }" |
|
105 |
apply (case_tac "b > 0") |
|
106 |
apply (insert pos_mod_conj [where a=a and b=b])[1] |
|
107 |
apply (insert pos_mod_conj [where a="-a" and b=b])[1] |
|
108 |
apply (auto simp: smod_int_alt_def algebra_simps sgn_if |
|
109 |
abs_if not_less add1_zle_eq [simplified add.commute])[1] |
|
110 |
apply (metis add_nonneg_nonneg int_one_le_iff_zero_less le_less less_add_same_cancel2 not_le pos_mod_conj) |
|
111 |
apply (metis (full_types) add.inverse_inverse eucl_rel_int eucl_rel_int_iff le_less_trans neg_0_le_iff_le) |
|
112 |
apply (insert neg_mod_conj [where a=a and b="b"])[1] |
|
113 |
apply (insert neg_mod_conj [where a="-a" and b="b"])[1] |
|
114 |
apply (clarsimp simp: smod_int_alt_def algebra_simps sgn_if |
|
115 |
abs_if not_less add1_zle_eq [simplified add.commute]) |
|
116 |
apply (metis neg_0_less_iff_less neg_mod_conj not_le not_less_iff_gr_or_eq order_trans pos_mod_conj) |
|
117 |
done |
|
118 |
||
119 |
lemma smod_int_compares: |
|
120 |
"\<lbrakk> 0 \<le> a; 0 < b \<rbrakk> \<Longrightarrow> (a :: int) smod b < b" |
|
121 |
"\<lbrakk> 0 \<le> a; 0 < b \<rbrakk> \<Longrightarrow> 0 \<le> (a :: int) smod b" |
|
122 |
"\<lbrakk> a \<le> 0; 0 < b \<rbrakk> \<Longrightarrow> -b < (a :: int) smod b" |
|
123 |
"\<lbrakk> a \<le> 0; 0 < b \<rbrakk> \<Longrightarrow> (a :: int) smod b \<le> 0" |
|
124 |
"\<lbrakk> 0 \<le> a; b < 0 \<rbrakk> \<Longrightarrow> (a :: int) smod b < - b" |
|
125 |
"\<lbrakk> 0 \<le> a; b < 0 \<rbrakk> \<Longrightarrow> 0 \<le> (a :: int) smod b" |
|
126 |
"\<lbrakk> a \<le> 0; b < 0 \<rbrakk> \<Longrightarrow> (a :: int) smod b \<le> 0" |
|
127 |
"\<lbrakk> a \<le> 0; b < 0 \<rbrakk> \<Longrightarrow> b \<le> (a :: int) smod b" |
|
128 |
apply (insert smod_int_range [where a=a and b=b]) |
|
129 |
apply (auto simp: add1_zle_eq smod_int_alt_def sgn_if) |
|
130 |
done |
|
131 |
||
132 |
lemma smod_int_mod_0 [simp]: |
|
133 |
"x smod (0 :: int) = x" |
|
134 |
by (clarsimp simp: signed_modulo_int_def) |
|
135 |
||
136 |
lemma smod_int_0_mod [simp]: |
|
137 |
"0 smod (x :: int) = 0" |
|
138 |
by (clarsimp simp: smod_int_alt_def) |
|
139 |
||
140 |
lemma smod_mod_positive: |
|
141 |
"\<lbrakk> 0 \<le> (a :: int); 0 \<le> b \<rbrakk> \<Longrightarrow> a smod b = a mod b" |
|
142 |
by (clarsimp simp: smod_int_alt_def zsgn_def) |
|
143 |
||
74592 | 144 |
lemma minus_sdiv_eq [simp]: |
145 |
\<open>- k sdiv l = - (k sdiv l)\<close> for k l :: int |
|
146 |
by (simp add: signed_divide_int_def) |
|
147 |
||
148 |
lemma sdiv_minus_eq [simp]: |
|
149 |
\<open>k sdiv - l = - (k sdiv l)\<close> for k l :: int |
|
150 |
by (simp add: signed_divide_int_def) |
|
151 |
||
152 |
lemma sdiv_int_numeral_numeral [simp]: |
|
153 |
\<open>numeral m sdiv numeral n = numeral m div (numeral n :: int)\<close> |
|
154 |
by (simp add: signed_divide_int_def) |
|
155 |
||
156 |
lemma minus_smod_eq [simp]: |
|
157 |
\<open>- k smod l = - (k smod l)\<close> for k l :: int |
|
158 |
by (simp add: smod_int_alt_def) |
|
159 |
||
160 |
lemma smod_minus_eq [simp]: |
|
161 |
\<open>k smod - l = k smod l\<close> for k l :: int |
|
162 |
by (simp add: smod_int_alt_def) |
|
163 |
||
164 |
lemma smod_int_numeral_numeral [simp]: |
|
165 |
\<open>numeral m smod numeral n = numeral m mod (numeral n :: int)\<close> |
|
166 |
by (simp add: smod_int_alt_def) |
|
167 |
||
72281
beeadb35e357
more thorough treatment of division, particularly signed division on int and word
haftmann
parents:
diff
changeset
|
168 |
end |