| author | paulson | 
| Fri, 24 Jan 2003 18:13:59 +0100 | |
| changeset 13786 | ab8f39f48a6f | 
| parent 13517 | 42efec18f5b2 | 
| child 13882 | 2266550ab316 | 
| permissions | -rw-r--r-- | 
| 3366 | 1 | (* Title: HOL/Divides.thy | 
| 2 | ID: $Id$ | |
| 3 | Author: Lawrence C Paulson, Cambridge University Computer Laboratory | |
| 6865 
5577ffe4c2f1
now div and mod are overloaded; dvd is polymorphic
 paulson parents: 
3366diff
changeset | 4 | Copyright 1999 University of Cambridge | 
| 3366 | 5 | |
| 6 | The division operators div, mod and the divides relation "dvd" | |
| 7 | *) | |
| 8 | ||
| 13152 | 9 | theory Divides = NatArith files("Divides_lemmas.ML"):
 | 
| 3366 | 10 | |
| 8902 | 11 | (*We use the same class for div and mod; | 
| 6865 
5577ffe4c2f1
now div and mod are overloaded; dvd is polymorphic
 paulson parents: 
3366diff
changeset | 12 | moreover, dvd is defined whenever multiplication is*) | 
| 
5577ffe4c2f1
now div and mod are overloaded; dvd is polymorphic
 paulson parents: 
3366diff
changeset | 13 | axclass | 
| 12338 
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
 wenzelm parents: 
10789diff
changeset | 14 | div < type | 
| 6865 
5577ffe4c2f1
now div and mod are overloaded; dvd is polymorphic
 paulson parents: 
3366diff
changeset | 15 | |
| 13152 | 16 | instance nat :: div .. | 
| 17 | instance nat :: plus_ac0 | |
| 18 | proof qed (rule add_commute add_assoc add_0)+ | |
| 6865 
5577ffe4c2f1
now div and mod are overloaded; dvd is polymorphic
 paulson parents: 
3366diff
changeset | 19 | |
| 3366 | 20 | consts | 
| 13152 | 21 | div :: "'a::div \<Rightarrow> 'a \<Rightarrow> 'a" (infixl 70) | 
| 22 | mod :: "'a::div \<Rightarrow> 'a \<Rightarrow> 'a" (infixl 70) | |
| 23 | dvd :: "'a::times \<Rightarrow> 'a \<Rightarrow> bool" (infixl 50) | |
| 3366 | 24 | |
| 25 | ||
| 26 | defs | |
| 6865 
5577ffe4c2f1
now div and mod are overloaded; dvd is polymorphic
 paulson parents: 
3366diff
changeset | 27 | |
| 13152 | 28 | mod_def: "m mod n == wfrec (trancl pred_nat) | 
| 7029 
08d4eb8500dd
new division laws taking advantage of (m div 0) = 0 and (m mod 0) = m
 paulson parents: 
6865diff
changeset | 29 | (%f j. if j<n | n=0 then j else f (j-n)) m" | 
| 6865 
5577ffe4c2f1
now div and mod are overloaded; dvd is polymorphic
 paulson parents: 
3366diff
changeset | 30 | |
| 13152 | 31 | div_def: "m div n == wfrec (trancl pred_nat) | 
| 7029 
08d4eb8500dd
new division laws taking advantage of (m div 0) = 0 and (m mod 0) = m
 paulson parents: 
6865diff
changeset | 32 | (%f j. if j<n | n=0 then 0 else Suc (f (j-n))) m" | 
| 3366 | 33 | |
| 6865 
5577ffe4c2f1
now div and mod are overloaded; dvd is polymorphic
 paulson parents: 
3366diff
changeset | 34 | (*The definition of dvd is polymorphic!*) | 
| 13152 | 35 | dvd_def: "m dvd n == EX k. n = m*k" | 
| 3366 | 36 | |
| 10559 
d3fd54fc659b
many new div and mod properties (borrowed from Integ/IntDiv)
 paulson parents: 
10214diff
changeset | 37 | (*This definition helps prove the harder properties of div and mod. | 
| 
d3fd54fc659b
many new div and mod properties (borrowed from Integ/IntDiv)
 paulson parents: 
10214diff
changeset | 38 | It is copied from IntDiv.thy; should it be overloaded?*) | 
| 
d3fd54fc659b
many new div and mod properties (borrowed from Integ/IntDiv)
 paulson parents: 
10214diff
changeset | 39 | constdefs | 
| 
d3fd54fc659b
many new div and mod properties (borrowed from Integ/IntDiv)
 paulson parents: 
10214diff
changeset | 40 | quorem :: "(nat*nat) * (nat*nat) => bool" | 
| 
d3fd54fc659b
many new div and mod properties (borrowed from Integ/IntDiv)
 paulson parents: 
10214diff
changeset | 41 | "quorem == %((a,b), (q,r)). | 
| 
d3fd54fc659b
many new div and mod properties (borrowed from Integ/IntDiv)
 paulson parents: 
10214diff
changeset | 42 | a = b*q + r & | 
| 
d3fd54fc659b
many new div and mod properties (borrowed from Integ/IntDiv)
 paulson parents: 
10214diff
changeset | 43 | (if 0<b then 0<=r & r<b else b<r & r <=0)" | 
| 
d3fd54fc659b
many new div and mod properties (borrowed from Integ/IntDiv)
 paulson parents: 
