| author | lcp | 
| Fri, 16 Dec 1994 17:41:49 +0100 | |
| changeset 800 | 23f55b829ccb | 
| parent 760 | f0200e91b272 | 
| child 1461 | 6bcb44e4d6e5 | 
| permissions | -rw-r--r-- | 
| 0 | 1 | (* Title: ZF/arith.ML | 
| 2 | ID: $Id$ | |
| 3 | Author: Lawrence C Paulson, Cambridge University Computer Laboratory | |
| 4 | Copyright 1992 University of Cambridge | |
| 5 | ||
| 6 | For arith.thy. Arithmetic operators and their definitions | |
| 7 | ||
| 8 | Proofs about elementary arithmetic: addition, multiplication, etc. | |
| 9 | ||
| 10 | Could prove def_rec_0, def_rec_succ... | |
| 11 | *) | |
| 12 | ||
| 13 | open Arith; | |
| 14 | ||
| 15 | (*"Difference" is subtraction of natural numbers. | |
| 16 | There are no negative numbers; we have | |
| 17 | m #- n = 0 iff m<=n and m #- n = succ(k) iff m>n. | |
| 18 | Also, rec(m, 0, %z w.z) is pred(m). | |
| 19 | *) | |
| 20 | ||
| 21 | (** rec -- better than nat_rec; the succ case has no type requirement! **) | |
| 22 | ||
| 23 | val rec_trans = rec_def RS def_transrec RS trans; | |
| 24 | ||
| 25 | goal Arith.thy "rec(0,a,b) = a"; | |
| 26 | by (rtac rec_trans 1); | |
| 27 | by (rtac nat_case_0 1); | |
| 760 | 28 | qed "rec_0"; | 
| 0 | 29 | |
| 30 | goal Arith.thy "rec(succ(m),a,b) = b(m, rec(m,a,b))"; | |
| 31 | by (rtac rec_trans 1); | |
| 14 
1c0926788772
ex/{bin.ML,comb.ML,prop.ML}: replaced NewSext by Syntax.simple_sext
 lcp parents: 
6diff
changeset | 32 | by (simp_tac (ZF_ss addsimps [nat_case_succ, nat_succI]) 1); | 
| 760 | 33 | qed "rec_succ"; | 
| 0 | 34 | |
| 35 | val major::prems = goal Arith.thy | |
| 36 | "[| n: nat; \ | |
| 37 | \ a: C(0); \ | |
| 38 | \ !!m z. [| m: nat; z: C(m) |] ==> b(m,z): C(succ(m)) \ | |
| 39 | \ |] ==> rec(n,a,b) : C(n)"; | |
| 40 | by (rtac (major RS nat_induct) 1); | |
| 41 | by (ALLGOALS | |
| 6 
8ce8c4d13d4d
Installation of new simplifier for ZF.  Deleted all congruence rules not
 lcp parents: 
0diff
changeset | 42 | (asm_simp_tac (ZF_ss addsimps (prems@[rec_0,rec_succ])))); | 
| 760 | 43 | qed "rec_type"; | 
| 0 | 44 | |
| 435 | 45 | val nat_le_refl = nat_into_Ord RS le_refl; | 
| 25 
3ac1c0c0016e
ordinal: DEFINITION of < and le to replace : and <= on ordinals!  Many
 lcp parents: 
14diff
changeset | 46 | |
| 
3ac1c0c0016e
ordinal: DEFINITION of < and le to replace : and <= on ordinals!  Many
 lcp parents: 
14diff
changeset | 47 | val nat_typechecks = [rec_type, nat_0I, nat_1I, nat_succI, Ord_nat]; | 
| 0 | 48 | |
| 25 
3ac1c0c0016e
ordinal: DEFINITION of < and le to replace : and <= on ordinals!  Many
 lcp parents: 
14diff
changeset | 49 | val nat_simps = [rec_0, rec_succ, not_lt0, nat_0_le, le0_iff, succ_le_iff, | 
| 
3ac1c0c0016e
ordinal: DEFINITION of < and le to replace : and <= on ordinals!  Many
 lcp parents: 
14diff
changeset | 50 | nat_le_refl]; | 
| 
3ac1c0c0016e
ordinal: DEFINITION of < and le to replace : and <= on ordinals!  Many
 lcp parents: 
14diff
changeset | 51 | |
| 
3ac1c0c0016e
ordinal: DEFINITION of < and le to replace : and <= on ordinals!  Many
 lcp parents: 
