| author | wenzelm |
| Sat, 03 Oct 2020 14:06:00 +0200 | |
| changeset 72367 | d3069e7e1175 |
| parent 72281 | beeadb35e357 |
| child 72768 | 4ab04bafae35 |
| 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 |
|
|
beeadb35e357
more thorough treatment of division, particularly signed division on int and word
haftmann
parents:
diff
changeset
|
27 |
end |