10214diff
changeset | 44 | |
| 13152 | 45 | use "Divides_lemmas.ML" | 
| 46 | ||
| 47 | lemma split_div: | |
| 13189 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 48 | "P(n div k :: nat) = | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 49 | ((k = 0 \<longrightarrow> P 0) \<and> (k \<noteq> 0 \<longrightarrow> (!i. !j<k. n = k*i + j \<longrightarrow> P i)))" | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 50 | (is "?P = ?Q" is "_ = (_ \<and> (_ \<longrightarrow> ?R))") | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 51 | proof | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 52 | assume P: ?P | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 53 | show ?Q | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 54 | proof (cases) | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 55 | assume "k = 0" | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 56 | with P show ?Q by(simp add:DIVISION_BY_ZERO_DIV) | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 57 | next | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 58 | assume not0: "k \<noteq> 0" | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 59 | thus ?Q | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 60 | proof (simp, intro allI impI) | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 61 | fix i j | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 62 | assume n: "n = k*i + j" and j: "j < k" | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 63 | show "P i" | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 64 | proof (cases) | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 65 | assume "i = 0" | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 66 | with n j P show "P i" by simp | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 67 | next | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 68 | assume "i \<noteq> 0" | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 69 | with not0 n j P show "P i" by(simp add:add_ac) | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 70 | qed | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 71 | qed | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 72 | qed | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 73 | next | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 74 | assume Q: ?Q | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 75 | show ?P | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 76 | proof (cases) | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 77 | assume "k = 0" | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 78 | with Q show ?P by(simp add:DIVISION_BY_ZERO_DIV) | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 79 | next | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 80 | assume not0: "k \<noteq> 0" | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 81 | with Q have R: ?R by simp | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 82 | from not0 R[THEN spec,of "n div k",THEN spec, of "n mod k"] | 
| 13517 | 83 | show ?P by simp | 
| 13189 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 84 | qed | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 85 | qed | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 86 | |
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 87 | lemma split_mod: | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 88 | "P(n mod k :: nat) = | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 89 | ((k = 0 \<longrightarrow> P n) \<and> (k \<noteq> 0 \<longrightarrow> (!i. !j<k. n = k*i + j \<longrightarrow> P j)))" | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 90 | (is "?P = ?Q" is "_ = (_ \<and> (_ \<longrightarrow> ?R))") | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 91 | proof | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 92 | assume P: ?P | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 93 | show ?Q | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 94 | proof (cases) | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 95 | assume "k = 0" | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 96 | with P show ?Q by(simp add:DIVISION_BY_ZERO_MOD) | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 97 | next | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 98 | assume not0: "k \<noteq> 0" | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 99 | thus ?Q | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 100 | proof (simp, intro allI impI) | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 101 | fix i j | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 102 | assume "n = k*i + j" "j < k" | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 103 | thus "P j" using not0 P by(simp add:add_ac mult_ac) | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 104 | qed | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 105 | qed | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 106 | next | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 107 | assume Q: ?Q | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 108 | show ?P | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 109 | proof (cases) | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 110 | assume "k = 0" | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 111 | with Q show ?P by(simp add:DIVISION_BY_ZERO_MOD) | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 112 | next | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 113 | assume not0: "k \<noteq> 0" | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 114 | with Q have R: ?R by simp | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 115 | from not0 R[THEN spec,of "n div k",THEN spec, of "n mod k"] | 
| 13517 | 116 | show ?P by simp | 
| 13189 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 117 | qed | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 118 | qed | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 119 | |
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 120 | (* | 
| 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 121 | lemma split_div: | 
| 13152 | 122 | assumes m: "m \<noteq> 0" | 
| 123 | shows "P(n div m :: nat) = (!i. !j<m. n = m*i + j \<longrightarrow> P i)" | |
| 124 | (is "?P = ?Q") | |
| 125 | proof | |
| 126 | assume P: ?P | |
| 127 | show ?Q | |
| 128 | proof (intro allI impI) | |
| 129 | fix i j | |
| 130 | assume n: "n = m*i + j" and j: "j < m" | |
| 131 | show "P i" | |
| 132 | proof (cases) | |
| 133 | assume "i = 0" | |
| 134 | with n j P show "P i" by simp | |
| 135 | next | |
| 136 | assume "i \<noteq> 0" | |
| 137 | with n j P show "P i" by (simp add:add_ac div_mult_self1) | |
| 138 | qed | |
| 139 | qed | |
| 140 | next | |
| 141 | assume Q: ?Q | |
| 142 | from m Q[THEN spec,of "n div m",THEN spec, of "n mod m"] | |
| 13517 | 143 | show ?P by simp | 
| 13152 | 144 | qed | 
| 145 | ||
| 146 | lemma split_mod: | |
| 147 | assumes m: "m \<noteq> 0" | |
| 148 | shows "P(n mod m :: nat) = (!i. !j<m. n = m*i + j \<longrightarrow> P j)" | |
| 149 | (is "?P = ?Q") | |
| 150 | proof | |
| 151 | assume P: ?P | |
| 152 | show ?Q | |
| 153 | proof (intro allI impI) | |
| 154 | fix i j | |
| 155 | assume "n = m*i + j" "j < m" | |
| 156 | thus "P j" using m P by(simp add:add_ac mult_ac) | |
| 157 | qed | |
| 158 | next | |
| 159 | assume Q: ?Q | |
| 160 | from m Q[THEN spec,of "n div m",THEN spec, of "n mod m"] | |
| 13517 | 161 | show ?P by simp | 
| 13152 | 162 | qed | 
| 13189 
81ed5c6de890
Now arith can deal with div/mod arbitrary nat numerals.
 nipkow parents: 
13152diff
changeset | 163 | *) | 
| 3366 | 164 | end |