14diff
changeset | 52 | val nat_ss = ZF_ss addsimps (nat_simps @ nat_typechecks); | 
| 0 | 53 | |
| 54 | ||
| 55 | (** Addition **) | |
| 56 | ||
| 760 | 57 | qed_goalw "add_type" Arith.thy [add_def] | 
| 0 | 58 | "[| m:nat; n:nat |] ==> m #+ n : nat" | 
| 59 | (fn prems=> [ (typechk_tac (prems@nat_typechecks@ZF_typechecks)) ]); | |
| 60 | ||
| 760 | 61 | qed_goalw "add_0" Arith.thy [add_def] | 
| 0 | 62 | "0 #+ n = n" | 
| 63 | (fn _ => [ (rtac rec_0 1) ]); | |
| 64 | ||
| 760 | 65 | qed_goalw "add_succ" Arith.thy [add_def] | 
| 0 | 66 | "succ(m) #+ n = succ(m #+ n)" | 
| 67 | (fn _=> [ (rtac rec_succ 1) ]); | |
| 68 | ||
| 69 | (** Multiplication **) | |
| 70 | ||
| 760 | 71 | qed_goalw "mult_type" Arith.thy [mult_def] | 
| 0 | 72 | "[| m:nat; n:nat |] ==> m #* n : nat" | 
| 73 | (fn prems=> | |
| 74 | [ (typechk_tac (prems@[add_type]@nat_typechecks@ZF_typechecks)) ]); | |
| 75 | ||
| 760 | 76 | qed_goalw "mult_0" Arith.thy [mult_def] | 
| 0 | 77 | "0 #* n = 0" | 
| 78 | (fn _ => [ (rtac rec_0 1) ]); | |
| 79 | ||
| 760 | 80 | qed_goalw "mult_succ" Arith.thy [mult_def] | 
| 0 | 81 | "succ(m) #* n = n #+ (m #* n)" | 
| 82 | (fn _ => [ (rtac rec_succ 1) ]); | |
| 83 | ||
| 84 | (** Difference **) | |
| 85 | ||
| 760 | 86 | qed_goalw "diff_type" Arith.thy [diff_def] | 
| 0 | 87 | "[| m:nat; n:nat |] ==> m #- n : nat" | 
| 88 | (fn prems=> [ (typechk_tac (prems@nat_typechecks@ZF_typechecks)) ]); | |
| 89 | ||
| 760 | 90 | qed_goalw "diff_0" Arith.thy [diff_def] | 
| 0 | 91 | "m #- 0 = m" | 
| 92 | (fn _ => [ (rtac rec_0 1) ]); | |
| 93 | ||
| 760 | 94 | qed_goalw "diff_0_eq_0" Arith.thy [diff_def] | 
| 0 | 95 | "n:nat ==> 0 #- n = 0" | 
| 96 | (fn [prem]=> | |
| 97 | [ (rtac (prem RS nat_induct) 1), | |
| 6 
8ce8c4d13d4d
Installation of new simplifier for ZF.  Deleted all congruence rules not
 lcp parents: 
0diff
changeset | 98 | (ALLGOALS (asm_simp_tac nat_ss)) ]); | 
| 0 | 99 | |
| 100 | (*Must simplify BEFORE the induction!! (Else we get a critical pair) | |
| 101 | succ(m) #- succ(n) rewrites to pred(succ(m) #- n) *) | |
| 760 | 102 | qed_goalw "diff_succ_succ" Arith.thy [diff_def] | 
| 0 | 103 | "[| m:nat; n:nat |] ==> succ(m) #- succ(n) = m #- n" | 
| 104 | (fn prems=> | |
| 6 
8ce8c4d13d4d
Installation of new simplifier for ZF.  Deleted all congruence rules not
 lcp parents: 
0diff
changeset | 105 | [ (asm_simp_tac (nat_ss addsimps prems) 1), | 
| 0 | 106 | (nat_ind_tac "n" prems 1), | 
| 6 
8ce8c4d13d4d
Installation of new simplifier for ZF.  Deleted all congruence rules not
 lcp parents: 
0diff
changeset | 107 | (ALLGOALS (asm_simp_tac (nat_ss addsimps prems))) ]); | 
| 0 | 108 | |
| 109 | val prems = goal Arith.thy | |
| 25 
3ac1c0c0016e
ordinal: DEFINITION of < and le to replace : and <= on ordinals!  Many
 lcp parents: 
14diff
changeset | 110 | "[| m:nat; n:nat |] ==> m #- n le m"; | 
| 
3ac1c0c0016e
ordinal: DEFINITION of < and le to replace : and <= on ordinals!  Many
 lcp parents: 
14diff
changeset | 111 | by (rtac (prems MRS diff_induct) 1); | 
| 
3ac1c0c0016e
ordinal: DEFINITION of < and le to replace : and <= on ordinals!  Many
 lcp parents: 
14diff
changeset | 112 | by (etac leE 3); | 
| 0 | 113 | by (ALLGOALS | 
| 6 
8ce8c4d13d4d
Installation of new simplifier for ZF.  Deleted all congruence rules not
 lcp parents: 
0diff
changeset | 114 | (asm_simp_tac | 
| 25 
3ac1c0c0016e
ordinal: DEFINITION of < and le to replace : and <= on ordinals!  Many
 lcp parents: 
14diff
changeset | 115 | (nat_ss addsimps (prems @ [le_iff, diff_0, diff_0_eq_0, | 
| 435 | 116 | diff_succ_succ, nat_into_Ord])))); | 
| 760 | 117 | qed "diff_le_self"; | 
| 0 | 118 | |
| 119 | (*** Simplification over add, mult, diff ***) | |
| 120 | ||
| 121 | val arith_typechecks = [add_type, mult_type, diff_type]; | |
| 25 
3ac1c0c0016e
ordinal: DEFINITION of < and le to replace : and <= on ordinals!  Many
 lcp parents: 
14diff
changeset | 122 | val arith_simps = [add_0, add_succ, | 
| 
3ac1c0c0016e
ordinal: DEFINITION of < and le to replace : and <= on ordinals!  Many
 lcp parents: 
14diff
changeset | 123 | mult_0, mult_succ, | 
| 
3ac1c0c0016e
ordinal: DEFINITION of < and le to replace : and <= on ordinals!  Many
 lcp parents: 
14diff
changeset | 124 | diff_0, diff_0_eq_0, diff_succ_succ]; | 
| 0 | 125 | |
| 25 
3ac1c0c0016e
ordinal: DEFINITION of < and le to replace : and <= on ordinals!  Many
 lcp parents: 
14diff
changeset | 126 | val arith_ss = nat_ss addsimps (arith_simps@arith_typechecks); | 
| 0 | 127 | |
| 128 | (*** Addition ***) | |
| 129 | ||
| 130 | (*Associative law for addition*) | |
| 760 | 131 | qed_goal "add_assoc" Arith.thy | 
| 0 | 132 | "m:nat ==> (m #+ n) #+ k = m #+ (n #+ k)" | 
| 133 | (fn prems=> | |
| 134 | [ (nat_ind_tac "m" prems 1), | |
| 6 
8ce8c4d13d4d
Installation of new simplifier for ZF.  Deleted all congruence rules not
 lcp parents: 
0diff
changeset | 135 | (ALLGOALS (asm_simp_tac (arith_ss addsimps prems))) ]); | 
| 0 | 136 | |
| 137 | (*The following two lemmas are used for add_commute and sometimes | |
| 138 | elsewhere, since they are safe for rewriting.*) | |
| 760 | 139 | qed_goal "add_0_right" Arith.thy | 
| 0 | 140 | "m:nat ==> m #+ 0 = m" | 
| 141 | (fn prems=> | |
| 142 | [ (nat_ind_tac "m" prems 1), | |
| 6 
8ce8c4d13d4d
Installation of new simplifier for ZF.  Deleted all congruence rules not
 lcp parents: 
0diff
changeset | 143 | (ALLGOALS (asm_simp_tac (arith_ss addsimps prems))) ]); | 
| 0 | 144 | |
| 760 | 145 | qed_goal "add_succ_right" Arith.thy | 
| 0 | 146 | "m:nat ==> m #+ succ(n) = succ(m #+ n)" | 
| 147 | (fn prems=> | |
| 148 | [ (nat_ind_tac "m" prems 1), | |
| 6 
8ce8c4d13d4d
Installation of new simplifier for ZF.  Deleted all congruence rules not
 lcp parents: 
0diff
changeset | 149 | (ALLGOALS (asm_simp_tac (arith_ss addsimps prems))) ]); | 
| 0 | 150 | |
| 151 | (*Commutative law for addition*) | |
| 760 | 152 | qed_goal "add_commute" Arith.thy | 
| 435 | 153 | "!!m n. [| m:nat; n:nat |] ==> m #+ n = n #+ m" | 
| 154 | (fn _ => | |
| 155 | [ (nat_ind_tac "n" [] 1), | |
| 0 | 156 | (ALLGOALS | 
| 435 | 157 | (asm_simp_tac (arith_ss addsimps [add_0_right, add_succ_right]))) ]); | 
| 158 | ||
| 437 | 159 | (*for a/c rewriting*) | 
| 760 | 160 | qed_goal "add_left_commute" Arith.thy | 
| 437 | 161 | "!!m n k. [| m:nat; n:nat |] ==> m#+(n#+k)=n#+(m#+k)" | 
| 162 | (fn _ => [asm_simp_tac (ZF_ss addsimps [add_assoc RS sym, add_commute]) 1]); | |
| 435 | 163 | |
| 164 | (*Addition is an AC-operator*) | |
| 165 | val add_ac = [add_assoc, add_commute, add_left_commute]; | |
| 0 | 166 | |
| 167 | (*Cancellation law on the left*) | |
| 437 | 168 | val [eqn,knat] = goal Arith.thy | 
| 169 | "[| k #+ m = k #+ n; k:nat |] ==> m=n"; | |
| 0 | 170 | by (rtac (eqn RS rev_mp) 1); | 
| 171 | by (nat_ind_tac "k" [knat] 1); | |
| 6 
8ce8c4d13d4d
Installation of new simplifier for ZF.  Deleted all congruence rules not
 lcp parents: 
0diff
changeset | 172 | by (ALLGOALS (simp_tac arith_ss)); | 
| 0 | 173 | by (fast_tac ZF_cs 1); | 
| 760 | 174 | qed "add_left_cancel"; | 
| 0 | 175 | |
| 176 | (*** Multiplication ***) | |
| 177 | ||
| 178 | (*right annihilation in product*) | |
| 760 | 179 | qed_goal "mult_0_right" Arith.thy | 
| 435 | 180 | "!!m. m:nat ==> m #* 0 = 0" | 
| 181 | (fn _=> | |
| 182 | [ (nat_ind_tac "m" [] 1), | |
| 183 | (ALLGOALS (asm_simp_tac arith_ss)) ]); | |
| 0 | 184 | |
| 185 | (*right successor law for multiplication*) | |
| 760 | 186 | qed_goal "mult_succ_right" Arith.thy | 
| 6 
8ce8c4d13d4d
Installation of new simplifier for ZF.  Deleted all congruence rules not
 lcp parents: 
0diff
changeset | 187 | "!!m n. [| m:nat; n:nat |] ==> m #* succ(n) = m #+ (m #* n)" | 
| 435 | 188 | (fn _ => | 
| 6 
8ce8c4d13d4d
Installation of new simplifier for ZF.  Deleted all congruence rules not
 lcp parents: 
0diff
changeset | 189 | [ (nat_ind_tac "m" [] 1), | 
| 435 | 190 | (ALLGOALS (asm_simp_tac (arith_ss addsimps add_ac))) ]); | 
| 0 | 191 | |
| 192 | (*Commutative law for multiplication*) | |
| 760 | 193 | qed_goal "mult_commute" Arith.thy | 
| 0 | 194 | "[| m:nat; n:nat |] ==> m #* n = n #* m" | 
| 195 | (fn prems=> | |
| 196 | [ (nat_ind_tac "m" prems 1), | |
| 6 
8ce8c4d13d4d
Installation of new simplifier for ZF.  Deleted all congruence rules not
 lcp parents: 
0diff
changeset | 197 | (ALLGOALS (asm_simp_tac | 
| 
8ce8c4d13d4d
Installation of new simplifier for ZF.  Deleted all congruence rules not
 lcp parents: 
0diff
changeset | 198 | (arith_ss addsimps (prems@[mult_0_right, mult_succ_right])))) ]); | 
| 0 | 199 | |
| 200 | (*addition distributes over multiplication*) | |
| 760 | 201 | qed_goal "add_mult_distrib" Arith.thy | 
| 14 
1c0926788772
ex/{bin.ML,comb.ML,prop.ML}: replaced NewSext by Syntax.simple_sext
 lcp parents: 
6diff
changeset | 202 | "!!m n. [| m:nat; k:nat |] ==> (m #+ n) #* k = (m #* k) #+ (n #* k)" | 
| 
1c0926788772
ex/{bin.ML,comb.ML,prop.ML}: replaced NewSext by Syntax.simple_sext
 lcp parents: 
6diff
changeset | 203 | (fn _=> | 
| 
1c0926788772
ex/{bin.ML,comb.ML,prop.ML}: replaced NewSext by Syntax.simple_sext
 lcp parents: 
6diff
changeset | 204 | [ (etac nat_induct 1), | 
| 
1c0926788772
ex/{bin.ML,comb.ML,prop.ML}: replaced NewSext by Syntax.simple_sext
 lcp parents: 
6diff
changeset | 205 | (ALLGOALS (asm_simp_tac (arith_ss addsimps [add_assoc RS sym]))) ]); | 
| 0 | 206 | |
| 207 | (*Distributive law on the left; requires an extra typing premise*) | |
| 760 | 208 | qed_goal "add_mult_distrib_left" Arith.thy | 
| 435 | 209 | "!!m. [| m:nat; n:nat; k:nat |] ==> k #* (m #+ n) = (k #* m) #+ (k #* n)" | 
| 0 | 210 | (fn prems=> | 
| 435 | 211 | [ (nat_ind_tac "m" [] 1), | 
| 212 | (asm_simp_tac (arith_ss addsimps [mult_0_right]) 1), | |
| 213 | (asm_simp_tac (arith_ss addsimps ([mult_succ_right] @ add_ac)) 1) ]); | |
| 0 | 214 | |
| 215 | (*Associative law for multiplication*) | |
| 760 | 216 | qed_goal "mult_assoc" Arith.thy | 
| 14 
1c0926788772
ex/{bin.ML,comb.ML,prop.ML}: replaced NewSext by Syntax.simple_sext
 lcp parents: 
6diff
changeset | 217 | "!!m n k. [| m:nat; n:nat; k:nat |] ==> (m #* n) #* k = m #* (n #* k)" | 
| 
1c0926788772
ex/{bin.ML,comb.ML,prop.ML}: replaced NewSext by Syntax.simple_sext
 lcp parents: 
6diff
changeset | 218 | (fn _=> | 
| 
1c0926788772
ex/{bin.ML,comb.ML,prop.ML}: replaced NewSext by Syntax.simple_sext
 lcp parents: 
6diff
changeset | 219 | [ (etac nat_induct 1), | 
| 
1c0926788772
ex/{bin.ML,comb.ML,prop.ML}: replaced NewSext by Syntax.simple_sext
 lcp parents: 
6diff
changeset | 220 | (ALLGOALS (asm_simp_tac (arith_ss addsimps [add_mult_distrib]))) ]); | 
| 0 | 221 | |
| 437 | 222 | (*for a/c rewriting*) | 
| 760 | 223 | qed_goal "mult_left_commute" Arith.thy | 
| 437 | 224 | "!!m n k. [| m:nat; n:nat; k:nat |] ==> m #* (n #* k) = n #* (m #* k)" | 
| 225 | (fn _ => [rtac (mult_commute RS trans) 1, | |
| 226 | rtac (mult_assoc RS trans) 3, | |
| 227 | rtac (mult_commute RS subst_context) 6, | |
| 228 | REPEAT (ares_tac [mult_type] 1)]); | |
| 229 | ||
| 230 | val mult_ac = [mult_assoc,mult_commute,mult_left_commute]; | |
| 231 | ||
| 0 | 232 | |
| 233 | (*** Difference ***) | |
| 234 | ||
| 760 | 235 | qed_goal "diff_self_eq_0" Arith.thy | 
| 0 | 236 | "m:nat ==> m #- m = 0" | 
| 237 | (fn prems=> | |
| 238 | [ (nat_ind_tac "m" prems 1), | |
| 6 
8ce8c4d13d4d
Installation of new simplifier for ZF.  Deleted all congruence rules not
 lcp parents: 
0diff
changeset | 239 | (ALLGOALS (asm_simp_tac (arith_ss addsimps prems))) ]); | 
| 0 | 240 | |
| 25 
3ac1c0c0016e
ordinal: DEFINITION of < and le to replace : and <= on ordinals!  Many
 lcp parents: 
14diff
changeset | 241 | (*Addition is the inverse of subtraction*) | 
| 
3ac1c0c0016e
ordinal: DEFINITION of < and le to replace : and <= on ordinals!  Many
 lcp parents: 
14diff
changeset | 242 | goal Arith.thy "!!m n. [| n le m; m:nat |] ==> n #+ (m#-n) = m"; | 
| 
3ac1c0c0016e
ordinal: DEFINITION of < and le to replace : and <= on ordinals!  Many
 lcp parents: 
14diff
changeset | 243 | by (forward_tac [lt_nat_in_nat] 1); | 
| 127 | 244 | by (etac nat_succI 1); | 
| 25 
3ac1c0c0016e
ordinal: DEFINITION of < and le to replace : and <= on ordinals!  Many
 lcp parents: 
14diff
changeset | 245 | by (etac rev_mp 1); | 
| 0 | 246 | by (res_inst_tac [("m","m"),("n","n")] diff_induct 1);
 | 
| 25 
3ac1c0c0016e
ordinal: DEFINITION of < and le to replace : and <= on ordinals!  Many
 lcp parents: 
14diff
changeset | 247 | by (ALLGOALS (asm_simp_tac arith_ss)); | 
| 760 | 248 | qed "add_diff_inverse"; | 
| 0 | 249 | |
| 250 | (*Subtraction is the inverse of addition. *) | |
| 251 | val [mnat,nnat] = goal Arith.thy | |
| 437 | 252 | "[| m:nat; n:nat |] ==> (n#+m) #- n = m"; | 
| 0 | 253 | by (rtac (nnat RS nat_induct) 1); | 
| 6 
8ce8c4d13d4d
Installation of new simplifier for ZF.  Deleted all congruence rules not
 lcp parents: 
0diff
changeset | 254 | by (ALLGOALS (asm_simp_tac (arith_ss addsimps [mnat]))); | 
| 760 | 255 | qed "diff_add_inverse"; | 
| 0 | 256 | |
| 437 | 257 | goal Arith.thy | 
| 258 | "!!m n. [| m:nat; n:nat |] ==> (m#+n) #- n = m"; | |
| 259 | by (res_inst_tac [("m1","m")] (add_commute RS ssubst) 1);
 | |
| 260 | by (REPEAT (ares_tac [diff_add_inverse] 1)); | |
| 760 | 261 | qed "diff_add_inverse2"; | 
| 437 | 262 | |
| 0 | 263 | val [mnat,nnat] = goal Arith.thy | 
| 264 | "[| m:nat; n:nat |] ==> n #- (n#+m) = 0"; | |
| 265 | by (rtac (nnat RS nat_induct) 1); | |
| 6 
8ce8c4d13d4d
Installation of new simplifier for ZF.  Deleted all congruence rules not
 lcp parents: 
0diff
changeset | 266 | by (ALLGOALS (asm_simp_tac (arith_ss addsimps [mnat]))); | 
| 760 | 267 | qed "diff_add_0"; | 
| 0 | 268 | |
| 269 | ||
| 270 | (*** Remainder ***) | |
| 271 | ||
| 25 
3ac1c0c0016e
ordinal: DEFINITION of < and le to replace : and <= on ordinals!  Many
 lcp parents: 
14diff
changeset | 272 | goal Arith.thy "!!m n. [| 0<n; n le m; m:nat |] ==> m #- n < m"; | 
| 
3ac1c0c0016e
ordinal: DEFINITION of < and le to replace : and <= on ordinals!  Many
 lcp parents: 
14diff
changeset | 273 | by (forward_tac [lt_nat_in_nat] 1 THEN etac nat_succI 1); | 
| 0 | 274 | by (etac rev_mp 1); | 
| 275 | by (etac rev_mp 1); | |
| 276 | by (res_inst_tac [("m","m"),("n","n")] diff_induct 1);
 | |
| 25 
3ac1c0c0016e
ordinal: DEFINITION of < and le to replace : and <= on ordinals!  Many
 lcp parents: 
14diff
changeset | 277 | by (ALLGOALS (asm_simp_tac (nat_ss addsimps [diff_le_self,diff_succ_succ]))); | 
| 760 | 278 | qed "div_termination"; | 
| 0 | 279 | |
| 25 
3ac1c0c0016e
ordinal: DEFINITION of < and le to replace : and <= on ordinals!  Many
 lcp parents: 
14diff
changeset | 280 | val div_rls = (*for mod and div*) | 
| 
3ac1c0c0016e
ordinal: DEFINITION of < and le to replace : and <= on ordinals!  Many
 lcp parents: 
14diff
changeset | 281 | nat_typechecks @ | 
| 
3ac1c0c0016e
ordinal: DEFINITION of < and le to replace : and <= on ordinals!  Many
 lcp parents: 
14diff
changeset | 282 | [Ord_transrec_type, apply_type, div_termination RS ltD, if_type, | 
| 435 | 283 | nat_into_Ord, not_lt_iff_le RS iffD1]; | 
| 25 
3ac1c0c0016e
ordinal: DEFINITION of < and le to replace : and <= on ordinals!  Many
 lcp parents: 
14diff
changeset | 284 | |
| 435 | 285 | val div_ss = ZF_ss addsimps [nat_into_Ord, div_termination RS ltD, | 
| 25 
3ac1c0c0016e
ordinal: DEFINITION of < and le to replace : and <= on ordinals!  Many
 lcp parents: 
14diff
changeset | 286 | not_lt_iff_le RS iffD2]; | 
| 0 | 287 | |
| 288 | (*Type checking depends upon termination!*) | |
| 25 
3ac1c0c0016e
ordinal: DEFINITION of < and le to replace : and <= on ordinals!  Many
 lcp parents: 
14diff
changeset | 289 | goalw Arith.thy [mod_def] "!!m n. [| 0<n; m:nat; n:nat |] ==> m mod n : nat"; | 
| 
3ac1c0c0016e
ordinal: DEFINITION of < and le to replace : and <= on ordinals!  Many
 lcp parents: 
14diff
changeset | 290 | by (REPEAT (ares_tac div_rls 1 ORELSE etac lt_trans 1)); | 
| 760 | 291 | qed "mod_type"; | 
| 0 | 292 | |
| 25 
3ac1c0c0016e
ordinal: DEFINITION of < and le to replace : and <= on ordinals!  Many
 lcp parents: 
14diff
changeset | 293 | goal Arith.thy "!!m n. [| 0<n; m<n |] ==> m mod n = m"; | 
| 0 | 294 | by (rtac (mod_def RS def_transrec RS trans) 1); | 
| 25 
3ac1c0c0016e
ordinal: DEFINITION of < and le to replace : and <= on ordinals!  Many
 lcp parents: 
14diff
changeset | 295 | by (asm_simp_tac div_ss 1); | 
| 760 | 296 | qed "mod_less"; | 
| 0 | 297 | |
| 25 
3ac1c0c0016e
ordinal: DEFINITION of < and le to replace : and <= on ordinals!  Many
 lcp parents: 
14diff
changeset | 298 | goal Arith.thy "!!m n. [| 0<n; n le m; m:nat |] ==> m mod n = (m#-n) mod n"; | 
| 
3ac1c0c0016e
ordinal: DEFINITION of < and le to replace : and <= on ordinals!  Many
 lcp parents: 
14diff
changeset | 299 | by (forward_tac [lt_nat_in_nat] 1 THEN etac nat_succI 1); | 
| 0 | 300 | by (rtac (mod_def RS def_transrec RS trans) 1); | 
| 25 
3ac1c0c0016e
ordinal: DEFINITION of < and le to replace : and <= on ordinals!  Many
 lcp parents: 
14diff
changeset | 301 | by (asm_simp_tac div_ss 1); | 
| 760 | 302 | qed "mod_geq"; | 
| 0 | 303 | |
| 304 | (*** Quotient ***) | |
| 305 | ||
| 306 | (*Type checking depends upon termination!*) | |
| 25 
3ac1c0c0016e
ordinal: DEFINITION of < and le to replace : and <= on ordinals!  Many
 lcp parents: 
14diff
changeset | 307 | goalw Arith.thy [div_def] | 
| 
3ac1c0c0016e
ordinal: DEFINITION of < and le to replace : and <= on ordinals!  Many
 lcp parents: 
14diff
changeset | 308 | "!!m n. [| 0<n; m:nat; n:nat |] ==> m div n : nat"; | 
| 
3ac1c0c0016e
ordinal: DEFINITION of < and le to replace : and <= on ordinals!  Many
 lcp parents: 
14diff
changeset | 309 | by (REPEAT (ares_tac div_rls 1 ORELSE etac lt_trans 1)); | 
| 760 | 310 | qed "div_type"; | 
| 0 | 311 | |
| 25 
3ac1c0c0016e
ordinal: DEFINITION of < and le to replace : and <= on ordinals!  Many
 lcp parents: 
14diff
changeset | 312 | goal Arith.thy "!!m n. [| 0<n; m<n |] ==> m div n = 0"; | 
| 0 | 313 | by (rtac (div_def RS def_transrec RS trans) 1); | 
| 25 
3ac1c0c0016e
ordinal: DEFINITION of < and le to replace : and <= on ordinals!  Many
 lcp parents: 
14diff
changeset | 314 | by (asm_simp_tac div_ss 1); | 
| 760 | 315 | qed "div_less"; | 
| 0 | 316 | |
| 25 
3ac1c0c0016e
ordinal: DEFINITION of < and le to replace : and <= on ordinals!  Many
 lcp parents: 
14diff
changeset | 317 | goal Arith.thy | 
| 
3ac1c0c0016e
ordinal: DEFINITION of < and le to replace : and <= on ordinals!  Many
 lcp parents: 
14diff
changeset | 318 | "!!m n. [| 0<n; n le m; m:nat |] ==> m div n = succ((m#-n) div n)"; | 
| 
3ac1c0c0016e
ordinal: DEFINITION of < and le to replace : and <= on ordinals!  Many
 lcp parents: 
14diff
changeset | 319 | by (forward_tac [lt_nat_in_nat] 1 THEN etac nat_succI 1); | 
| 0 | 320 | by (rtac (div_def RS def_transrec RS trans) 1); | 
| 25 
3ac1c0c0016e
ordinal: DEFINITION of < and le to replace : and <= on ordinals!  Many
 lcp parents: 
14diff
changeset | 321 | by (asm_simp_tac div_ss 1); | 
| 760 | 322 | qed "div_geq"; | 
| 0 | 323 | |
| 324 | (*Main Result.*) | |
| 25 
3ac1c0c0016e
ordinal: DEFINITION of < and le to replace : and <= on ordinals!  Many
 lcp parents: 
14diff
changeset | 325 | goal Arith.thy | 
| 
3ac1c0c0016e
ordinal: DEFINITION of < and le to replace : and <= on ordinals!  Many
 lcp parents: 
14diff
changeset | 326 | "!!m n. [| 0<n; m:nat; n:nat |] ==> (m div n)#*n #+ m mod n = m"; | 
| 
3ac1c0c0016e
ordinal: DEFINITION of < and le to replace : and <= on ordinals!  Many
 lcp parents: 
14diff
changeset | 327 | by (etac complete_induct 1); | 
| 437 | 328 | by (excluded_middle_tac "x<n" 1); | 
| 25 
3ac1c0c0016e
ordinal: DEFINITION of < and le to replace : and <= on ordinals!  Many
 lcp parents: 
14diff
changeset | 329 | (*case x<n*) | 
| 
3ac1c0c0016e
ordinal: DEFINITION of < and le to replace : and <= on ordinals!  Many
 lcp parents: 
14diff
changeset | 330 | by (asm_simp_tac (arith_ss addsimps [mod_less, div_less]) 2); | 
| 
3ac1c0c0016e
ordinal: DEFINITION of < and le to replace : and <= on ordinals!  Many
 lcp parents: 
14diff
changeset | 331 | (*case n le x*) | 
| 
3ac1c0c0016e
ordinal: DEFINITION of < and le to replace : and <= on ordinals!  Many
 lcp parents: 
14diff
changeset | 332 | by (asm_full_simp_tac | 
| 435 | 333 | (arith_ss addsimps [not_lt_iff_le, nat_into_Ord, | 
| 25 
3ac1c0c0016e
ordinal: DEFINITION of < and le to replace : and <= on ordinals!  Many
 lcp parents: 
14diff
changeset | 334 | mod_geq, div_geq, add_assoc, | 
| 
3ac1c0c0016e
ordinal: DEFINITION of < and le to replace : and <= on ordinals!  Many
 lcp parents: 
14diff
changeset | 335 | div_termination RS ltD, add_diff_inverse]) 1); | 
| 760 | 336 | qed "mod_div_equality"; | 
| 0 | 337 | |
| 338 | ||
| 25 
3ac1c0c0016e
ordinal: DEFINITION of < and le to replace : and <= on ordinals!  Many
 lcp parents: 
14diff
changeset | 339 | (**** Additional theorems about "le" ****) | 
| 0 | 340 | |
| 25 
3ac1c0c0016e
ordinal: DEFINITION of < and le to replace : and <= on ordinals!  Many
 lcp parents: 
14diff
changeset | 341 | goal Arith.thy "!!m n. [| m:nat; n:nat |] ==> m le m #+ n"; | 
| 
3ac1c0c0016e
ordinal: DEFINITION of < and le to replace : and <= on ordinals!  Many
 lcp parents: 
14diff
changeset | 342 | by (etac nat_induct 1); | 
| 
3ac1c0c0016e
ordinal: DEFINITION of < and le to replace : and <= on ordinals!  Many
 lcp parents: 
14diff
changeset | 343 | by (ALLGOALS (asm_simp_tac arith_ss)); | 
| 760 | 344 | qed "add_le_self"; | 
| 14 
1c0926788772
ex/{bin.ML,comb.ML,prop.ML}: replaced NewSext by Syntax.simple_sext
 lcp parents: 
6diff
changeset | 345 | |
| 25 
3ac1c0c0016e
ordinal: DEFINITION of < and le to replace : and <= on ordinals!  Many
 lcp parents: 
14diff
changeset | 346 | goal Arith.thy "!!m n. [| m:nat; n:nat |] ==> m le n #+ m"; | 
| 14 
1c0926788772
ex/{bin.ML,comb.ML,prop.ML}: replaced NewSext by Syntax.simple_sext
 lcp parents: 
6diff
changeset | 347 | by (rtac (add_commute RS ssubst) 1); | 
| 25 
3ac1c0c0016e
ordinal: DEFINITION of < and le to replace : and <= on ordinals!  Many
 lcp parents: 
14diff
changeset | 348 | by (REPEAT (ares_tac [add_le_self] 1)); | 
| 760 | 349 | qed "add_le_self2"; | 
| 14 
1c0926788772
ex/{bin.ML,comb.ML,prop.ML}: replaced NewSext by Syntax.simple_sext
 lcp parents: 
6diff
changeset | 350 | |
| 
1c0926788772
ex/{bin.ML,comb.ML,prop.ML}: replaced NewSext by Syntax.simple_sext
 lcp parents: 
6diff
changeset | 351 | (** Monotonicity of addition **) | 
| 
1c0926788772
ex/{bin.ML,comb.ML,prop.ML}: replaced NewSext by Syntax.simple_sext
 lcp parents: 
6diff
changeset | 352 | |
| 
1c0926788772
ex/{bin.ML,comb.ML,prop.ML}: replaced NewSext by Syntax.simple_sext
 lcp parents: 
6diff
changeset | 353 | (*strict, in 1st argument*) | 
| 25 
3ac1c0c0016e
ordinal: DEFINITION of < and le to replace : and <= on ordinals!  Many
 lcp parents: 
14diff
changeset | 354 | goal Arith.thy "!!i j k. [| i<j; j:nat; k:nat |] ==> i#+k < j#+k"; | 
| 
3ac1c0c0016e
ordinal: DEFINITION of < and le to replace : and <= on ordinals!  Many
 lcp parents: 
14diff
changeset | 355 | by (forward_tac [lt_nat_in_nat] 1); | 
| 127 | 356 | by (assume_tac 1); | 
| 25 
3ac1c0c0016e
ordinal: DEFINITION of < and le to replace : and <= on ordinals!  Many
 lcp parents: 
14diff
changeset | 357 | by (etac succ_lt_induct 1); | 
| 
3ac1c0c0016e
ordinal: DEFINITION of < and le to replace : and <= on ordinals!  Many
 lcp parents: 
14diff
changeset | 358 | by (ALLGOALS (asm_simp_tac (arith_ss addsimps [leI]))); | 
| 760 | 359 | qed "add_lt_mono1"; | 
| 14 
1c0926788772
ex/{bin.ML,comb.ML,prop.ML}: replaced NewSext by Syntax.simple_sext
 lcp parents: 
6diff
changeset | 360 | |
| 
1c0926788772
ex/{bin.ML,comb.ML,prop.ML}: replaced NewSext by Syntax.simple_sext
 lcp parents: 
6diff
changeset | 361 | (*strict, in both arguments*) | 
| 25 
3ac1c0c0016e
ordinal: DEFINITION of < and le to replace : and <= on ordinals!  Many
 lcp parents: 
14diff
changeset | 362 | goal Arith.thy "!!i j k l. [| i<j; k<l; j:nat; l:nat |] ==> i#+k < j#+l"; | 
| 
3ac1c0c0016e
ordinal: DEFINITION of < and le to replace : and <= on ordinals!  Many
 lcp parents: 
14diff
changeset | 363 | by (rtac (add_lt_mono1 RS lt_trans) 1); | 
| 
3ac1c0c0016e
ordinal: DEFINITION of < and le to replace : and <= on ordinals!  Many
 lcp parents: 
14diff
changeset | 364 | by (REPEAT (eresolve_tac [asm_rl, lt_nat_in_nat] 1)); | 
| 14 
1c0926788772
ex/{bin.ML,comb.ML,prop.ML}: replaced NewSext by Syntax.simple_sext
 lcp parents: 
6diff
changeset | 365 | by (EVERY [rtac (add_commute RS ssubst) 1, | 
| 
1c0926788772
ex/{bin.ML,comb.ML,prop.ML}: replaced NewSext by Syntax.simple_sext
 lcp parents: 
6diff
changeset | 366 | rtac (add_commute RS ssubst) 3, | 
| 25 
3ac1c0c0016e
ordinal: DEFINITION of < and le to replace : and <= on ordinals!  Many
 lcp parents: 
14diff
changeset | 367 | rtac add_lt_mono1 5]); | 
| 
3ac1c0c0016e
ordinal: DEFINITION of < and le to replace : and <= on ordinals!  Many
 lcp parents: 
14diff
changeset | 368 | by (REPEAT (eresolve_tac [asm_rl, lt_nat_in_nat] 1)); | 
| 760 | 369 | qed "add_lt_mono"; | 
| 14 
1c0926788772
ex/{bin.ML,comb.ML,prop.ML}: replaced NewSext by Syntax.simple_sext
 lcp parents: 
6diff
changeset | 370 | |
| 25 
3ac1c0c0016e
ordinal: DEFINITION of < and le to replace : and <= on ordinals!  Many
 lcp parents: 
14diff
changeset | 371 | (*A [clumsy] way of lifting < monotonicity to le monotonicity *) | 
| 435 | 372 | val lt_mono::ford::prems = goal Ordinal.thy | 
| 25 
3ac1c0c0016e
ordinal: DEFINITION of < and le to replace : and <= on ordinals!  Many
 lcp parents: 
14diff
changeset | 373 | "[| !!i j. [| i<j; j:k |] ==> f(i) < f(j); \ | 
| 
3ac1c0c0016e
ordinal: DEFINITION of < and le to replace : and <= on ordinals!  Many
 lcp parents: 
14diff
changeset | 374 | \ !!i. i:k ==> Ord(f(i)); \ | 
| 
3ac1c0c0016e
ordinal: DEFINITION of < and le to replace : and <= on ordinals!  Many
 lcp parents: 
14diff
changeset | 375 | \ i le j; j:k \ | 
| 
3ac1c0c0016e
ordinal: DEFINITION of < and le to replace : and <= on ordinals!  Many
 lcp parents: 
14diff
changeset | 376 | \ |] ==> f(i) le f(j)"; | 
| 14 
1c0926788772
ex/{bin.ML,comb.ML,prop.ML}: replaced NewSext by Syntax.simple_sext
 lcp parents: 
