author | wenzelm |
Mon, 22 Oct 2001 17:58:11 +0200 | |
changeset 11879 | 1a386a1e002c |
parent 11386 | cf8d81cf8034 |
child 12789 | 459b5de466b2 |
permissions | -rw-r--r-- |
1793 | 1 |
(* Title: ZF/Arith.ML |
0 | 2 |
ID: $Id$ |
1461 | 3 |
Author: Lawrence C Paulson, Cambridge University Computer Laboratory |
0 | 4 |
Copyright 1992 University of Cambridge |
5 |
||
1609 | 6 |
Arithmetic operators and their definitions |
0 | 7 |
|
8 |
Proofs about elementary arithmetic: addition, multiplication, etc. |
|
9 |
*) |
|
10 |
||
11 |
(*"Difference" is subtraction of natural numbers. |
|
12 |
There are no negative numbers; we have |
|
13 |
m #- n = 0 iff m<=n and m #- n = succ(k) iff m>n. |
|
14 |
Also, rec(m, 0, %z w.z) is pred(m). |
|
15 |
*) |
|
16 |
||
8127
68c6159440f1
new lemmas for Ntree recursor example; more simprules; more lemmas borrowed
paulson
parents:
7499
diff
changeset
|
17 |
Addsimps [rec_type, nat_0_le]; |
9907 | 18 |
bind_thms ("nat_typechecks", [rec_type, nat_0I, nat_1I, nat_succI, Ord_nat]); |
0 | 19 |
|
5137 | 20 |
Goal "[| 0<k; k: nat |] ==> EX j: nat. k = succ(j)"; |
1708 | 21 |
by (etac rev_mp 1); |
6070 | 22 |
by (induct_tac "k" 1); |
2469 | 23 |
by (Simp_tac 1); |
3016 | 24 |
by (Blast_tac 1); |
1708 | 25 |
val lemma = result(); |
26 |
||
27 |
(* [| 0 < k; k: nat; !!j. [| j: nat; k = succ(j) |] ==> Q |] ==> Q *) |
|
28 |
bind_thm ("zero_lt_natE", lemma RS bexE); |
|
29 |
||
30 |
||
9491
1a36151ee2fc
natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents:
9301
diff
changeset
|
31 |
(** natify: coercion to "nat" **) |
1a36151ee2fc
natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents:
9301
diff
changeset
|
32 |
|
1a36151ee2fc
natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents:
9301
diff
changeset
|
33 |
Goalw [pred_def] "pred(succ(y)) = y"; |
1a36151ee2fc
natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents:
9301
diff
changeset
|
34 |
by Auto_tac; |
1a36151ee2fc
natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents:
9301
diff
changeset
|
35 |
qed "pred_succ_eq"; |
1a36151ee2fc
natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents:
9301
diff
changeset
|
36 |
Addsimps [pred_succ_eq]; |
1a36151ee2fc
natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents:
9301
diff
changeset
|
37 |
|
1a36151ee2fc
natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents:
9301
diff
changeset
|
38 |
Goal "natify(succ(x)) = succ(natify(x))"; |
1a36151ee2fc
natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents:
9301
diff
changeset
|
39 |
by (rtac (natify_def RS def_Vrecursor RS trans) 1); |
1a36151ee2fc
natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents:
9301
diff
changeset
|
40 |
by Auto_tac; |
1a36151ee2fc
natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents:
9301
diff
changeset
|
41 |
qed "natify_succ"; |
1a36151ee2fc
natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents:
9301
diff
changeset
|
42 |
|
1a36151ee2fc
natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents:
9301
diff
changeset
|
43 |
Goal "natify(0) = 0"; |
1a36151ee2fc
natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents:
9301
diff
changeset
|
44 |
by (rtac (natify_def RS def_Vrecursor RS trans) 1); |
1a36151ee2fc
natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents:
9301
diff
changeset
|
45 |
by Auto_tac; |
1a36151ee2fc
natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents:
9301
diff
changeset
|
46 |
qed "natify_0"; |
1a36151ee2fc
natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents:
9301
diff
changeset
|
47 |
Addsimps [natify_0]; |
1a36151ee2fc
natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents:
9301
diff
changeset
|
48 |
|
1a36151ee2fc
natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents:
9301
diff
changeset
|
49 |
Goal "ALL z. x ~= succ(z) ==> natify(x) = 0"; |
1a36151ee2fc
natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents:
9301
diff
changeset
|
50 |
by (rtac (natify_def RS def_Vrecursor RS trans) 1); |
1a36151ee2fc
natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents:
9301
diff
changeset
|
51 |
by Auto_tac; |
1a36151ee2fc
natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents:
9301
diff
changeset
|
52 |
qed "natify_non_succ"; |
1a36151ee2fc
natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents:
9301
diff
changeset
|
53 |
|
1a36151ee2fc
natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents:
9301
diff
changeset
|
54 |
Goal "natify(x) : nat"; |
1a36151ee2fc
natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents:
9301
diff
changeset
|
55 |
by (eps_ind_tac "x" 1); |
1a36151ee2fc
natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents:
9301
diff
changeset
|
56 |
by (case_tac "EX z. x1 = succ(z)" 1); |
9548 | 57 |
by (auto_tac (claset(), simpset() addsimps [natify_succ, natify_non_succ])); |
9491
1a36151ee2fc
natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents:
9301
diff
changeset
|
58 |
qed "natify_in_nat"; |
1a36151ee2fc
natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents:
9301
diff
changeset
|
59 |
AddIffs [natify_in_nat]; |
1a36151ee2fc
natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents:
9301
diff
changeset
|
60 |
AddTCs [natify_in_nat]; |
1a36151ee2fc
natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents:
9301
diff
changeset
|
61 |
|
1a36151ee2fc
natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents:
9301
diff
changeset
|
62 |
Goal "n : nat ==> natify(n) = n"; |
1a36151ee2fc
natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents:
9301
diff
changeset
|
63 |
by (induct_tac "n" 1); |
9548 | 64 |
by (auto_tac (claset(), simpset() addsimps [natify_succ])); |
9491
1a36151ee2fc
natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents:
9301
diff
changeset
|
65 |
qed "natify_ident"; |
1a36151ee2fc
natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents:
9301
diff
changeset
|
66 |
Addsimps [natify_ident]; |
1a36151ee2fc
natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents:
9301
diff
changeset
|
67 |
|
11386 | 68 |
Goal "[|natify(x) = y; x: nat|] ==> x=y"; |
69 |
by Auto_tac; |
|
70 |
qed "natify_eqE"; |
|
71 |
||
9491
1a36151ee2fc
natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents:
9301
diff
changeset
|
72 |
|
1a36151ee2fc
natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents:
9301
diff
changeset
|
73 |
(*** Collapsing rules: to remove natify from arithmetic expressions ***) |
1a36151ee2fc
natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents:
9301
diff
changeset
|
74 |
|
1a36151ee2fc
natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents:
9301
diff
changeset
|
75 |
Goal "natify(natify(x)) = natify(x)"; |
1a36151ee2fc
natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents:
9301
diff
changeset
|
76 |
by (Simp_tac 1); |
1a36151ee2fc
natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents:
9301
diff
changeset
|
77 |
qed "natify_idem"; |
1a36151ee2fc
natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents:
9301
diff
changeset
|
78 |
Addsimps [natify_idem]; |
1a36151ee2fc
natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents:
9301
diff
changeset
|
79 |
|
0 | 80 |
(** Addition **) |
81 |
||
9491
1a36151ee2fc
natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents:
9301
diff
changeset
|
82 |
Goal "natify(m) #+ n = m #+ n"; |
1a36151ee2fc
natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents:
9301
diff
changeset
|
83 |
by (simp_tac (simpset() addsimps [add_def]) 1); |
1a36151ee2fc
natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents:
9301
diff
changeset
|
84 |
qed "add_natify1"; |
1a36151ee2fc
natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents:
9301
diff
changeset
|
85 |
|
1a36151ee2fc
natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents:
9301
diff
changeset
|
86 |
Goal "m #+ natify(n) = m #+ n"; |
1a36151ee2fc
natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents:
9301
diff
changeset
|
87 |
by (simp_tac (simpset() addsimps [add_def]) 1); |
1a36151ee2fc
natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents:
9301
diff
changeset
|
88 |
qed "add_natify2"; |
1a36151ee2fc
natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents:
9301
diff
changeset
|
89 |
Addsimps [add_natify1, add_natify2]; |
1a36151ee2fc
natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents:
9301
diff
changeset
|
90 |
|
1a36151ee2fc
natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents:
9301
diff
changeset
|
91 |
(** Multiplication **) |
1a36151ee2fc
natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents:
9301
diff
changeset
|
92 |
|
1a36151ee2fc
natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents:
9301
diff
changeset
|
93 |
Goal "natify(m) #* n = m #* n"; |
1a36151ee2fc
natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents:
9301
diff
changeset
|
94 |
by (simp_tac (simpset() addsimps [mult_def]) 1); |
1a36151ee2fc
natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents:
9301
diff
changeset
|
95 |
qed "mult_natify1"; |
1a36151ee2fc
natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents:
9301
diff
changeset
|
96 |
|
1a36151ee2fc
natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents:
9301
diff
changeset
|
97 |
Goal "m #* natify(n) = m #* n"; |
1a36151ee2fc
natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents:
9301
diff
changeset
|
98 |
by (simp_tac (simpset() addsimps [mult_def]) 1); |
1a36151ee2fc
natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents:
9301
diff
changeset
|
99 |
qed "mult_natify2"; |
1a36151ee2fc
natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents:
9301
diff
changeset
|
100 |
Addsimps [mult_natify1, mult_natify2]; |
1a36151ee2fc
natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents:
9301
diff
changeset
|
101 |
|
1a36151ee2fc
natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents:
9301
diff
changeset
|
102 |
(** Difference **) |
1a36151ee2fc
natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents:
9301
diff
changeset
|
103 |
|
1a36151ee2fc
natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents:
9301
diff
changeset
|
104 |
Goal "natify(m) #- n = m #- n"; |
1a36151ee2fc
natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents:
9301
diff
changeset
|
105 |
by (simp_tac (simpset() addsimps [diff_def]) 1); |
1a36151ee2fc
natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents:
9301
diff
changeset
|
106 |
qed "diff_natify1"; |
1a36151ee2fc
natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents:
9301
diff
changeset
|
107 |
|
1a36151ee2fc
natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents:
9301
diff
changeset
|
108 |
Goal "m #- natify(n) = m #- n"; |
1a36151ee2fc
natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents:
9301
diff
changeset
|
109 |
by (simp_tac (simpset() addsimps [diff_def]) 1); |
1a36151ee2fc
natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents:
9301
diff
changeset
|
110 |
qed "diff_natify2"; |
1a36151ee2fc
natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents:
9301
diff
changeset
|
111 |
Addsimps [diff_natify1, diff_natify2]; |
1a36151ee2fc
natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents:
9301
diff
changeset
|
112 |
|
9492
72e429c66608
used natify with div and mod; also put in the divide-by-zero trick
paulson
parents:
9491
diff
changeset
|
113 |
(** Remainder **) |
72e429c66608
used natify with div and mod; also put in the divide-by-zero trick
paulson
parents:
9491
diff
changeset
|
114 |
|
72e429c66608
used natify with div and mod; also put in the divide-by-zero trick
paulson
parents:
9491
diff
changeset
|
115 |
Goal "natify(m) mod n = m mod n"; |
72e429c66608
used natify with div and mod; also put in the divide-by-zero trick
paulson
parents:
9491
diff
changeset
|
116 |
by (simp_tac (simpset() addsimps [mod_def]) 1); |
72e429c66608
used natify with div and mod; also put in the divide-by-zero trick
paulson
parents:
9491
diff
changeset
|
117 |
qed "mod_natify1"; |
72e429c66608
used natify with div and mod; also put in the divide-by-zero trick
paulson
parents:
9491
diff
changeset
|
118 |
|
72e429c66608
used natify with div and mod; also put in the divide-by-zero trick
paulson
parents:
9491
diff
changeset
|
119 |
Goal "m mod natify(n) = m mod n"; |
72e429c66608
used natify with div and mod; also put in the divide-by-zero trick
paulson
parents:
9491
diff
changeset
|
120 |
by (simp_tac (simpset() addsimps [mod_def]) 1); |
72e429c66608
used natify with div and mod; also put in the divide-by-zero trick
paulson
parents:
9491
diff
changeset
|
121 |
qed "mod_natify2"; |
72e429c66608
used natify with div and mod; also put in the divide-by-zero trick
paulson
parents:
9491
diff
changeset
|
122 |
Addsimps [mod_natify1, mod_natify2]; |
72e429c66608
used natify with div and mod; also put in the divide-by-zero trick
paulson
parents:
9491
diff
changeset
|
123 |
|
72e429c66608
used natify with div and mod; also put in the divide-by-zero trick
paulson
parents:
9491
diff
changeset
|
124 |
(** Quotient **) |
72e429c66608
used natify with div and mod; also put in the divide-by-zero trick
paulson
parents:
9491
diff
changeset
|
125 |
|
72e429c66608
used natify with div and mod; also put in the divide-by-zero trick
paulson
parents:
9491
diff
changeset
|
126 |
Goal "natify(m) div n = m div n"; |
72e429c66608
used natify with div and mod; also put in the divide-by-zero trick
paulson
parents:
9491
diff
changeset
|
127 |
by (simp_tac (simpset() addsimps [div_def]) 1); |
72e429c66608
used natify with div and mod; also put in the divide-by-zero trick
paulson
parents:
9491
diff
changeset
|
128 |
qed "div_natify1"; |
72e429c66608
used natify with div and mod; also put in the divide-by-zero trick
paulson
parents:
9491
diff
changeset
|
129 |
|
72e429c66608
used natify with div and mod; also put in the divide-by-zero trick
paulson
parents:
9491
diff
changeset
|
130 |
Goal "m div natify(n) = m div n"; |
72e429c66608
used natify with div and mod; also put in the divide-by-zero trick
paulson
parents:
9491
diff
changeset
|
131 |
by (simp_tac (simpset() addsimps [div_def]) 1); |
72e429c66608
used natify with div and mod; also put in the divide-by-zero trick
paulson
parents:
9491
diff
changeset
|
132 |
qed "div_natify2"; |
72e429c66608
used natify with div and mod; also put in the divide-by-zero trick
paulson
parents:
9491
diff
changeset
|
133 |
Addsimps [div_natify1, div_natify2]; |
72e429c66608
used natify with div and mod; also put in the divide-by-zero trick
paulson
parents:
9491
diff
changeset
|
134 |
|
9491
1a36151ee2fc
natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents:
9301
diff
changeset
|
135 |
|
1a36151ee2fc
natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents:
9301
diff
changeset
|
136 |
(*** Typing rules ***) |
1a36151ee2fc
natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents:
9301
diff
changeset
|
137 |
|
1a36151ee2fc
natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents:
9301
diff
changeset
|
138 |
(** Addition **) |
1a36151ee2fc
natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents:
9301
diff
changeset
|
139 |
|
9492
72e429c66608
used natify with div and mod; also put in the divide-by-zero trick
paulson
parents:
9491
diff
changeset
|
140 |
Goal "[| m:nat; n:nat |] ==> raw_add (m, n) : nat"; |
6070 | 141 |
by (induct_tac "m" 1); |
142 |
by Auto_tac; |
|
9491
1a36151ee2fc
natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents:
9301
diff
changeset
|
143 |
qed "raw_add_type"; |
1a36151ee2fc
natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents:
9301
diff
changeset
|
144 |
|
1a36151ee2fc
natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents:
9301
diff
changeset
|
145 |
Goal "m #+ n : nat"; |
1a36151ee2fc
natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents:
9301
diff
changeset
|
146 |
by (simp_tac (simpset() addsimps [add_def, raw_add_type]) 1); |
6070 | 147 |
qed "add_type"; |
9491
1a36151ee2fc
natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents:
9301
diff
changeset
|
148 |
AddIffs [add_type]; |
6153 | 149 |
AddTCs [add_type]; |
2469 | 150 |
|
0 | 151 |
(** Multiplication **) |
152 |
||
9492
72e429c66608
used natify with div and mod; also put in the divide-by-zero trick
paulson
parents:
9491
diff
changeset
|
153 |
Goal "[| m:nat; n:nat |] ==> raw_mult (m, n) : nat"; |
6070 | 154 |
by (induct_tac "m" 1); |
9491
1a36151ee2fc
natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents:
9301
diff
changeset
|
155 |
by (ALLGOALS (asm_simp_tac (simpset() addsimps [raw_add_type]))); |
1a36151ee2fc
natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents:
9301
diff
changeset
|
156 |
qed "raw_mult_type"; |
1a36151ee2fc
natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents:
9301
diff
changeset
|
157 |
|
1a36151ee2fc
natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents:
9301
diff
changeset
|
158 |
Goal "m #* n : nat"; |
1a36151ee2fc
natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents:
9301
diff
changeset
|
159 |
by (simp_tac (simpset() addsimps [mult_def, raw_mult_type]) 1); |
6070 | 160 |
qed "mult_type"; |
9491
1a36151ee2fc
natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents:
9301
diff
changeset
|
161 |
AddIffs [mult_type]; |
6153 | 162 |
AddTCs [mult_type]; |
0 | 163 |
|
2469 | 164 |
|
0 | 165 |
(** Difference **) |
166 |
||
9492
72e429c66608
used natify with div and mod; also put in the divide-by-zero trick
paulson
parents:
9491
diff
changeset
|
167 |
Goal "[| m:nat; n:nat |] ==> raw_diff (m, n) : nat"; |
6070 | 168 |
by (induct_tac "n" 1); |
169 |
by Auto_tac; |
|
170 |
by (fast_tac (claset() addIs [nat_case_type]) 1); |
|
9491
1a36151ee2fc
natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents:
9301
diff
changeset
|
171 |
qed "raw_diff_type"; |
1a36151ee2fc
natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents:
9301
diff
changeset
|
172 |
|
1a36151ee2fc
natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents:
9301
diff
changeset
|
173 |
Goal "m #- n : nat"; |
1a36151ee2fc
natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents:
9301
diff
changeset
|
174 |
by (simp_tac (simpset() addsimps [diff_def, raw_diff_type]) 1); |
6070 | 175 |
qed "diff_type"; |
9491
1a36151ee2fc
natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents:
9301
diff
changeset
|
176 |
AddIffs [diff_type]; |
6153 | 177 |
AddTCs [diff_type]; |
0 | 178 |
|
9491
1a36151ee2fc
natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents:
9301
diff
changeset
|
179 |
Goalw [diff_def] "0 #- n = 0"; |
1a36151ee2fc
natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents:
9301
diff
changeset
|
180 |
by (rtac (natify_in_nat RS nat_induct) 1); |
6070 | 181 |
by Auto_tac; |
182 |
qed "diff_0_eq_0"; |
|
0 | 183 |
|
6070 | 184 |
(*Must simplify BEFORE the induction: else we get a critical pair*) |
9491
1a36151ee2fc
natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents:
9301
diff
changeset
|
185 |
Goal "succ(m) #- succ(n) = m #- n"; |
9548 | 186 |
by (simp_tac (simpset() addsimps [natify_succ, diff_def]) 1); |
9491
1a36151ee2fc
natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents:
9301
diff
changeset
|
187 |
by (res_inst_tac [("x1","n")] (natify_in_nat RS nat_induct) 1); |
6070 | 188 |
by Auto_tac; |
189 |
qed "diff_succ_succ"; |
|
0 | 190 |
|
9491
1a36151ee2fc
natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents:
9301
diff
changeset
|
191 |
(*This defining property is no longer wanted*) |
1a36151ee2fc
natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents:
9301
diff
changeset
|
192 |
Delsimps [raw_diff_succ]; |
2469 | 193 |
|
9491
1a36151ee2fc
natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents:
9301
diff
changeset
|
194 |
(*Natify has weakened this law, compared with the older approach*) |
1a36151ee2fc
natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents:
9301
diff
changeset
|
195 |
Goal "m #- 0 = natify(m)"; |
1a36151ee2fc
natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents:
9301
diff
changeset
|
196 |
by (asm_simp_tac (simpset() addsimps [diff_def]) 1); |
1a36151ee2fc
natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents:
9301
diff
changeset
|
197 |
qed "diff_0"; |
0 | 198 |
|
9491
1a36151ee2fc
natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents:
9301
diff
changeset
|
199 |
Addsimps [diff_0, diff_0_eq_0, diff_succ_succ]; |
1a36151ee2fc
natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents:
9301
diff
changeset
|
200 |
|
1a36151ee2fc
natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents:
9301
diff
changeset
|
201 |
Goal "m:nat ==> (m #- n) le m"; |
1a36151ee2fc
natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents:
9301
diff
changeset
|
202 |
by (subgoal_tac "(m #- natify(n)) le m" 1); |
1a36151ee2fc
natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents:
9301
diff
changeset
|
203 |
by (res_inst_tac [("m","m"), ("n","natify(n)")] diff_induct 2); |
1a36151ee2fc
natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents:
9301
diff
changeset
|
204 |
by (etac leE 6); |
1a36151ee2fc
natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents:
9301
diff
changeset
|
205 |
by (ALLGOALS (asm_full_simp_tac (simpset() addsimps [le_iff]))); |
760 | 206 |
qed "diff_le_self"; |
0 | 207 |
|
208 |
||
209 |
(*** Addition ***) |
|
210 |
||
9491
1a36151ee2fc
natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents:
9301
diff
changeset
|
211 |
(*Natify has weakened this law, compared with the older approach*) |
1a36151ee2fc
natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents:
9301
diff
changeset
|
212 |
Goal "0 #+ m = natify(m)"; |
1a36151ee2fc
natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents:
9301
diff
changeset
|
213 |
by (asm_simp_tac (simpset() addsimps [add_def]) 1); |
9548 | 214 |
qed "add_0_natify"; |
9491
1a36151ee2fc
natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents:
9301
diff
changeset
|
215 |
|
1a36151ee2fc
natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents:
9301
diff
changeset
|
216 |
Goal "succ(m) #+ n = succ(m #+ n)"; |
9548 | 217 |
by (asm_simp_tac (simpset() addsimps [natify_succ, add_def]) 1); |
9491
1a36151ee2fc
natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents:
9301
diff
changeset
|
218 |
qed "add_succ"; |
1a36151ee2fc
natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents:
9301
diff
changeset
|
219 |
|
9548 | 220 |
Addsimps [add_0_natify, add_succ]; |
221 |
||
222 |
Goal "m: nat ==> 0 #+ m = m"; |
|
223 |
by (Asm_simp_tac 1); |
|
224 |
qed "add_0"; |
|
9491
1a36151ee2fc
natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents:
9301
diff
changeset
|
225 |
|
0 | 226 |
(*Associative law for addition*) |
9491
1a36151ee2fc
natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents:
9301
diff
changeset
|
227 |
Goal "(m #+ n) #+ k = m #+ (n #+ k)"; |
1a36151ee2fc
natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents:
9301
diff
changeset
|
228 |
by (subgoal_tac "(natify(m) #+ natify(n)) #+ natify(k) = \ |
1a36151ee2fc
natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents:
9301
diff
changeset
|
229 |
\ natify(m) #+ (natify(n) #+ natify(k))" 1); |
1a36151ee2fc
natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents:
9301
diff
changeset
|
230 |
by (res_inst_tac [("n","natify(m)")] nat_induct 2); |
6070 | 231 |
by Auto_tac; |
232 |
qed "add_assoc"; |
|
0 | 233 |
|
234 |
(*The following two lemmas are used for add_commute and sometimes |
|
235 |
elsewhere, since they are safe for rewriting.*) |
|
9491
1a36151ee2fc
natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents:
9301
diff
changeset
|
236 |
Goal "m #+ 0 = natify(m)"; |
1a36151ee2fc
natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents:
9301
diff
changeset
|
237 |
by (subgoal_tac "natify(m) #+ 0 = natify(m)" 1); |
1a36151ee2fc
natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents:
9301
diff
changeset
|
238 |
by (res_inst_tac [("n","natify(m)")] nat_induct 2); |
1a36151ee2fc
natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents:
9301
diff
changeset
|
239 |
by Auto_tac; |
1a36151ee2fc
natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents:
9301
diff
changeset
|
240 |
qed "add_0_right_natify"; |
1a36151ee2fc
natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents:
9301
diff
changeset
|
241 |
|
1a36151ee2fc
natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents:
9301
diff
changeset
|
242 |
Goalw [add_def] "m #+ succ(n) = succ(m #+ n)"; |
1a36151ee2fc
natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents:
9301
diff
changeset
|
243 |
by (res_inst_tac [("n","natify(m)")] nat_induct 1); |
9548 | 244 |
by (auto_tac (claset(), simpset() addsimps [natify_succ])); |
9491
1a36151ee2fc
natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents:
9301
diff
changeset
|
245 |
qed "add_succ_right"; |
1a36151ee2fc
natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents:
9301
diff
changeset
|
246 |
|
1a36151ee2fc
natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents:
9301
diff
changeset
|
247 |
Addsimps [add_0_right_natify, add_succ_right]; |
1a36151ee2fc
natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents:
9301
diff
changeset
|
248 |
|
1a36151ee2fc
natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents:
9301
diff
changeset
|
249 |
Goal "m: nat ==> m #+ 0 = m"; |
6070 | 250 |
by Auto_tac; |
251 |
qed "add_0_right"; |
|
0 | 252 |
|
253 |
(*Commutative law for addition*) |
|
9491
1a36151ee2fc
natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents:
9301
diff
changeset
|
254 |
Goal "m #+ n = n #+ m"; |
1a36151ee2fc
natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents:
9301
diff
changeset
|
255 |
by (subgoal_tac "natify(m) #+ natify(n) = natify(n) #+ natify(m)" 1); |
1a36151ee2fc
natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents:
9301
diff
changeset
|
256 |
by (res_inst_tac [("n","natify(m)")] nat_induct 2); |
6070 | 257 |
by Auto_tac; |
258 |
qed "add_commute"; |
|
435 | 259 |
|
437 | 260 |
(*for a/c rewriting*) |
9491
1a36151ee2fc
natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents:
9301
diff
changeset
|
261 |
Goal "m#+(n#+k)=n#+(m#+k)"; |
1a36151ee2fc
natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents:
9301
diff
changeset
|
262 |
by (rtac (add_commute RS trans) 1); |
1a36151ee2fc
natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents:
9301
diff
changeset
|
263 |
by (rtac (add_assoc RS trans) 1); |
1a36151ee2fc
natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents:
9301
diff
changeset
|
264 |
by (rtac (add_commute RS subst_context) 1); |
6070 | 265 |
qed "add_left_commute"; |
435 | 266 |
|
267 |
(*Addition is an AC-operator*) |
|
9907 | 268 |
bind_thms ("add_ac", [add_assoc, add_commute, add_left_commute]); |
0 | 269 |
|
270 |
(*Cancellation law on the left*) |
|
9492
72e429c66608
used natify with div and mod; also put in the divide-by-zero trick
paulson
parents:
9491
diff
changeset
|
271 |
Goal "[| raw_add(k, m) = raw_add(k, n); k:nat |] ==> m=n"; |
6070 | 272 |
by (etac rev_mp 1); |
273 |
by (induct_tac "k" 1); |
|
274 |
by Auto_tac; |
|
9491
1a36151ee2fc
natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents:
9301
diff
changeset
|
275 |
qed "raw_add_left_cancel"; |
1a36151ee2fc
natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents:
9301
diff
changeset
|
276 |
|
1a36151ee2fc
natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents:
9301
diff
changeset
|
277 |
Goalw [add_def] "k #+ m = k #+ n ==> natify(m) = natify(n)"; |
1a36151ee2fc
natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents:
9301
diff
changeset
|
278 |
by (dtac raw_add_left_cancel 1); |
1a36151ee2fc
natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents:
9301
diff
changeset
|
279 |
by Auto_tac; |
1a36151ee2fc
natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents:
9301
diff
changeset
|
280 |
qed "add_left_cancel_natify"; |
1a36151ee2fc
natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents:
9301
diff
changeset
|
281 |
|
9548 | 282 |
Goal "[| i = j; i #+ m = j #+ n; m:nat; n:nat |] ==> m = n"; |
283 |
by (force_tac (claset() addSDs [add_left_cancel_natify], simpset()) 1); |
|
760 | 284 |
qed "add_left_cancel"; |
0 | 285 |
|
9548 | 286 |
(*Thanks to Sten Agerholm*) |
287 |
Goal "k#+m le k#+n ==> natify(m) le natify(n)"; |
|
288 |
by (res_inst_tac [("P", "natify(k)#+m le natify(k)#+n")] rev_mp 1); |
|
289 |
by (res_inst_tac [("n","natify(k)")] nat_induct 2); |
|
290 |
by Auto_tac; |
|
291 |
qed "add_le_elim1_natify"; |
|
9491
1a36151ee2fc
natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents:
9301
diff
changeset
|
292 |
|
9548 | 293 |
Goal "[| k#+m le k#+n; m: nat; n: nat |] ==> m le n"; |
294 |
by (dtac add_le_elim1_natify 1); |
|
295 |
by Auto_tac; |
|
296 |
qed "add_le_elim1"; |
|
297 |
||
298 |
Goal "k#+m < k#+n ==> natify(m) < natify(n)"; |
|
299 |
by (res_inst_tac [("P", "natify(k)#+m < natify(k)#+n")] rev_mp 1); |
|
300 |
by (res_inst_tac [("n","natify(k)")] nat_induct 2); |
|
301 |
by Auto_tac; |
|
302 |
qed "add_lt_elim1_natify"; |
|
303 |
||
304 |
Goal "[| k#+m < k#+n; m: nat; n: nat |] ==> m < n"; |
|
305 |
by (dtac add_lt_elim1_natify 1); |
|
306 |
by Auto_tac; |
|
307 |
qed "add_lt_elim1"; |
|
308 |
||
309 |
||
310 |
(*** Monotonicity of Addition ***) |
|
311 |
||
312 |
(*strict, in 1st argument; proof is by rule induction on 'less than'. |
|
313 |
Still need j:nat, for consider j = omega. Then we can have i<omega, |
|
314 |
which is the same as i:nat, but natify(j)=0, so the conclusion fails.*) |
|
315 |
Goal "[| i<j; j:nat |] ==> i#+k < j#+k"; |
|
316 |
by (ftac lt_nat_in_nat 1); |
|
317 |
by (assume_tac 1); |
|
318 |
by (etac succ_lt_induct 1); |
|
319 |
by (ALLGOALS (asm_simp_tac (simpset() addsimps [leI]))); |
|
320 |
qed "add_lt_mono1"; |
|
321 |
||
322 |
(*strict, in both arguments*) |
|
323 |
Goal "[| i<j; k<l; j:nat; l:nat |] ==> i#+k < j#+l"; |
|
324 |
by (rtac (add_lt_mono1 RS lt_trans) 1); |
|
325 |
by (REPEAT (assume_tac 1)); |
|
326 |
by (EVERY [stac add_commute 1, |
|
327 |
stac add_commute 1, |
|
328 |
rtac add_lt_mono1 1]); |
|
329 |
by (REPEAT (assume_tac 1)); |
|
330 |
qed "add_lt_mono"; |
|
331 |
||
332 |
(*A [clumsy] way of lifting < monotonicity to le monotonicity *) |
|
333 |
val lt_mono::ford::prems = Goal |
|
334 |
"[| !!i j. [| i<j; j:k |] ==> f(i) < f(j); \ |
|
335 |
\ !!i. i:k ==> Ord(f(i)); \ |
|
336 |
\ i le j; j:k \ |
|
337 |
\ |] ==> f(i) le f(j)"; |
|
338 |
by (cut_facts_tac prems 1); |
|
339 |
by (blast_tac (le_cs addSIs [lt_mono,ford] addSEs [leE]) 1); |
|
340 |
qed "Ord_lt_mono_imp_le_mono"; |
|
341 |
||
342 |
(*le monotonicity, 1st argument*) |
|
343 |
Goal "[| i le j; j:nat |] ==> i#+k le j#+k"; |
|
344 |
by (res_inst_tac [("f", "%j. j#+k")] Ord_lt_mono_imp_le_mono 1); |
|
345 |
by (REPEAT (ares_tac [add_lt_mono1, add_type RS nat_into_Ord] 1)); |
|
346 |
qed "add_le_mono1"; |
|
347 |
||
348 |
(* le monotonicity, BOTH arguments*) |
|
349 |
Goal "[| i le j; k le l; j:nat; l:nat |] ==> i#+k le j#+l"; |
|
350 |
by (rtac (add_le_mono1 RS le_trans) 1); |
|
351 |
by (REPEAT (assume_tac 1)); |
|
352 |
by (EVERY [stac add_commute 1, |
|
353 |
stac add_commute 1, |
|
354 |
rtac add_le_mono1 1]); |
|
355 |
by (REPEAT (assume_tac 1)); |
|
356 |
qed "add_le_mono"; |
|
357 |
||
358 |
(** Subtraction is the inverse of addition. **) |
|
359 |
||
360 |
Goal "(n#+m) #- n = natify(m)"; |
|
361 |
by (subgoal_tac "(natify(n) #+ m) #- natify(n) = natify(m)" 1); |
|
362 |
by (res_inst_tac [("n","natify(n)")] nat_induct 2); |
|
363 |
by Auto_tac; |
|
364 |
qed "diff_add_inverse"; |
|
365 |
||
366 |
Goal "(m#+n) #- n = natify(m)"; |
|
367 |
by (simp_tac (simpset() addsimps [inst "m" "m" add_commute, |
|
368 |
diff_add_inverse]) 1); |
|
369 |
qed "diff_add_inverse2"; |
|
370 |
||
371 |
Goal "(k#+m) #- (k#+n) = m #- n"; |
|
372 |
by (subgoal_tac "(natify(k) #+ natify(m)) #- (natify(k) #+ natify(n)) = \ |
|
373 |
\ natify(m) #- natify(n)" 1); |
|
374 |
by (res_inst_tac [("n","natify(k)")] nat_induct 2); |
|
375 |
by Auto_tac; |
|
376 |
qed "diff_cancel"; |
|
377 |
||
378 |
Goal "(m#+k) #- (n#+k) = m #- n"; |
|
379 |
by (simp_tac (simpset() addsimps [inst "n" "k" add_commute, diff_cancel]) 1); |
|
380 |
qed "diff_cancel2"; |
|
381 |
||
382 |
Goal "n #- (n#+m) = 0"; |
|
383 |
by (subgoal_tac "natify(n) #- (natify(n) #+ natify(m)) = 0" 1); |
|
384 |
by (res_inst_tac [("n","natify(n)")] nat_induct 2); |
|
385 |
by Auto_tac; |
|
386 |
qed "diff_add_0"; |
|
387 |
||
388 |
||
389 |
(** Lemmas for the CancelNumerals simproc **) |
|
390 |
||
391 |
Goal "(u #+ m = u #+ n) <-> (0 #+ m = natify(n))"; |
|
392 |
by Auto_tac; |
|
393 |
by (blast_tac (claset() addDs [add_left_cancel_natify]) 1); |
|
394 |
by (asm_full_simp_tac (simpset() addsimps [add_def]) 1); |
|
395 |
qed "eq_add_iff"; |
|
396 |
||
397 |
Goal "(u #+ m < u #+ n) <-> (0 #+ m < natify(n))"; |
|
398 |
by (auto_tac (claset(), simpset() addsimps [add_lt_elim1_natify])); |
|
399 |
by (dtac add_lt_mono1 1); |
|
400 |
by (auto_tac (claset(), simpset() addsimps [inst "m" "u" add_commute])); |
|
401 |
qed "less_add_iff"; |
|
402 |
||
403 |
Goal "((u #+ m) #- (u #+ n)) = ((0 #+ m) #- n)"; |
|
404 |
by (asm_simp_tac (simpset() addsimps [diff_cancel]) 1); |
|
405 |
qed "diff_add_eq"; |
|
406 |
||
407 |
(*To tidy up the result of a simproc. Only the RHS will be simplified.*) |
|
408 |
Goal "u = u' ==> (t==u) == (t==u')"; |
|
409 |
by Auto_tac; |
|
410 |
qed "eq_cong2"; |
|
411 |
||
412 |
Goal "u <-> u' ==> (t==u) == (t==u')"; |
|
413 |
by Auto_tac; |
|
414 |
qed "iff_cong2"; |
|
415 |
||
416 |
||
417 |
(*** Multiplication [the simprocs need these laws] ***) |
|
0 | 418 |
|
9491
1a36151ee2fc
natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents:
9301
diff
changeset
|
419 |
Goal "0 #* m = 0"; |
1a36151ee2fc
natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents:
9301
diff
changeset
|
420 |
by (simp_tac (simpset() addsimps [mult_def]) 1); |
1a36151ee2fc
natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents:
9301
diff
changeset
|
421 |
qed "mult_0"; |
1a36151ee2fc
natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents:
9301
diff
changeset
|
422 |
|
1a36151ee2fc
natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents:
9301
diff
changeset
|
423 |
Goal "succ(m) #* n = n #+ (m #* n)"; |
9548 | 424 |
by (simp_tac (simpset() addsimps [add_def, mult_def, natify_succ, |
425 |
raw_mult_type]) 1); |
|
9491
1a36151ee2fc
natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents:
9301
diff
changeset
|
426 |
qed "mult_succ"; |
1a36151ee2fc
natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents:
9301
diff
changeset
|
427 |
|
1a36151ee2fc
natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents:
9301
diff
changeset
|
428 |
Addsimps [mult_0, mult_succ]; |
1a36151ee2fc
natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents:
9301
diff
changeset
|
429 |
|
0 | 430 |
(*right annihilation in product*) |
9491
1a36151ee2fc
natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents:
9301
diff
changeset
|
431 |
Goalw [mult_def] "m #* 0 = 0"; |
1a36151ee2fc
natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents:
9301
diff
changeset
|
432 |
by (res_inst_tac [("n","natify(m)")] nat_induct 1); |
6070 | 433 |
by Auto_tac; |
434 |
qed "mult_0_right"; |
|
0 | 435 |
|
436 |
(*right successor law for multiplication*) |
|
9491
1a36151ee2fc
natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents:
9301
diff
changeset
|
437 |
Goal "m #* succ(n) = m #+ (m #* n)"; |
1a36151ee2fc
natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents:
9301
diff
changeset
|
438 |
by (subgoal_tac "natify(m) #* succ(natify(n)) = \ |
1a36151ee2fc
natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents:
9301
diff
changeset
|
439 |
\ natify(m) #+ (natify(m) #* natify(n))" 1); |
9548 | 440 |
by (full_simp_tac (simpset() addsimps [natify_succ, add_def, mult_def]) 1); |
9491
1a36151ee2fc
natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents:
9301
diff
changeset
|
441 |
by (res_inst_tac [("n","natify(m)")] nat_induct 1); |
6070 | 442 |
by (ALLGOALS (asm_simp_tac (simpset() addsimps add_ac))); |
443 |
qed "mult_succ_right"; |
|
2469 | 444 |
|
445 |
Addsimps [mult_0_right, mult_succ_right]; |
|
0 | 446 |
|
9491
1a36151ee2fc
natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents:
9301
diff
changeset
|
447 |
Goal "1 #* n = natify(n)"; |
1a36151ee2fc
natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents:
9301
diff
changeset
|
448 |
by Auto_tac; |
1a36151ee2fc
natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents:
9301
diff
changeset
|
449 |
qed "mult_1_natify"; |
1a36151ee2fc
natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents:
9301
diff
changeset
|
450 |
|
1a36151ee2fc
natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents:
9301
diff
changeset
|
451 |
Goal "n #* 1 = natify(n)"; |
1a36151ee2fc
natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents:
9301
diff
changeset
|
452 |
by Auto_tac; |
1a36151ee2fc
natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents:
9301
diff
changeset
|
453 |
qed "mult_1_right_natify"; |
1a36151ee2fc
natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents:
9301
diff
changeset
|
454 |
|
1a36151ee2fc
natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents:
9301
diff
changeset
|
455 |
Addsimps [mult_1_natify, mult_1_right_natify]; |
1a36151ee2fc
natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents:
9301
diff
changeset
|
456 |
|
1a36151ee2fc
natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents:
9301
diff
changeset
|
457 |
Goal "n : nat ==> 1 #* n = n"; |
2469 | 458 |
by (Asm_simp_tac 1); |
1793 | 459 |
qed "mult_1"; |
460 |
||
9491
1a36151ee2fc
natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents:
9301
diff
changeset
|
461 |
Goal "n : nat ==> n #* 1 = n"; |
2469 | 462 |
by (Asm_simp_tac 1); |
1793 | 463 |
qed "mult_1_right"; |
464 |
||
0 | 465 |
(*Commutative law for multiplication*) |
9491
1a36151ee2fc
natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents:
9301
diff
changeset
|
466 |
Goal "m #* n = n #* m"; |
1a36151ee2fc
natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents:
9301
diff
changeset
|
467 |
by (subgoal_tac "natify(m) #* natify(n) = natify(n) #* natify(m)" 1); |
1a36151ee2fc
natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents:
9301
diff
changeset
|
468 |
by (res_inst_tac [("n","natify(m)")] nat_induct 2); |
6070 | 469 |
by Auto_tac; |
470 |
qed "mult_commute"; |
|
0 | 471 |
|
472 |
(*addition distributes over multiplication*) |
|
9491
1a36151ee2fc
natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents:
9301
diff
changeset
|
473 |
Goal "(m #+ n) #* k = (m #* k) #+ (n #* k)"; |
1a36151ee2fc
natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents:
9301
diff
changeset
|
474 |
by (subgoal_tac "(natify(m) #+ natify(n)) #* natify(k) = \ |
1a36151ee2fc
natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents:
9301
diff
changeset
|
475 |
\ (natify(m) #* natify(k)) #+ (natify(n) #* natify(k))" 1); |
1a36151ee2fc
natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents:
9301
diff
changeset
|
476 |
by (res_inst_tac [("n","natify(m)")] nat_induct 2); |
1a36151ee2fc
natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents:
9301
diff
changeset
|
477 |
by (ALLGOALS (asm_full_simp_tac (simpset() addsimps [add_assoc RS sym]))); |
6070 | 478 |
qed "add_mult_distrib"; |
0 | 479 |
|
9491
1a36151ee2fc
natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents:
9301
diff
changeset
|
480 |
(*Distributive law on the left*) |
1a36151ee2fc
natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents:
9301
diff
changeset
|
481 |
Goal "k #* (m #+ n) = (k #* m) #+ (k #* n)"; |
1a36151ee2fc
natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents:
9301
diff
changeset
|
482 |
by (subgoal_tac "natify(k) #* (natify(m) #+ natify(n)) = \ |
1a36151ee2fc
natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents:
9301
diff
changeset
|
483 |
\ (natify(k) #* natify(m)) #+ (natify(k) #* natify(n))" 1); |
1a36151ee2fc
natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents:
9301
diff
changeset
|
484 |
by (res_inst_tac [("n","natify(m)")] nat_induct 2); |
1a36151ee2fc
natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents:
9301
diff
changeset
|
485 |
by (ALLGOALS (asm_full_simp_tac (simpset() addsimps add_ac))); |
6070 | 486 |
qed "add_mult_distrib_left"; |
0 | 487 |
|
488 |
(*Associative law for multiplication*) |
|
9491
1a36151ee2fc
natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents:
9301
diff
changeset
|
489 |
Goal "(m #* n) #* k = m #* (n #* k)"; |
1a36151ee2fc
natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents:
9301
diff
changeset
|
490 |
by (subgoal_tac "(natify(m) #* natify(n)) #* natify(k) = \ |
1a36151ee2fc
natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents:
9301
diff
changeset
|
491 |
\ natify(m) #* (natify(n) #* natify(k))" 1); |
1a36151ee2fc
natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents:
9301
diff
changeset
|
492 |
by (res_inst_tac [("n","natify(m)")] nat_induct 2); |
1a36151ee2fc
natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents:
9301
diff
changeset
|
493 |
by (ALLGOALS (asm_full_simp_tac (simpset() addsimps [add_mult_distrib]))); |
6070 | 494 |
qed "mult_assoc"; |
0 | 495 |
|
437 | 496 |
(*for a/c rewriting*) |
9491
1a36151ee2fc
natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents:
9301
diff
changeset
|
497 |
Goal "m #* (n #* k) = n #* (m #* k)"; |
6070 | 498 |
by (rtac (mult_commute RS trans) 1); |
9491
1a36151ee2fc
natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents:
9301
diff
changeset
|
499 |
by (rtac (mult_assoc RS trans) 1); |
1a36151ee2fc
natify, a coercion to reduce the number of type constraints in arithmetic
paulson
parents:
9301
diff
changeset
|
500 |
by (rtac (mult_commute RS subst_context) 1); |
6070 | 501 |
qed "mult_left_commute"; |
437 | 502 |
|
9907 | 503 |
bind_thms ("mult_ac", [mult_assoc,mult_commute,mult_left_commute]); |