author | desharna |
Fri, 04 Apr 2025 15:30:03 +0200 | |
changeset 82399 | 9d457dfb56c5 |
parent 78791 | 4f7dce5c1a81 |
permissions | -rw-r--r-- |
41777 | 1 |
(* Title: ZF/ArithSimp.thy |
9548 | 2 |
Author: Lawrence C Paulson, Cambridge University Computer Laboratory |
3 |
Copyright 2000 University of Cambridge |
|
4 |
*) |
|
5 |
||
60770 | 6 |
section\<open>Arithmetic with simplification\<close> |
13328 | 7 |
|
46820 | 8 |
theory ArithSimp |
15481 | 9 |
imports Arith |
10 |
begin |
|
13259 | 11 |
|
78791 | 12 |
subsection \<open>Arithmetic simplification\<close> |
13 |
||
69605 | 14 |
ML_file \<open>~~/src/Provers/Arith/cancel_numerals.ML\<close> |
15 |
ML_file \<open>~~/src/Provers/Arith/combine_numerals.ML\<close> |
|
16 |
ML_file \<open>arith_data.ML\<close> |
|
48891 | 17 |
|
78791 | 18 |
simproc_setup nateq_cancel_numerals |
19 |
("l #+ m = n" | "l = m #+ n" | "l #* m = n" | "l = m #* n" | "succ(m) = n" | "m = succ(n)") = |
|
20 |
\<open>K ArithData.nateq_cancel_numerals_proc\<close> |
|
21 |
||
22 |
simproc_setup natless_cancel_numerals |
|
23 |
("l #+ m < n" | "l < m #+ n" | "l #* m < n" | "l < m #* n" | "succ(m) < n" | "m < succ(n)") = |
|
24 |
\<open>K ArithData.natless_cancel_numerals_proc\<close> |
|
25 |
||
26 |
simproc_setup natdiff_cancel_numerals |
|
27 |
("(l #+ m) #- n" | "l #- (m #+ n)" | "(l #* m) #- n" | "l #- (m #* n)" | |
|
28 |
"succ(m) #- n" | "m #- succ(n)") = |
|
29 |
\<open>K ArithData.natdiff_cancel_numerals_proc\<close> |
|
30 |
||
31 |
||
32 |
subsubsection \<open>Examples\<close> |
|
33 |
||
34 |
lemma "x #+ y = x #+ z" apply simp oops |
|
35 |
lemma "y #+ x = x #+ z" apply simp oops |
|
36 |
lemma "x #+ y #+ z = x #+ z" apply simp oops |
|
37 |
lemma "y #+ (z #+ x) = z #+ x" apply simp oops |
|
38 |
lemma "x #+ y #+ z = (z #+ y) #+ (x #+ w)" apply simp oops |
|
39 |
lemma "x#*y #+ z = (z #+ y) #+ (y#*x #+ w)" apply simp oops |
|
40 |
||
41 |
lemma "x #+ succ(y) = x #+ z" apply simp oops |
|
42 |
lemma "x #+ succ(y) = succ(z #+ x)" apply simp oops |
|
43 |
lemma "succ(x) #+ succ(y) #+ z = succ(z #+ y) #+ succ(x #+ w)" apply simp oops |
|
44 |
||
45 |
lemma "(x #+ y) #- (x #+ z) = w" apply simp oops |
|
46 |
lemma "(y #+ x) #- (x #+ z) = dd" apply simp oops |
|
47 |
lemma "(x #+ y #+ z) #- (x #+ z) = dd" apply simp oops |
|
48 |
lemma "(y #+ (z #+ x)) #- (z #+ x) = dd" apply simp oops |
|
49 |
lemma "(x #+ y #+ z) #- ((z #+ y) #+ (x #+ w)) = dd" apply simp oops |
|
50 |
lemma "(x#*y #+ z) #- ((z #+ y) #+ (y#*x #+ w)) = dd" apply simp oops |
|
51 |
||
52 |
(*BAD occurrence of natify*) |
|
53 |
lemma "(x #+ succ(y)) #- (x #+ z) = dd" apply simp oops |
|
54 |
||
55 |
lemma "x #* y2 #+ y #* x2 = y #* x2 #+ x #* y2" apply simp oops |
|
56 |
||
57 |
lemma "(x #+ succ(y)) #- (succ(z #+ x)) = dd" apply simp oops |
|
58 |
lemma "(succ(x) #+ succ(y) #+ z) #- (succ(z #+ y) #+ succ(x #+ w)) = dd" apply simp oops |
|
59 |
||
60 |
(*use of typing information*) |
|
61 |
lemma "x : nat ==> x #+ y = x" apply simp oops |
|
62 |
lemma "x : nat --> x #+ y = x" apply simp oops |
|
63 |
lemma "x : nat ==> x #+ y < x" apply simp oops |
|
64 |
lemma "x : nat ==> x < y#+x" apply simp oops |
|
65 |
lemma "x : nat ==> x \<le> succ(x)" apply simp oops |
|
66 |
||
67 |
(*fails: no typing information isn't visible*) |
|
68 |
lemma "x #+ y = x" apply simp? oops |
|
69 |
||
70 |
lemma "x #+ y < x #+ z" apply simp oops |
|
71 |
lemma "y #+ x < x #+ z" apply simp oops |
|
72 |
lemma "x #+ y #+ z < x #+ z" apply simp oops |
|
73 |
lemma "y #+ z #+ x < x #+ z" apply simp oops |
|
74 |
lemma "y #+ (z #+ x) < z #+ x" apply simp oops |
|
75 |
lemma "x #+ y #+ z < (z #+ y) #+ (x #+ w)" apply simp oops |
|
76 |
lemma "x#*y #+ z < (z #+ y) #+ (y#*x #+ w)" apply simp oops |
|
77 |
||
78 |
lemma "x #+ succ(y) < x #+ z" apply simp oops |
|
79 |
lemma "x #+ succ(y) < succ(z #+ x)" apply simp oops |
|
80 |
lemma "succ(x) #+ succ(y) #+ z < succ(z #+ y) #+ succ(x #+ w)" apply simp oops |
|
81 |
||
82 |
lemma "x #+ succ(y) \<le> succ(z #+ x)" apply simp oops |
|
83 |
||
48891 | 84 |
|
60770 | 85 |
subsection\<open>Difference\<close> |
13259 | 86 |
|
14046 | 87 |
lemma diff_self_eq_0 [simp]: "m #- m = 0" |
13259 | 88 |
apply (subgoal_tac "natify (m) #- natify (m) = 0") |
89 |
apply (rule_tac [2] natify_in_nat [THEN nat_induct], auto) |
|
90 |
done |
|
91 |
||
92 |
(**Addition is the inverse of subtraction**) |
|
93 |
||
94 |
(*We need m:nat even if we replace the RHS by natify(m), for consider e.g. |
|
46820 | 95 |
n=2, m=omega; then n + (m-n) = 2 + (0-2) = 2 \<noteq> 0 = natify(m).*) |
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69605
diff
changeset
|
96 |
lemma add_diff_inverse: "\<lbrakk>n \<le> m; m:nat\<rbrakk> \<Longrightarrow> n #+ (m#-n) = m" |
13259 | 97 |
apply (frule lt_nat_in_nat, erule nat_succI) |
98 |
apply (erule rev_mp) |
|
13784 | 99 |
apply (rule_tac m = m and n = n in diff_induct, auto) |
13259 | 100 |
done |
101 |
||
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69605
diff
changeset
|
102 |
lemma add_diff_inverse2: "\<lbrakk>n \<le> m; m:nat\<rbrakk> \<Longrightarrow> (m#-n) #+ n = m" |
13259 | 103 |
apply (frule lt_nat_in_nat, erule nat_succI) |
104 |
apply (simp (no_asm_simp) add: add_commute add_diff_inverse) |
|
105 |
done |
|
106 |
||
107 |
(*Proof is IDENTICAL to that of add_diff_inverse*) |
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69605
diff
changeset
|
108 |
lemma diff_succ: "\<lbrakk>n \<le> m; m:nat\<rbrakk> \<Longrightarrow> succ(m) #- n = succ(m#-n)" |
13259 | 109 |
apply (frule lt_nat_in_nat, erule nat_succI) |
110 |
apply (erule rev_mp) |
|
13784 | 111 |
apply (rule_tac m = m and n = n in diff_induct) |
13259 | 112 |
apply (simp_all (no_asm_simp)) |
113 |
done |
|
114 |
||
115 |
lemma zero_less_diff [simp]: |
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69605
diff
changeset
|
116 |
"\<lbrakk>m: nat; n: nat\<rbrakk> \<Longrightarrow> 0 < (n #- m) \<longleftrightarrow> m<n" |
13784 | 117 |
apply (rule_tac m = m and n = n in diff_induct) |
13259 | 118 |
apply (simp_all (no_asm_simp)) |
119 |
done |
|
120 |
||
121 |
||
122 |
(** Difference distributes over multiplication **) |
|
123 |
||
124 |
lemma diff_mult_distrib: "(m #- n) #* k = (m #* k) #- (n #* k)" |
|
125 |
apply (subgoal_tac " (natify (m) #- natify (n)) #* natify (k) = (natify (m) #* natify (k)) #- (natify (n) #* natify (k))") |
|
126 |
apply (rule_tac [2] m = "natify (m) " and n = "natify (n) " in diff_induct) |
|
127 |
apply (simp_all add: diff_cancel) |
|
128 |
done |
|
129 |
||
130 |
lemma diff_mult_distrib2: "k #* (m #- n) = (k #* m) #- (k #* n)" |
|
131 |
apply (simp (no_asm) add: mult_commute [of k] diff_mult_distrib) |
|
132 |
done |
|
133 |
||
134 |
||
60770 | 135 |
subsection\<open>Remainder\<close> |
13259 | 136 |
|
137 |
(*We need m:nat even with natify*) |
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69605
diff
changeset
|
138 |
lemma div_termination: "\<lbrakk>0<n; n \<le> m; m:nat\<rbrakk> \<Longrightarrow> m #- n < m" |
13259 | 139 |
apply (frule lt_nat_in_nat, erule nat_succI) |
140 |
apply (erule rev_mp) |
|
141 |
apply (erule rev_mp) |
|
13784 | 142 |
apply (rule_tac m = m and n = n in diff_induct) |
13259 | 143 |
apply (simp_all (no_asm_simp) add: diff_le_self) |
144 |
done |
|
145 |
||
146 |
(*for mod and div*) |
|
46820 | 147 |
lemmas div_rls = |
148 |
nat_typechecks Ord_transrec_type apply_funtype |
|
13259 | 149 |
div_termination [THEN ltD] |
150 |
nat_into_Ord not_lt_iff_le [THEN iffD1] |
|
151 |
||
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69605
diff
changeset
|
152 |
lemma raw_mod_type: "\<lbrakk>m:nat; n:nat\<rbrakk> \<Longrightarrow> raw_mod (m, n) \<in> nat" |
76216
9fc34f76b4e8
getting rid of apply (unfold ...)
paulson <lp15@cam.ac.uk>
parents:
76214
diff
changeset
|
153 |
unfolding raw_mod_def |
13259 | 154 |
apply (rule Ord_transrec_type) |
155 |
apply (auto simp add: nat_into_Ord [THEN Ord_0_lt_iff]) |
|
46820 | 156 |
apply (blast intro: div_rls) |
13259 | 157 |
done |
158 |
||
46820 | 159 |
lemma mod_type [TC,iff]: "m mod n \<in> nat" |
76216
9fc34f76b4e8
getting rid of apply (unfold ...)
paulson <lp15@cam.ac.uk>
parents:
76214
diff
changeset
|
160 |
unfolding mod_def |
13259 | 161 |
apply (simp (no_asm) add: mod_def raw_mod_type) |
162 |
done |
|
163 |
||
164 |
||
46820 | 165 |
(** Aribtrary definitions for division by zero. Useful to simplify |
13259 | 166 |
certain equations **) |
167 |
||
168 |
lemma DIVISION_BY_ZERO_DIV: "a div 0 = 0" |
|
76216
9fc34f76b4e8
getting rid of apply (unfold ...)
paulson <lp15@cam.ac.uk>
parents:
76214
diff
changeset
|
169 |
unfolding div_def |
13259 | 170 |
apply (rule raw_div_def [THEN def_transrec, THEN trans]) |
171 |
apply (simp (no_asm_simp)) |
|
172 |
done (*NOT for adding to default simpset*) |
|
173 |
||
174 |
lemma DIVISION_BY_ZERO_MOD: "a mod 0 = natify(a)" |
|
76216
9fc34f76b4e8
getting rid of apply (unfold ...)
paulson <lp15@cam.ac.uk>
parents:
76214
diff
changeset
|
175 |
unfolding mod_def |
13259 | 176 |
apply (rule raw_mod_def [THEN def_transrec, THEN trans]) |
177 |
apply (simp (no_asm_simp)) |
|
178 |
done (*NOT for adding to default simpset*) |
|
179 |
||
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69605
diff
changeset
|
180 |
lemma raw_mod_less: "m<n \<Longrightarrow> raw_mod (m,n) = m" |
13259 | 181 |
apply (rule raw_mod_def [THEN def_transrec, THEN trans]) |
182 |
apply (simp (no_asm_simp) add: div_termination [THEN ltD]) |
|
183 |
done |
|
184 |
||
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69605
diff
changeset
|
185 |
lemma mod_less [simp]: "\<lbrakk>m<n; n \<in> nat\<rbrakk> \<Longrightarrow> m mod n = m" |
13259 | 186 |
apply (frule lt_nat_in_nat, assumption) |
187 |
apply (simp (no_asm_simp) add: mod_def raw_mod_less) |
|
188 |
done |
|
189 |
||
190 |
lemma raw_mod_geq: |
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69605
diff
changeset
|
191 |
"\<lbrakk>0<n; n \<le> m; m:nat\<rbrakk> \<Longrightarrow> raw_mod (m, n) = raw_mod (m#-n, n)" |
13259 | 192 |
apply (frule lt_nat_in_nat, erule nat_succI) |
193 |
apply (rule raw_mod_def [THEN def_transrec, THEN trans]) |
|
13611 | 194 |
apply (simp (no_asm_simp) add: div_termination [THEN ltD] not_lt_iff_le [THEN iffD2], blast) |
13259 | 195 |
done |
196 |
||
197 |
||
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69605
diff
changeset
|
198 |
lemma mod_geq: "\<lbrakk>n \<le> m; m:nat\<rbrakk> \<Longrightarrow> m mod n = (m#-n) mod n" |
13259 | 199 |
apply (frule lt_nat_in_nat, erule nat_succI) |
200 |
apply (case_tac "n=0") |
|
201 |
apply (simp add: DIVISION_BY_ZERO_MOD) |
|
202 |
apply (simp add: mod_def raw_mod_geq nat_into_Ord [THEN Ord_0_lt_iff]) |
|
203 |
done |
|
204 |
||
205 |
||
60770 | 206 |
subsection\<open>Division\<close> |
13259 | 207 |
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69605
diff
changeset
|
208 |
lemma raw_div_type: "\<lbrakk>m:nat; n:nat\<rbrakk> \<Longrightarrow> raw_div (m, n) \<in> nat" |
76216
9fc34f76b4e8
getting rid of apply (unfold ...)
paulson <lp15@cam.ac.uk>
parents:
76214
diff
changeset
|
209 |
unfolding raw_div_def |
13259 | 210 |
apply (rule Ord_transrec_type) |
211 |
apply (auto simp add: nat_into_Ord [THEN Ord_0_lt_iff]) |
|
46820 | 212 |
apply (blast intro: div_rls) |
13259 | 213 |
done |
214 |
||
46820 | 215 |
lemma div_type [TC,iff]: "m div n \<in> nat" |
76216
9fc34f76b4e8
getting rid of apply (unfold ...)
paulson <lp15@cam.ac.uk>
parents:
76214
diff
changeset
|
216 |
unfolding div_def |
13259 | 217 |
apply (simp (no_asm) add: div_def raw_div_type) |
218 |
done |
|
219 |
||
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69605
diff
changeset
|
220 |
lemma raw_div_less: "m<n \<Longrightarrow> raw_div (m,n) = 0" |
13259 | 221 |
apply (rule raw_div_def [THEN def_transrec, THEN trans]) |
222 |
apply (simp (no_asm_simp) add: div_termination [THEN ltD]) |
|
223 |
done |
|
224 |
||
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69605
diff
changeset
|
225 |
lemma div_less [simp]: "\<lbrakk>m<n; n \<in> nat\<rbrakk> \<Longrightarrow> m div n = 0" |
13259 | 226 |
apply (frule lt_nat_in_nat, assumption) |
227 |
apply (simp (no_asm_simp) add: div_def raw_div_less) |
|
228 |
done |
|
229 |
||
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69605
diff
changeset
|
230 |
lemma raw_div_geq: "\<lbrakk>0<n; n \<le> m; m:nat\<rbrakk> \<Longrightarrow> raw_div(m,n) = succ(raw_div(m#-n, n))" |
46820 | 231 |
apply (subgoal_tac "n \<noteq> 0") |
13259 | 232 |
prefer 2 apply blast |
233 |
apply (frule lt_nat_in_nat, erule nat_succI) |
|
234 |
apply (rule raw_div_def [THEN def_transrec, THEN trans]) |
|
46820 | 235 |
apply (simp (no_asm_simp) add: div_termination [THEN ltD] not_lt_iff_le [THEN iffD2] ) |
13259 | 236 |
done |
237 |
||
238 |
lemma div_geq [simp]: |
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69605
diff
changeset
|
239 |
"\<lbrakk>0<n; n \<le> m; m:nat\<rbrakk> \<Longrightarrow> m div n = succ ((m#-n) div n)" |
13259 | 240 |
apply (frule lt_nat_in_nat, erule nat_succI) |
241 |
apply (simp (no_asm_simp) add: div_def raw_div_geq) |
|
242 |
done |
|
243 |
||
244 |
declare div_less [simp] div_geq [simp] |
|
245 |
||
246 |
||
247 |
(*A key result*) |
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69605
diff
changeset
|
248 |
lemma mod_div_lemma: "\<lbrakk>m: nat; n: nat\<rbrakk> \<Longrightarrow> (m div n)#*n #+ m mod n = m" |
13259 | 249 |
apply (case_tac "n=0") |
250 |
apply (simp add: DIVISION_BY_ZERO_MOD) |
|
251 |
apply (simp add: nat_into_Ord [THEN Ord_0_lt_iff]) |
|
252 |
apply (erule complete_induct) |
|
253 |
apply (case_tac "x<n") |
|
60770 | 254 |
txt\<open>case x<n\<close> |
13259 | 255 |
apply (simp (no_asm_simp)) |
69593 | 256 |
txt\<open>case \<^term>\<open>n \<le> x\<close>\<close> |
13259 | 257 |
apply (simp add: not_lt_iff_le add_assoc mod_geq div_termination [THEN ltD] add_diff_inverse) |
258 |
done |
|
259 |
||
260 |
lemma mod_div_equality_natify: "(m div n)#*n #+ m mod n = natify(m)" |
|
261 |
apply (subgoal_tac " (natify (m) div natify (n))#*natify (n) #+ natify (m) mod natify (n) = natify (m) ") |
|
46820 | 262 |
apply force |
13259 | 263 |
apply (subst mod_div_lemma, auto) |
264 |
done |
|
265 |
||
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69605
diff
changeset
|
266 |
lemma mod_div_equality: "m: nat \<Longrightarrow> (m div n)#*n #+ m mod n = m" |
13259 | 267 |
apply (simp (no_asm_simp) add: mod_div_equality_natify) |
268 |
done |
|
269 |
||
270 |
||
60770 | 271 |
subsection\<open>Further Facts about Remainder\<close> |
13356 | 272 |
|
60770 | 273 |
text\<open>(mainly for mutilated chess board)\<close> |
13259 | 274 |
|
275 |
lemma mod_succ_lemma: |
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69605
diff
changeset
|
276 |
"\<lbrakk>0<n; m:nat; n:nat\<rbrakk> |
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69605
diff
changeset
|
277 |
\<Longrightarrow> succ(m) mod n = (if succ(m mod n) = n then 0 else succ(m mod n))" |
13259 | 278 |
apply (erule complete_induct) |
279 |
apply (case_tac "succ (x) <n") |
|
60770 | 280 |
txt\<open>case succ(x) < n\<close> |
13259 | 281 |
apply (simp (no_asm_simp) add: nat_le_refl [THEN lt_trans] succ_neq_self) |
282 |
apply (simp add: ltD [THEN mem_imp_not_eq]) |
|
69593 | 283 |
txt\<open>case \<^term>\<open>n \<le> succ(x)\<close>\<close> |
13259 | 284 |
apply (simp add: mod_geq not_lt_iff_le) |
285 |
apply (erule leE) |
|
286 |
apply (simp (no_asm_simp) add: mod_geq div_termination [THEN ltD] diff_succ) |
|
60770 | 287 |
txt\<open>equality case\<close> |
13259 | 288 |
apply (simp add: diff_self_eq_0) |
289 |
done |
|
290 |
||
291 |
lemma mod_succ: |
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69605
diff
changeset
|
292 |
"n:nat \<Longrightarrow> succ(m) mod n = (if succ(m mod n) = n then 0 else succ(m mod n))" |
13259 | 293 |
apply (case_tac "n=0") |
294 |
apply (simp (no_asm_simp) add: natify_succ DIVISION_BY_ZERO_MOD) |
|
295 |
apply (subgoal_tac "natify (succ (m)) mod n = (if succ (natify (m) mod n) = n then 0 else succ (natify (m) mod n))") |
|
296 |
prefer 2 |
|
297 |
apply (subst natify_succ) |
|
298 |
apply (rule mod_succ_lemma) |
|
299 |
apply (auto simp del: natify_succ simp add: nat_into_Ord [THEN Ord_0_lt_iff]) |
|
300 |
done |
|
301 |
||
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69605
diff
changeset
|
302 |
lemma mod_less_divisor: "\<lbrakk>0<n; n:nat\<rbrakk> \<Longrightarrow> m mod n < n" |
13259 | 303 |
apply (subgoal_tac "natify (m) mod n < n") |
304 |
apply (rule_tac [2] i = "natify (m) " in complete_induct) |
|
46820 | 305 |
apply (case_tac [3] "x<n", auto) |
69593 | 306 |
txt\<open>case \<^term>\<open>n \<le> x\<close>\<close> |
13259 | 307 |
apply (simp add: mod_geq not_lt_iff_le div_termination [THEN ltD]) |
308 |
done |
|
309 |
||
310 |
lemma mod_1_eq [simp]: "m mod 1 = 0" |
|
13784 | 311 |
by (cut_tac n = 1 in mod_less_divisor, auto) |
13259 | 312 |
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69605
diff
changeset
|
313 |
lemma mod2_cases: "b<2 \<Longrightarrow> k mod 2 = b | k mod 2 = (if b=1 then 0 else 1)" |
13259 | 314 |
apply (subgoal_tac "k mod 2: 2") |
315 |
prefer 2 apply (simp add: mod_less_divisor [THEN ltD]) |
|
316 |
apply (drule ltD, auto) |
|
317 |
done |
|
318 |
||
319 |
lemma mod2_succ_succ [simp]: "succ(succ(m)) mod 2 = m mod 2" |
|
320 |
apply (subgoal_tac "m mod 2: 2") |
|
321 |
prefer 2 apply (simp add: mod_less_divisor [THEN ltD]) |
|
322 |
apply (auto simp add: mod_succ) |
|
323 |
done |
|
324 |
||
325 |
lemma mod2_add_more [simp]: "(m#+m#+n) mod 2 = n mod 2" |
|
326 |
apply (subgoal_tac " (natify (m) #+natify (m) #+n) mod 2 = n mod 2") |
|
327 |
apply (rule_tac [2] n = "natify (m) " in nat_induct) |
|
328 |
apply auto |
|
329 |
done |
|
330 |
||
331 |
lemma mod2_add_self [simp]: "(m#+m) mod 2 = 0" |
|
13784 | 332 |
by (cut_tac n = 0 in mod2_add_more, auto) |
13259 | 333 |
|
334 |
||
61798 | 335 |
subsection\<open>Additional theorems about \<open>\<le>\<close>\<close> |
13259 | 336 |
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69605
diff
changeset
|
337 |
lemma add_le_self: "m:nat \<Longrightarrow> m \<le> (m #+ n)" |
13259 | 338 |
apply (simp (no_asm_simp)) |
339 |
done |
|
340 |
||
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69605
diff
changeset
|
341 |
lemma add_le_self2: "m:nat \<Longrightarrow> m \<le> (n #+ m)" |
13259 | 342 |
apply (simp (no_asm_simp)) |
343 |
done |
|
344 |
||
345 |
(*** Monotonicity of Multiplication ***) |
|
346 |
||
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69605
diff
changeset
|
347 |
lemma mult_le_mono1: "\<lbrakk>i \<le> j; j:nat\<rbrakk> \<Longrightarrow> (i#*k) \<le> (j#*k)" |
46820 | 348 |
apply (subgoal_tac "natify (i) #*natify (k) \<le> j#*natify (k) ") |
13259 | 349 |
apply (frule_tac [2] lt_nat_in_nat) |
350 |
apply (rule_tac [3] n = "natify (k) " in nat_induct) |
|
351 |
apply (simp_all add: add_le_mono) |
|
352 |
done |
|
353 |
||
46820 | 354 |
(* @{text"\<le>"} monotonicity, BOTH arguments*) |
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69605
diff
changeset
|
355 |
lemma mult_le_mono: "\<lbrakk>i \<le> j; k \<le> l; j:nat; l:nat\<rbrakk> \<Longrightarrow> i#*k \<le> j#*l" |
13259 | 356 |
apply (rule mult_le_mono1 [THEN le_trans], assumption+) |
357 |
apply (subst mult_commute, subst mult_commute, rule mult_le_mono1, assumption+) |
|
358 |
done |
|
359 |
||
360 |
(*strict, in 1st argument; proof is by induction on k>0. |
|
361 |
I can't see how to relax the typing conditions.*) |
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69605
diff
changeset
|
362 |
lemma mult_lt_mono2: "\<lbrakk>i<j; 0<k; j:nat; k:nat\<rbrakk> \<Longrightarrow> k#*i < k#*j" |
13259 | 363 |
apply (erule zero_lt_natE) |
364 |
apply (frule_tac [2] lt_nat_in_nat) |
|
365 |
apply (simp_all (no_asm_simp)) |
|
366 |
apply (induct_tac "x") |
|
367 |
apply (simp_all (no_asm_simp) add: add_lt_mono) |
|
368 |
done |
|
369 |
||
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69605
diff
changeset
|
370 |
lemma mult_lt_mono1: "\<lbrakk>i<j; 0<k; j:nat; k:nat\<rbrakk> \<Longrightarrow> i#*k < j#*k" |
13259 | 371 |
apply (simp (no_asm_simp) add: mult_lt_mono2 mult_commute [of _ k]) |
372 |
done |
|
373 |
||
76214 | 374 |
lemma add_eq_0_iff [iff]: "m#+n = 0 \<longleftrightarrow> natify(m)=0 \<and> natify(n)=0" |
375 |
apply (subgoal_tac "natify (m) #+ natify (n) = 0 \<longleftrightarrow> natify (m) =0 \<and> natify (n) =0") |
|
13259 | 376 |
apply (rule_tac [2] n = "natify (m) " in natE) |
377 |
apply (rule_tac [4] n = "natify (n) " in natE) |
|
378 |
apply auto |
|
379 |
done |
|
380 |
||
76214 | 381 |
lemma zero_lt_mult_iff [iff]: "0 < m#*n \<longleftrightarrow> 0 < natify(m) \<and> 0 < natify(n)" |
382 |
apply (subgoal_tac "0 < natify (m) #*natify (n) \<longleftrightarrow> 0 < natify (m) \<and> 0 < natify (n) ") |
|
13259 | 383 |
apply (rule_tac [2] n = "natify (m) " in natE) |
384 |
apply (rule_tac [4] n = "natify (n) " in natE) |
|
385 |
apply (rule_tac [3] n = "natify (n) " in natE) |
|
386 |
apply auto |
|
387 |
done |
|
388 |
||
76214 | 389 |
lemma mult_eq_1_iff [iff]: "m#*n = 1 \<longleftrightarrow> natify(m)=1 \<and> natify(n)=1" |
390 |
apply (subgoal_tac "natify (m) #* natify (n) = 1 \<longleftrightarrow> natify (m) =1 \<and> natify (n) =1") |
|
13259 | 391 |
apply (rule_tac [2] n = "natify (m) " in natE) |
392 |
apply (rule_tac [4] n = "natify (n) " in natE) |
|
393 |
apply auto |
|
394 |
done |
|
395 |
||
396 |
||
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69605
diff
changeset
|
397 |
lemma mult_is_zero: "\<lbrakk>m: nat; n: nat\<rbrakk> \<Longrightarrow> (m #* n = 0) \<longleftrightarrow> (m = 0 | n = 0)" |
13259 | 398 |
apply auto |
399 |
apply (erule natE) |
|
400 |
apply (erule_tac [2] natE, auto) |
|
401 |
done |
|
402 |
||
403 |
lemma mult_is_zero_natify [iff]: |
|
46821
ff6b0c1087f2
Using mathematical notation for <-> and cardinal arithmetic
paulson
parents:
46820
diff
changeset
|
404 |
"(m #* n = 0) \<longleftrightarrow> (natify(m) = 0 | natify(n) = 0)" |
13259 | 405 |
apply (cut_tac m = "natify (m) " and n = "natify (n) " in mult_is_zero) |
406 |
apply auto |
|
407 |
done |
|
408 |
||
409 |
||
60770 | 410 |
subsection\<open>Cancellation Laws for Common Factors in Comparisons\<close> |
13259 | 411 |
|
412 |
lemma mult_less_cancel_lemma: |
|
76214 | 413 |
"\<lbrakk>k: nat; m: nat; n: nat\<rbrakk> \<Longrightarrow> (m#*k < n#*k) \<longleftrightarrow> (0<k \<and> m<n)" |
13259 | 414 |
apply (safe intro!: mult_lt_mono1) |
415 |
apply (erule natE, auto) |
|
416 |
apply (rule not_le_iff_lt [THEN iffD1]) |
|
417 |
apply (drule_tac [3] not_le_iff_lt [THEN [2] rev_iffD2]) |
|
418 |
prefer 5 apply (blast intro: mult_le_mono1, auto) |
|
419 |
done |
|
420 |
||
421 |
lemma mult_less_cancel2 [simp]: |
|
76214 | 422 |
"(m#*k < n#*k) \<longleftrightarrow> (0 < natify(k) \<and> natify(m) < natify(n))" |
13259 | 423 |
apply (rule iff_trans) |
424 |
apply (rule_tac [2] mult_less_cancel_lemma, auto) |
|
425 |
done |
|
426 |
||
427 |
lemma mult_less_cancel1 [simp]: |
|
76214 | 428 |
"(k#*m < k#*n) \<longleftrightarrow> (0 < natify(k) \<and> natify(m) < natify(n))" |
13259 | 429 |
apply (simp (no_asm) add: mult_less_cancel2 mult_commute [of k]) |
430 |
done |
|
431 |
||
46821
ff6b0c1087f2
Using mathematical notation for <-> and cardinal arithmetic
paulson
parents:
46820
diff
changeset
|
432 |
lemma mult_le_cancel2 [simp]: "(m#*k \<le> n#*k) \<longleftrightarrow> (0 < natify(k) \<longrightarrow> natify(m) \<le> natify(n))" |
13259 | 433 |
apply (simp (no_asm_simp) add: not_lt_iff_le [THEN iff_sym]) |
434 |
apply auto |
|
435 |
done |
|
436 |
||
46821
ff6b0c1087f2
Using mathematical notation for <-> and cardinal arithmetic
paulson
parents:
46820
diff
changeset
|
437 |
lemma mult_le_cancel1 [simp]: "(k#*m \<le> k#*n) \<longleftrightarrow> (0 < natify(k) \<longrightarrow> natify(m) \<le> natify(n))" |
13259 | 438 |
apply (simp (no_asm_simp) add: not_lt_iff_le [THEN iff_sym]) |
439 |
apply auto |
|
440 |
done |
|
441 |
||
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69605
diff
changeset
|
442 |
lemma mult_le_cancel_le1: "k \<in> nat \<Longrightarrow> k #* m \<le> k \<longleftrightarrow> (0 < k \<longrightarrow> natify(m) \<le> 1)" |
13784 | 443 |
by (cut_tac k = k and m = m and n = 1 in mult_le_cancel1, auto) |
13259 | 444 |
|
76214 | 445 |
lemma Ord_eq_iff_le: "\<lbrakk>Ord(m); Ord(n)\<rbrakk> \<Longrightarrow> m=n \<longleftrightarrow> (m \<le> n \<and> n \<le> m)" |
13259 | 446 |
by (blast intro: le_anti_sym) |
447 |
||
448 |
lemma mult_cancel2_lemma: |
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69605
diff
changeset
|
449 |
"\<lbrakk>k: nat; m: nat; n: nat\<rbrakk> \<Longrightarrow> (m#*k = n#*k) \<longleftrightarrow> (m=n | k=0)" |
13259 | 450 |
apply (simp (no_asm_simp) add: Ord_eq_iff_le [of "m#*k"] Ord_eq_iff_le [of m]) |
451 |
apply (auto simp add: Ord_0_lt_iff) |
|
452 |
done |
|
453 |
||
454 |
lemma mult_cancel2 [simp]: |
|
46821
ff6b0c1087f2
Using mathematical notation for <-> and cardinal arithmetic
paulson
parents:
46820
diff
changeset
|
455 |
"(m#*k = n#*k) \<longleftrightarrow> (natify(m) = natify(n) | natify(k) = 0)" |
13259 | 456 |
apply (rule iff_trans) |
457 |
apply (rule_tac [2] mult_cancel2_lemma, auto) |
|
458 |
done |
|
459 |
||
460 |
lemma mult_cancel1 [simp]: |
|
46821
ff6b0c1087f2
Using mathematical notation for <-> and cardinal arithmetic
paulson
parents:
46820
diff
changeset
|
461 |
"(k#*m = k#*n) \<longleftrightarrow> (natify(m) = natify(n) | natify(k) = 0)" |
13259 | 462 |
apply (simp (no_asm) add: mult_cancel2 mult_commute [of k]) |
463 |
done |
|
464 |
||
465 |
||
466 |
(** Cancellation law for division **) |
|
467 |
||
468 |
lemma div_cancel_raw: |
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69605
diff
changeset
|
469 |
"\<lbrakk>0<n; 0<k; k:nat; m:nat; n:nat\<rbrakk> \<Longrightarrow> (k#*m) div (k#*n) = m div n" |
13784 | 470 |
apply (erule_tac i = m in complete_induct) |
13259 | 471 |
apply (case_tac "x<n") |
472 |
apply (simp add: div_less zero_lt_mult_iff mult_lt_mono2) |
|
473 |
apply (simp add: not_lt_iff_le zero_lt_mult_iff le_refl [THEN mult_le_mono] |
|
474 |
div_geq diff_mult_distrib2 [symmetric] div_termination [THEN ltD]) |
|
475 |
done |
|
476 |
||
477 |
lemma div_cancel: |
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69605
diff
changeset
|
478 |
"\<lbrakk>0 < natify(n); 0 < natify(k)\<rbrakk> \<Longrightarrow> (k#*m) div (k#*n) = m div n" |
46820 | 479 |
apply (cut_tac k = "natify (k) " and m = "natify (m)" and n = "natify (n)" |
13259 | 480 |
in div_cancel_raw) |
481 |
apply auto |
|
482 |
done |
|
483 |
||
484 |
||
60770 | 485 |
subsection\<open>More Lemmas about Remainder\<close> |
13259 | 486 |
|
487 |
lemma mult_mod_distrib_raw: |
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69605
diff
changeset
|
488 |
"\<lbrakk>k:nat; m:nat; n:nat\<rbrakk> \<Longrightarrow> (k#*m) mod (k#*n) = k #* (m mod n)" |
13259 | 489 |
apply (case_tac "k=0") |
490 |
apply (simp add: DIVISION_BY_ZERO_MOD) |
|
491 |
apply (case_tac "n=0") |
|
492 |
apply (simp add: DIVISION_BY_ZERO_MOD) |
|
493 |
apply (simp add: nat_into_Ord [THEN Ord_0_lt_iff]) |
|
13784 | 494 |
apply (erule_tac i = m in complete_induct) |
13259 | 495 |
apply (case_tac "x<n") |
496 |
apply (simp (no_asm_simp) add: mod_less zero_lt_mult_iff mult_lt_mono2) |
|
46820 | 497 |
apply (simp add: not_lt_iff_le zero_lt_mult_iff le_refl [THEN mult_le_mono] |
13259 | 498 |
mod_geq diff_mult_distrib2 [symmetric] div_termination [THEN ltD]) |
499 |
done |
|
500 |
||
501 |
lemma mod_mult_distrib2: "k #* (m mod n) = (k#*m) mod (k#*n)" |
|
46820 | 502 |
apply (cut_tac k = "natify (k) " and m = "natify (m)" and n = "natify (n)" |
13259 | 503 |
in mult_mod_distrib_raw) |
504 |
apply auto |
|
505 |
done |
|
506 |
||
507 |
lemma mult_mod_distrib: "(m mod n) #* k = (m#*k) mod (n#*k)" |
|
508 |
apply (simp (no_asm) add: mult_commute mod_mult_distrib2) |
|
509 |
done |
|
510 |
||
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69605
diff
changeset
|
511 |
lemma mod_add_self2_raw: "n \<in> nat \<Longrightarrow> (m #+ n) mod n = m mod n" |
13259 | 512 |
apply (subgoal_tac " (n #+ m) mod n = (n #+ m #- n) mod n") |
46820 | 513 |
apply (simp add: add_commute) |
514 |
apply (subst mod_geq [symmetric], auto) |
|
13259 | 515 |
done |
516 |
||
517 |
lemma mod_add_self2 [simp]: "(m #+ n) mod n = m mod n" |
|
518 |
apply (cut_tac n = "natify (n) " in mod_add_self2_raw) |
|
519 |
apply auto |
|
520 |
done |
|
521 |
||
522 |
lemma mod_add_self1 [simp]: "(n#+m) mod n = m mod n" |
|
523 |
apply (simp (no_asm_simp) add: add_commute mod_add_self2) |
|
524 |
done |
|
525 |
||
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69605
diff
changeset
|
526 |
lemma mod_mult_self1_raw: "k \<in> nat \<Longrightarrow> (m #+ k#*n) mod n = m mod n" |
13259 | 527 |
apply (erule nat_induct) |
528 |
apply (simp_all (no_asm_simp) add: add_left_commute [of _ n]) |
|
529 |
done |
|
530 |
||
531 |
lemma mod_mult_self1 [simp]: "(m #+ k#*n) mod n = m mod n" |
|
532 |
apply (cut_tac k = "natify (k) " in mod_mult_self1_raw) |
|
533 |
apply auto |
|
534 |
done |
|
535 |
||
536 |
lemma mod_mult_self2 [simp]: "(m #+ n#*k) mod n = m mod n" |
|
537 |
apply (simp (no_asm) add: mult_commute mod_mult_self1) |
|
538 |
done |
|
539 |
||
540 |
(*Lemma for gcd*) |
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69605
diff
changeset
|
541 |
lemma mult_eq_self_implies_10: "m = m#*n \<Longrightarrow> natify(n)=1 | m=0" |
13259 | 542 |
apply (subgoal_tac "m: nat") |
46820 | 543 |
prefer 2 |
13259 | 544 |
apply (erule ssubst) |
46820 | 545 |
apply simp |
13259 | 546 |
apply (rule disjCI) |
547 |
apply (drule sym) |
|
548 |
apply (rule Ord_linear_lt [of "natify(n)" 1]) |
|
46820 | 549 |
apply simp_all |
550 |
apply (subgoal_tac "m #* n = 0", simp) |
|
13259 | 551 |
apply (subst mult_natify2 [symmetric]) |
552 |
apply (simp del: mult_natify2) |
|
553 |
apply (drule nat_into_Ord [THEN Ord_0_lt, THEN [2] mult_lt_mono2], auto) |
|
554 |
done |
|
555 |
||
556 |
lemma less_imp_succ_add [rule_format]: |
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69605
diff
changeset
|
557 |
"\<lbrakk>m<n; n: nat\<rbrakk> \<Longrightarrow> \<exists>k\<in>nat. n = succ(m#+k)" |
13259 | 558 |
apply (frule lt_nat_in_nat, assumption) |
559 |
apply (erule rev_mp) |
|
560 |
apply (induct_tac "n") |
|
561 |
apply (simp_all (no_asm) add: le_iff) |
|
562 |
apply (blast elim!: leE intro!: add_0_right [symmetric] add_succ_right [symmetric]) |
|
563 |
done |
|
564 |
||
565 |
lemma less_iff_succ_add: |
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69605
diff
changeset
|
566 |
"\<lbrakk>m: nat; n: nat\<rbrakk> \<Longrightarrow> (m<n) \<longleftrightarrow> (\<exists>k\<in>nat. n = succ(m#+k))" |
13259 | 567 |
by (auto intro: less_imp_succ_add) |
568 |
||
14055 | 569 |
lemma add_lt_elim2: |
570 |
"\<lbrakk>a #+ d = b #+ c; a < b; b \<in> nat; c \<in> nat; d \<in> nat\<rbrakk> \<Longrightarrow> c < d" |
|
46820 | 571 |
by (drule less_imp_succ_add, auto) |
14055 | 572 |
|
573 |
lemma add_le_elim2: |
|
46820 | 574 |
"\<lbrakk>a #+ d = b #+ c; a \<le> b; b \<in> nat; c \<in> nat; d \<in> nat\<rbrakk> \<Longrightarrow> c \<le> d" |
575 |
by (drule less_imp_succ_add, auto) |
|
14055 | 576 |
|
13356 | 577 |
|
60770 | 578 |
subsubsection\<open>More Lemmas About Difference\<close> |
13259 | 579 |
|
580 |
lemma diff_is_0_lemma: |
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69605
diff
changeset
|
581 |
"\<lbrakk>m: nat; n: nat\<rbrakk> \<Longrightarrow> m #- n = 0 \<longleftrightarrow> m \<le> n" |
13784 | 582 |
apply (rule_tac m = m and n = n in diff_induct, simp_all) |
13259 | 583 |
done |
584 |
||
46821
ff6b0c1087f2
Using mathematical notation for <-> and cardinal arithmetic
paulson
parents:
46820
diff
changeset
|
585 |
lemma diff_is_0_iff: "m #- n = 0 \<longleftrightarrow> natify(m) \<le> natify(n)" |
13259 | 586 |
by (simp add: diff_is_0_lemma [symmetric]) |
587 |
||
588 |
lemma nat_lt_imp_diff_eq_0: |
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69605
diff
changeset
|
589 |
"\<lbrakk>a:nat; b:nat; a<b\<rbrakk> \<Longrightarrow> a #- b = 0" |
46820 | 590 |
by (simp add: diff_is_0_iff le_iff) |
13259 | 591 |
|
14055 | 592 |
lemma raw_nat_diff_split: |
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69605
diff
changeset
|
593 |
"\<lbrakk>a:nat; b:nat\<rbrakk> \<Longrightarrow> |
76214 | 594 |
(P(a #- b)) \<longleftrightarrow> ((a < b \<longrightarrow>P(0)) \<and> (\<forall>d\<in>nat. a = b #+ d \<longrightarrow> P(d)))" |
13259 | 595 |
apply (case_tac "a < b") |
596 |
apply (force simp add: nat_lt_imp_diff_eq_0) |
|
46820 | 597 |
apply (rule iffI, force, simp) |
13259 | 598 |
apply (drule_tac x="a#-b" in bspec) |
46820 | 599 |
apply (simp_all add: Ordinal.not_lt_iff_le add_diff_inverse) |
13259 | 600 |
done |
601 |
||
14055 | 602 |
lemma nat_diff_split: |
46821
ff6b0c1087f2
Using mathematical notation for <-> and cardinal arithmetic
paulson
parents:
46820
diff
changeset
|
603 |
"(P(a #- b)) \<longleftrightarrow> |
76214 | 604 |
(natify(a) < natify(b) \<longrightarrow>P(0)) \<and> (\<forall>d\<in>nat. natify(a) = b #+ d \<longrightarrow> P(d))" |
14055 | 605 |
apply (cut_tac P=P and a="natify(a)" and b="natify(b)" in raw_nat_diff_split) |
606 |
apply simp_all |
|
607 |
done |
|
608 |
||
60770 | 609 |
text\<open>Difference and less-than\<close> |
14060
c0c4af41fa3b
Adding the theory UNITY/AllocImpl.thy, with supporting lemmas
paulson
parents:
14055
diff
changeset
|
610 |
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69605
diff
changeset
|
611 |
lemma diff_lt_imp_lt: "\<lbrakk>(k#-i) < (k#-j); i\<in>nat; j\<in>nat; k\<in>nat\<rbrakk> \<Longrightarrow> j<i" |
14060
c0c4af41fa3b
Adding the theory UNITY/AllocImpl.thy, with supporting lemmas
paulson
parents:
14055
diff
changeset
|
612 |
apply (erule rev_mp) |
63648 | 613 |
apply (simp split: nat_diff_split, auto) |
14060
c0c4af41fa3b
Adding the theory UNITY/AllocImpl.thy, with supporting lemmas
paulson
parents:
14055
diff
changeset
|
614 |
apply (blast intro: add_le_self lt_trans1) |
c0c4af41fa3b
Adding the theory UNITY/AllocImpl.thy, with supporting lemmas
paulson
parents:
14055
diff
changeset
|
615 |
apply (rule not_le_iff_lt [THEN iffD1], auto) |
c0c4af41fa3b
Adding the theory UNITY/AllocImpl.thy, with supporting lemmas
paulson
parents:
14055
diff
changeset
|
616 |
apply (subgoal_tac "i #+ da < j #+ d", force) |
46820 | 617 |
apply (blast intro: add_le_lt_mono) |
14060
c0c4af41fa3b
Adding the theory UNITY/AllocImpl.thy, with supporting lemmas
paulson
parents:
14055
diff
changeset
|
618 |
done |
c0c4af41fa3b
Adding the theory UNITY/AllocImpl.thy, with supporting lemmas
paulson
parents:
14055
diff
changeset
|
619 |
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69605
diff
changeset
|
620 |
lemma lt_imp_diff_lt: "\<lbrakk>j<i; i\<le>k; k\<in>nat\<rbrakk> \<Longrightarrow> (k#-i) < (k#-j)" |
14060
c0c4af41fa3b
Adding the theory UNITY/AllocImpl.thy, with supporting lemmas
paulson
parents:
14055
diff
changeset
|
621 |
apply (frule le_in_nat, assumption) |
c0c4af41fa3b
Adding the theory UNITY/AllocImpl.thy, with supporting lemmas
paulson
parents:
14055
diff
changeset
|
622 |
apply (frule lt_nat_in_nat, assumption) |
63648 | 623 |
apply (simp split: nat_diff_split, auto) |
14060
c0c4af41fa3b
Adding the theory UNITY/AllocImpl.thy, with supporting lemmas
paulson
parents:
14055
diff
changeset
|
624 |
apply (blast intro: lt_asym lt_trans2) |
c0c4af41fa3b
Adding the theory UNITY/AllocImpl.thy, with supporting lemmas
paulson
parents:
14055
diff
changeset
|
625 |
apply (blast intro: lt_irrefl lt_trans2) |
c0c4af41fa3b
Adding the theory UNITY/AllocImpl.thy, with supporting lemmas
paulson
parents:
14055
diff
changeset
|
626 |
apply (rule not_le_iff_lt [THEN iffD1], auto) |
c0c4af41fa3b
Adding the theory UNITY/AllocImpl.thy, with supporting lemmas
paulson
parents:
14055
diff
changeset
|
627 |
apply (subgoal_tac "j #+ d < i #+ da", force) |
46820 | 628 |
apply (blast intro: add_lt_le_mono) |
14060
c0c4af41fa3b
Adding the theory UNITY/AllocImpl.thy, with supporting lemmas
paulson
parents:
14055
diff
changeset
|
629 |
done |
c0c4af41fa3b
Adding the theory UNITY/AllocImpl.thy, with supporting lemmas
paulson
parents:
14055
diff
changeset
|
630 |
|
c0c4af41fa3b
Adding the theory UNITY/AllocImpl.thy, with supporting lemmas
paulson
parents:
14055
diff
changeset
|
631 |
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
69605
diff
changeset
|
632 |
lemma diff_lt_iff_lt: "\<lbrakk>i\<le>k; j\<in>nat; k\<in>nat\<rbrakk> \<Longrightarrow> (k#-i) < (k#-j) \<longleftrightarrow> j<i" |
14060
c0c4af41fa3b
Adding the theory UNITY/AllocImpl.thy, with supporting lemmas
paulson
parents:
14055
diff
changeset
|
633 |
apply (frule le_in_nat, assumption) |
46820 | 634 |
apply (blast intro: lt_imp_diff_lt diff_lt_imp_lt) |
14060
c0c4af41fa3b
Adding the theory UNITY/AllocImpl.thy, with supporting lemmas
paulson
parents:
14055
diff
changeset
|
635 |
done |
c0c4af41fa3b
Adding the theory UNITY/AllocImpl.thy, with supporting lemmas
paulson
parents:
14055
diff
changeset
|
636 |
|
9548 | 637 |
end |