6diff
changeset | 377 | by (cut_facts_tac prems 1); | 
| 25 
3ac1c0c0016e
ordinal: DEFINITION of < and le to replace : and <= on ordinals!  Many
 lcp parents: 
14diff
changeset | 378 | by (fast_tac (lt_cs addSIs [lt_mono,ford] addSEs [leE]) 1); | 
| 760 | 379 | qed "Ord_lt_mono_imp_le_mono"; | 
| 14 
1c0926788772
ex/{bin.ML,comb.ML,prop.ML}: replaced NewSext by Syntax.simple_sext
 lcp parents: 
6diff
changeset | 380 | |
| 25 
3ac1c0c0016e
ordinal: DEFINITION of < and le to replace : and <= on ordinals!  Many
 lcp parents: 
14diff
changeset | 381 | (*le monotonicity, 1st argument*) | 
| 14 
1c0926788772
ex/{bin.ML,comb.ML,prop.ML}: replaced NewSext by Syntax.simple_sext
 lcp parents: 
6diff
changeset | 382 | goal Arith.thy | 
| 25 
3ac1c0c0016e
ordinal: DEFINITION of < and le to replace : and <= on ordinals!  Many
 lcp parents: 
14diff
changeset | 383 | "!!i j k. [| i le j; j:nat; k:nat |] ==> i#+k le j#+k"; | 
| 
3ac1c0c0016e
ordinal: DEFINITION of < and le to replace : and <= on ordinals!  Many
 lcp parents: 
14diff
changeset | 384 | by (res_inst_tac [("f", "%j.j#+k")] Ord_lt_mono_imp_le_mono 1);
 | 
| 435 | 385 | by (REPEAT (ares_tac [add_lt_mono1, add_type RS nat_into_Ord] 1)); | 
| 760 | 386 | qed "add_le_mono1"; | 
| 14 
1c0926788772
ex/{bin.ML,comb.ML,prop.ML}: replaced NewSext by Syntax.simple_sext
 lcp parents: 
6diff
changeset | 387 | |
| 25 
3ac1c0c0016e
ordinal: DEFINITION of < and le to replace : and <= on ordinals!  Many
 lcp parents: 
14diff
changeset | 388 | (* le monotonicity, BOTH arguments*) | 
| 14 
1c0926788772
ex/{bin.ML,comb.ML,prop.ML}: replaced NewSext by Syntax.simple_sext
 lcp parents: 
6diff
changeset | 389 | goal Arith.thy | 
| 25 
3ac1c0c0016e
ordinal: DEFINITION of < and le to replace : and <= on ordinals!  Many
 lcp parents: 
14diff
changeset | 390 | "!!i j k. [| i le j; k le l; j:nat; l:nat |] ==> i#+k le j#+l"; | 
| 
3ac1c0c0016e
ordinal: DEFINITION of < and le to replace : and <= on ordinals!  Many
 lcp parents: 
14diff
changeset | 391 | by (rtac (add_le_mono1 RS le_trans) 1); | 
| 
3ac1c0c0016e
ordinal: DEFINITION of < and le to replace : and <= on ordinals!  Many
 lcp parents: 
14diff
changeset | 392 | by (REPEAT (eresolve_tac [asm_rl, lt_nat_in_nat, nat_succI] 1)); | 
| 14 
1c0926788772
ex/{bin.ML,comb.ML,prop.ML}: replaced NewSext by Syntax.simple_sext
 lcp parents: 
6diff
changeset | 393 | by (EVERY [rtac (add_commute RS ssubst) 1, | 
| 
1c0926788772
ex/{bin.ML,comb.ML,prop.ML}: replaced NewSext by Syntax.simple_sext
 lcp parents: 
6diff
changeset | 394 | rtac (add_commute RS ssubst) 3, | 
| 25 
3ac1c0c0016e
ordinal: DEFINITION of < and le to replace : and <= on ordinals!  Many
 lcp parents: 
14diff
changeset | 395 | rtac add_le_mono1 5]); | 
| 
3ac1c0c0016e
ordinal: DEFINITION of < and le to replace : and <= on ordinals!  Many
 lcp parents: 
14diff
changeset | 396 | by (REPEAT (eresolve_tac [asm_rl, lt_nat_in_nat, nat_succI] 1)); | 
| 760 | 397 | qed "add_le_mono"; |