author | paulson |
Tue, 20 May 1997 11:38:50 +0200 | |
changeset 3235 | 351565b7321b |
parent 2683 | be7b439baef2 |
child 3725 | c7fa890d0d92 |
permissions | -rw-r--r-- |
1465 | 1 |
(* Title: Integ.ML |
925
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
2 |
ID: $Id$ |
2215 | 3 |
Authors: Lawrence C Paulson, Cambridge University Computer Laboratory |
925
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
4 |
Copyright 1993 University of Cambridge |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
5 |
|
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
6 |
The integers as equivalence classes over nat*nat. |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
7 |
|
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
8 |
Could also prove... |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
9 |
"znegative(z) ==> $# zmagnitude(z) = $~ z" |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
10 |
"~ znegative(z) ==> $# zmagnitude(z) = z" |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
11 |
< is a linear ordering |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
12 |
+ and * are monotonic wrt < |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
13 |
*) |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
14 |
|
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
15 |
open Integ; |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
16 |
|
1894 | 17 |
Delrules [equalityI]; |
18 |
||
925
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
19 |
|
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
20 |
(*** Proving that intrel is an equivalence relation ***) |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
21 |
|
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
22 |
val eqa::eqb::prems = goal Arith.thy |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
23 |
"[| (x1::nat) + y2 = x2 + y1; x2 + y3 = x3 + y2 |] ==> \ |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
24 |
\ x1 + y3 = x3 + y1"; |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
25 |
by (res_inst_tac [("k2","x2")] (add_left_cancel RS iffD1) 1); |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
26 |
by (rtac (add_left_commute RS trans) 1); |
2036 | 27 |
by (stac eqb 1); |
925
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
28 |
by (rtac (add_left_commute RS trans) 1); |
2036 | 29 |
by (stac eqa 1); |
925
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
30 |
by (rtac (add_left_commute) 1); |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
31 |
qed "integ_trans_lemma"; |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
32 |
|
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
33 |
(** Natural deduction for intrel **) |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
34 |
|
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
35 |
val prems = goalw Integ.thy [intrel_def] |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
36 |
"[| x1+y2 = x2+y1|] ==> \ |
972
e61b058d58d2
changed syntax of tuples from <..., ...> to (..., ...)
clasohm
parents:
925
diff
changeset
|
37 |
\ ((x1,y1),(x2,y2)): intrel"; |
1894 | 38 |
by (fast_tac (!claset addIs prems) 1); |
925
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
39 |
qed "intrelI"; |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
40 |
|
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
41 |
(*intrelE is hard to derive because fast_tac tries hyp_subst_tac so soon*) |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
42 |
goalw Integ.thy [intrel_def] |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
43 |
"p: intrel --> (EX x1 y1 x2 y2. \ |
972
e61b058d58d2
changed syntax of tuples from <..., ...> to (..., ...)
clasohm
parents:
925
diff
changeset
|
44 |
\ p = ((x1,y1),(x2,y2)) & x1+y2 = x2+y1)"; |
1894 | 45 |
by (Fast_tac 1); |
925
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
46 |
qed "intrelE_lemma"; |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
47 |
|
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
48 |
val [major,minor] = goal Integ.thy |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
49 |
"[| p: intrel; \ |
972
e61b058d58d2
changed syntax of tuples from <..., ...> to (..., ...)
clasohm
parents:
925
diff
changeset
|
50 |
\ !!x1 y1 x2 y2. [| p = ((x1,y1),(x2,y2)); x1+y2 = x2+y1|] ==> Q |] \ |
925
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
51 |
\ ==> Q"; |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
52 |
by (cut_facts_tac [major RS (intrelE_lemma RS mp)] 1); |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
53 |
by (REPEAT (eresolve_tac [asm_rl,exE,conjE,minor] 1)); |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
54 |
qed "intrelE"; |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
55 |
|
1894 | 56 |
AddSIs [intrelI]; |
57 |
AddSEs [intrelE]; |
|
925
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
58 |
|
972
e61b058d58d2
changed syntax of tuples from <..., ...> to (..., ...)
clasohm
parents:
925
diff
changeset
|
59 |
goal Integ.thy "((x1,y1),(x2,y2)): intrel = (x1+y2 = x2+y1)"; |
1894 | 60 |
by (Fast_tac 1); |
925
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
61 |
qed "intrel_iff"; |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
62 |
|
972
e61b058d58d2
changed syntax of tuples from <..., ...> to (..., ...)
clasohm
parents:
925
diff
changeset
|
63 |
goal Integ.thy "(x,x): intrel"; |
2036 | 64 |
by (stac surjective_pairing 1 THEN rtac (refl RS intrelI) 1); |
925
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
65 |
qed "intrel_refl"; |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
66 |
|
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
67 |
goalw Integ.thy [equiv_def, refl_def, sym_def, trans_def] |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
68 |
"equiv {x::(nat*nat).True} intrel"; |
1894 | 69 |
by (fast_tac (!claset addSIs [intrel_refl] |
925
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
70 |
addSEs [sym, integ_trans_lemma]) 1); |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
71 |
qed "equiv_intrel"; |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
72 |
|
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
73 |
val equiv_intrel_iff = |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
74 |
[TrueI, TrueI] MRS |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
75 |
([CollectI, CollectI] MRS |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
76 |
(equiv_intrel RS eq_equiv_class_iff)); |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
77 |
|
972
e61b058d58d2
changed syntax of tuples from <..., ...> to (..., ...)
clasohm
parents:
925
diff
changeset
|
78 |
goalw Integ.thy [Integ_def,intrel_def,quotient_def] "intrel^^{(x,y)}:Integ"; |
1894 | 79 |
by (Fast_tac 1); |
925
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
80 |
qed "intrel_in_integ"; |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
81 |
|
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
82 |
goal Integ.thy "inj_onto Abs_Integ Integ"; |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
83 |
by (rtac inj_onto_inverseI 1); |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
84 |
by (etac Abs_Integ_inverse 1); |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
85 |
qed "inj_onto_Abs_Integ"; |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
86 |
|
1266 | 87 |
Addsimps [equiv_intrel_iff, inj_onto_Abs_Integ RS inj_onto_iff, |
1465 | 88 |
intrel_iff, intrel_in_integ, Abs_Integ_inverse]; |
925
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
89 |
|
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
90 |
goal Integ.thy "inj(Rep_Integ)"; |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
91 |
by (rtac inj_inverseI 1); |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
92 |
by (rtac Rep_Integ_inverse 1); |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
93 |
qed "inj_Rep_Integ"; |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
94 |
|
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
95 |
|
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
96 |
|
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
97 |
|
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
98 |
(** znat: the injection from nat to Integ **) |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
99 |
|
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
100 |
goal Integ.thy "inj(znat)"; |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
101 |
by (rtac injI 1); |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
102 |
by (rewtac znat_def); |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
103 |
by (dtac (inj_onto_Abs_Integ RS inj_ontoD) 1); |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
104 |
by (REPEAT (rtac intrel_in_integ 1)); |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
105 |
by (dtac eq_equiv_class 1); |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
106 |
by (rtac equiv_intrel 1); |
1894 | 107 |
by (Fast_tac 1); |
108 |
by (safe_tac (!claset)); |
|
1266 | 109 |
by (Asm_full_simp_tac 1); |
925
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
110 |
qed "inj_znat"; |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
111 |
|
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
112 |
|
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
113 |
(**** zminus: unary negation on Integ ****) |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
114 |
|
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
115 |
goalw Integ.thy [congruent_def] |
972
e61b058d58d2
changed syntax of tuples from <..., ...> to (..., ...)
clasohm
parents:
925
diff
changeset
|
116 |
"congruent intrel (%p. split (%x y. intrel^^{(y,x)}) p)"; |
1894 | 117 |
by (safe_tac (!claset)); |
1266 | 118 |
by (asm_simp_tac (!simpset addsimps add_ac) 1); |
925
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
119 |
qed "zminus_congruent"; |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
120 |
|
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
121 |
|
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
122 |
(*Resolve th against the corresponding facts for zminus*) |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
123 |
val zminus_ize = RSLIST [equiv_intrel, zminus_congruent]; |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
124 |
|
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
125 |
goalw Integ.thy [zminus_def] |
972
e61b058d58d2
changed syntax of tuples from <..., ...> to (..., ...)
clasohm
parents:
925
diff
changeset
|
126 |
"$~ Abs_Integ(intrel^^{(x,y)}) = Abs_Integ(intrel ^^ {(y,x)})"; |
925
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
127 |
by (res_inst_tac [("f","Abs_Integ")] arg_cong 1); |
1266 | 128 |
by (simp_tac (!simpset addsimps |
925
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
129 |
[intrel_in_integ RS Abs_Integ_inverse,zminus_ize UN_equiv_class]) 1); |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
130 |
qed "zminus"; |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
131 |
|
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
132 |
(*by lcp*) |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
133 |
val [prem] = goal Integ.thy |
972
e61b058d58d2
changed syntax of tuples from <..., ...> to (..., ...)
clasohm
parents:
925
diff
changeset
|
134 |
"(!!x y. z = Abs_Integ(intrel^^{(x,y)}) ==> P) ==> P"; |
925
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
135 |
by (res_inst_tac [("x1","z")] |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
136 |
(rewrite_rule [Integ_def] Rep_Integ RS quotientE) 1); |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
137 |
by (dres_inst_tac [("f","Abs_Integ")] arg_cong 1); |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
138 |
by (res_inst_tac [("p","x")] PairE 1); |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
139 |
by (rtac prem 1); |
1266 | 140 |
by (asm_full_simp_tac (!simpset addsimps [Rep_Integ_inverse]) 1); |
925
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
141 |
qed "eq_Abs_Integ"; |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
142 |
|
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
143 |
goal Integ.thy "$~ ($~ z) = z"; |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
144 |
by (res_inst_tac [("z","z")] eq_Abs_Integ 1); |
1266 | 145 |
by (asm_simp_tac (!simpset addsimps [zminus]) 1); |
925
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
146 |
qed "zminus_zminus"; |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
147 |
|
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
148 |
goal Integ.thy "inj(zminus)"; |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
149 |
by (rtac injI 1); |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
150 |
by (dres_inst_tac [("f","zminus")] arg_cong 1); |
1266 | 151 |
by (asm_full_simp_tac (!simpset addsimps [zminus_zminus]) 1); |
925
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
152 |
qed "inj_zminus"; |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
153 |
|
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
154 |
goalw Integ.thy [znat_def] "$~ ($#0) = $#0"; |
1266 | 155 |
by (simp_tac (!simpset addsimps [zminus]) 1); |
925
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
156 |
qed "zminus_0"; |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
157 |
|
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
158 |
|
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
159 |
(**** znegative: the test for negative integers ****) |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
160 |
|
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
161 |
goal Arith.thy "!!m x n::nat. n+m=x ==> m<=x"; |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
162 |
by (dtac (disjI2 RS less_or_eq_imp_le) 1); |
1266 | 163 |
by (asm_full_simp_tac (!simpset addsimps add_ac) 1); |
925
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
164 |
by (dtac add_leD1 1); |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
165 |
by (assume_tac 1); |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
166 |
qed "not_znegative_znat_lemma"; |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
167 |
|
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
168 |
|
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
169 |
goalw Integ.thy [znegative_def, znat_def] |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
170 |
"~ znegative($# n)"; |
1266 | 171 |
by (Simp_tac 1); |
1894 | 172 |
by (safe_tac (!claset)); |
925
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
173 |
by (rtac ccontr 1); |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
174 |
by (etac notE 1); |
1266 | 175 |
by (Asm_full_simp_tac 1); |
925
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
176 |
by (dtac not_znegative_znat_lemma 1); |
1894 | 177 |
by (fast_tac (!claset addDs [leD]) 1); |
925
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
178 |
qed "not_znegative_znat"; |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
179 |
|
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
180 |
goalw Integ.thy [znegative_def, znat_def] "znegative($~ $# Suc(n))"; |
1266 | 181 |
by (simp_tac (!simpset addsimps [zminus]) 1); |
925
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
182 |
by (REPEAT (ares_tac [exI, conjI] 1)); |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
183 |
by (rtac (intrelI RS ImageI) 2); |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
184 |
by (rtac singletonI 3); |
1266 | 185 |
by (Simp_tac 2); |
925
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
186 |
by (rtac less_add_Suc1 1); |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
187 |
qed "znegative_zminus_znat"; |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
188 |
|
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
189 |
|
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
190 |
(**** zmagnitude: magnitide of an integer, as a natural number ****) |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
191 |
|
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
192 |
goal Arith.thy "!!n::nat. n - Suc(n+m)=0"; |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
193 |
by (nat_ind_tac "n" 1); |
1266 | 194 |
by (ALLGOALS Asm_simp_tac); |
925
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
195 |
qed "diff_Suc_add_0"; |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
196 |
|
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
197 |
goal Arith.thy "Suc((n::nat)+m)-n=Suc(m)"; |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
198 |
by (nat_ind_tac "n" 1); |
1266 | 199 |
by (ALLGOALS Asm_simp_tac); |
925
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
200 |
qed "diff_Suc_add_inverse"; |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
201 |
|
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
202 |
goalw Integ.thy [congruent_def] |
972
e61b058d58d2
changed syntax of tuples from <..., ...> to (..., ...)
clasohm
parents:
925
diff
changeset
|
203 |
"congruent intrel (split (%x y. intrel^^{((y-x) + (x-(y::nat)),0)}))"; |
1894 | 204 |
by (safe_tac (!claset)); |
1266 | 205 |
by (Asm_simp_tac 1); |
925
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
206 |
by (etac rev_mp 1); |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
207 |
by (res_inst_tac [("m","x1"),("n","y1")] diff_induct 1); |
1266 | 208 |
by (asm_simp_tac (!simpset addsimps [inj_Suc RS inj_eq]) 3); |
209 |
by (asm_simp_tac (!simpset addsimps [diff_add_inverse,diff_add_0]) 2); |
|
210 |
by (Asm_simp_tac 1); |
|
925
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
211 |
by (rtac impI 1); |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
212 |
by (etac subst 1); |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
213 |
by (res_inst_tac [("m1","x")] (add_commute RS ssubst) 1); |
1266 | 214 |
by (asm_simp_tac (!simpset addsimps [diff_add_inverse,diff_add_0]) 1); |
925
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
215 |
by (rtac impI 1); |
1266 | 216 |
by (asm_simp_tac (!simpset addsimps |
1465 | 217 |
[diff_add_inverse, diff_add_0, diff_Suc_add_0, |
218 |
diff_Suc_add_inverse]) 1); |
|
925
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
219 |
qed "zmagnitude_congruent"; |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
220 |
|
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
221 |
(*Resolve th against the corresponding facts for zmagnitude*) |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
222 |
val zmagnitude_ize = RSLIST [equiv_intrel, zmagnitude_congruent]; |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
223 |
|
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
224 |
|
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
225 |
goalw Integ.thy [zmagnitude_def] |
972
e61b058d58d2
changed syntax of tuples from <..., ...> to (..., ...)
clasohm
parents:
925
diff
changeset
|
226 |
"zmagnitude (Abs_Integ(intrel^^{(x,y)})) = \ |
e61b058d58d2
changed syntax of tuples from <..., ...> to (..., ...)
clasohm
parents:
925
diff
changeset
|
227 |
\ Abs_Integ(intrel^^{((y - x) + (x - y),0)})"; |
925
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
228 |
by (res_inst_tac [("f","Abs_Integ")] arg_cong 1); |
1266 | 229 |
by (asm_simp_tac (!simpset addsimps [zmagnitude_ize UN_equiv_class]) 1); |
925
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
230 |
qed "zmagnitude"; |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
231 |
|
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
232 |
goalw Integ.thy [znat_def] "zmagnitude($# n) = $#n"; |
1266 | 233 |
by (asm_simp_tac (!simpset addsimps [zmagnitude]) 1); |
925
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
234 |
qed "zmagnitude_znat"; |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
235 |
|
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
236 |
goalw Integ.thy [znat_def] "zmagnitude($~ $# n) = $#n"; |
1266 | 237 |
by (asm_simp_tac (!simpset addsimps [zmagnitude, zminus]) 1); |
925
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
238 |
qed "zmagnitude_zminus_znat"; |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
239 |
|
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
240 |
|
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
241 |
(**** zadd: addition on Integ ****) |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
242 |
|
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
243 |
(** Congruence property for addition **) |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
244 |
|
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
245 |
goalw Integ.thy [congruent2_def] |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
246 |
"congruent2 intrel (%p1 p2. \ |
972
e61b058d58d2
changed syntax of tuples from <..., ...> to (..., ...)
clasohm
parents:
925
diff
changeset
|
247 |
\ split (%x1 y1. split (%x2 y2. intrel^^{(x1+x2, y1+y2)}) p2) p1)"; |
925
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
248 |
(*Proof via congruent2_commuteI seems longer*) |
1894 | 249 |
by (safe_tac (!claset)); |
1266 | 250 |
by (asm_simp_tac (!simpset addsimps [add_assoc]) 1); |
925
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
251 |
(*The rest should be trivial, but rearranging terms is hard*) |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
252 |
by (res_inst_tac [("x1","x1a")] (add_left_commute RS ssubst) 1); |
1266 | 253 |
by (asm_simp_tac (!simpset addsimps [add_assoc RS sym]) 1); |
254 |
by (asm_simp_tac (!simpset addsimps add_ac) 1); |
|
925
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
255 |
qed "zadd_congruent2"; |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
256 |
|
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
257 |
(*Resolve th against the corresponding facts for zadd*) |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
258 |
val zadd_ize = RSLIST [equiv_intrel, zadd_congruent2]; |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
259 |
|
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
260 |
goalw Integ.thy [zadd_def] |
972
e61b058d58d2
changed syntax of tuples from <..., ...> to (..., ...)
clasohm
parents:
925
diff
changeset
|
261 |
"Abs_Integ(intrel^^{(x1,y1)}) + Abs_Integ(intrel^^{(x2,y2)}) = \ |
e61b058d58d2
changed syntax of tuples from <..., ...> to (..., ...)
clasohm
parents:
925
diff
changeset
|
262 |
\ Abs_Integ(intrel^^{(x1+x2, y1+y2)})"; |
925
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
263 |
by (asm_simp_tac |
1266 | 264 |
(!simpset addsimps [zadd_ize UN_equiv_class2]) 1); |
925
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
265 |
qed "zadd"; |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
266 |
|
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
267 |
goalw Integ.thy [znat_def] "$#0 + z = z"; |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
268 |
by (res_inst_tac [("z","z")] eq_Abs_Integ 1); |
1266 | 269 |
by (asm_simp_tac (!simpset addsimps [zadd]) 1); |
925
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
270 |
qed "zadd_0"; |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
271 |
|
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
272 |
goal Integ.thy "$~ (z + w) = $~ z + $~ w"; |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
273 |
by (res_inst_tac [("z","z")] eq_Abs_Integ 1); |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
274 |
by (res_inst_tac [("z","w")] eq_Abs_Integ 1); |
1266 | 275 |
by (asm_simp_tac (!simpset addsimps [zminus,zadd]) 1); |
925
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
276 |
qed "zminus_zadd_distrib"; |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
277 |
|
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
278 |
goal Integ.thy "(z::int) + w = w + z"; |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
279 |
by (res_inst_tac [("z","z")] eq_Abs_Integ 1); |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
280 |
by (res_inst_tac [("z","w")] eq_Abs_Integ 1); |
1266 | 281 |
by (asm_simp_tac (!simpset addsimps (add_ac @ [zadd])) 1); |
925
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
282 |
qed "zadd_commute"; |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
283 |
|
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
284 |
goal Integ.thy "((z1::int) + z2) + z3 = z1 + (z2 + z3)"; |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
285 |
by (res_inst_tac [("z","z1")] eq_Abs_Integ 1); |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
286 |
by (res_inst_tac [("z","z2")] eq_Abs_Integ 1); |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
287 |
by (res_inst_tac [("z","z3")] eq_Abs_Integ 1); |
1266 | 288 |
by (asm_simp_tac (!simpset addsimps [zadd, add_assoc]) 1); |
925
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
289 |
qed "zadd_assoc"; |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
290 |
|
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
291 |
(*For AC rewriting*) |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
292 |
goal Integ.thy "(x::int)+(y+z)=y+(x+z)"; |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
293 |
by (rtac (zadd_commute RS trans) 1); |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
294 |
by (rtac (zadd_assoc RS trans) 1); |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
295 |
by (rtac (zadd_commute RS arg_cong) 1); |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
296 |
qed "zadd_left_commute"; |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
297 |
|
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
298 |
(*Integer addition is an AC operator*) |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
299 |
val zadd_ac = [zadd_assoc,zadd_commute,zadd_left_commute]; |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
300 |
|
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
301 |
goalw Integ.thy [znat_def] "$# (m + n) = ($#m) + ($#n)"; |
1266 | 302 |
by (asm_simp_tac (!simpset addsimps [zadd]) 1); |
925
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
303 |
qed "znat_add"; |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
304 |
|
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
305 |
goalw Integ.thy [znat_def] "z + ($~ z) = $#0"; |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
306 |
by (res_inst_tac [("z","z")] eq_Abs_Integ 1); |
1266 | 307 |
by (asm_simp_tac (!simpset addsimps [zminus, zadd, add_commute]) 1); |
925
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
308 |
qed "zadd_zminus_inverse"; |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
309 |
|
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
310 |
goal Integ.thy "($~ z) + z = $#0"; |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
311 |
by (rtac (zadd_commute RS trans) 1); |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
312 |
by (rtac zadd_zminus_inverse 1); |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
313 |
qed "zadd_zminus_inverse2"; |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
314 |
|
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
315 |
goal Integ.thy "z + $#0 = z"; |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
316 |
by (rtac (zadd_commute RS trans) 1); |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
317 |
by (rtac zadd_0 1); |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
318 |
qed "zadd_0_right"; |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
319 |
|
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
320 |
|
2224
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
paulson
parents:
2215
diff
changeset
|
321 |
(** Lemmas **) |
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
paulson
parents:
2215
diff
changeset
|
322 |
|
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
paulson
parents:
2215
diff
changeset
|
323 |
qed_goal "zadd_assoc_cong" Integ.thy |
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
paulson
parents:
2215
diff
changeset
|
324 |
"!!z. (z::int) + v = z' + v' ==> z + (v + w) = z' + (v' + w)" |
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
paulson
parents:
2215
diff
changeset
|
325 |
(fn _ => [(asm_simp_tac (!simpset addsimps [zadd_assoc RS sym]) 1)]); |
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
paulson
parents:
2215
diff
changeset
|
326 |
|
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
paulson
parents:
2215
diff
changeset
|
327 |
qed_goal "zadd_assoc_swap" Integ.thy "(z::int) + (v + w) = v + (z + w)" |
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
paulson
parents:
2215
diff
changeset
|
328 |
(fn _ => [(REPEAT (ares_tac [zadd_commute RS zadd_assoc_cong] 1))]); |
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
paulson
parents:
2215
diff
changeset
|
329 |
|
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
paulson
parents:
2215
diff
changeset
|
330 |
|
925
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
331 |
(*Need properties of subtraction? Or use $- just as an abbreviation!*) |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
332 |
|
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
333 |
(**** zmult: multiplication on Integ ****) |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
334 |
|
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
335 |
(** Congruence property for multiplication **) |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
336 |
|
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
337 |
goal Integ.thy "((k::nat) + l) + (m + n) = (k + m) + (n + l)"; |
1266 | 338 |
by (simp_tac (!simpset addsimps add_ac) 1); |
925
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
339 |
qed "zmult_congruent_lemma"; |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
340 |
|
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
341 |
goal Integ.thy |
1465 | 342 |
"congruent2 intrel (%p1 p2. \ |
343 |
\ split (%x1 y1. split (%x2 y2. \ |
|
972
e61b058d58d2
changed syntax of tuples from <..., ...> to (..., ...)
clasohm
parents:
925
diff
changeset
|
344 |
\ intrel^^{(x1*x2 + y1*y2, x1*y2 + y1*x2)}) p2) p1)"; |
925
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
345 |
by (rtac (equiv_intrel RS congruent2_commuteI) 1); |
1894 | 346 |
by (safe_tac (!claset)); |
925
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
347 |
by (rewtac split_def); |
1266 | 348 |
by (simp_tac (!simpset addsimps add_ac@mult_ac) 1); |
349 |
by (asm_simp_tac (!simpset delsimps [equiv_intrel_iff] |
|
350 |
addsimps add_ac@mult_ac) 1); |
|
925
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
351 |
by (rtac (intrelI RS(equiv_intrel RS equiv_class_eq)) 1); |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
352 |
by (rtac (zmult_congruent_lemma RS trans) 1); |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
353 |
by (rtac (zmult_congruent_lemma RS trans RS sym) 1); |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
354 |
by (rtac (zmult_congruent_lemma RS trans RS sym) 1); |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
355 |
by (rtac (zmult_congruent_lemma RS trans RS sym) 1); |
1798
c055505f36d1
Explicitly included add_mult_distrib & add_mult_distrib2
paulson
parents:
1619
diff
changeset
|
356 |
by (asm_simp_tac (!simpset addsimps [add_mult_distrib RS sym]) 1); |
1266 | 357 |
by (asm_simp_tac (!simpset addsimps add_ac@mult_ac) 1); |
925
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
358 |
qed "zmult_congruent2"; |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
359 |
|
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
360 |
(*Resolve th against the corresponding facts for zmult*) |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
361 |
val zmult_ize = RSLIST [equiv_intrel, zmult_congruent2]; |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
362 |
|
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
363 |
goalw Integ.thy [zmult_def] |
1465 | 364 |
"Abs_Integ((intrel^^{(x1,y1)})) * Abs_Integ((intrel^^{(x2,y2)})) = \ |
972
e61b058d58d2
changed syntax of tuples from <..., ...> to (..., ...)
clasohm
parents:
925
diff
changeset
|
365 |
\ Abs_Integ(intrel ^^ {(x1*x2 + y1*y2, x1*y2 + y1*x2)})"; |
1266 | 366 |
by (simp_tac (!simpset addsimps [zmult_ize UN_equiv_class2]) 1); |
925
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
367 |
qed "zmult"; |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
368 |
|
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
369 |
goalw Integ.thy [znat_def] "$#0 * z = $#0"; |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
370 |
by (res_inst_tac [("z","z")] eq_Abs_Integ 1); |
1266 | 371 |
by (asm_simp_tac (!simpset addsimps [zmult]) 1); |
925
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
372 |
qed "zmult_0"; |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
373 |
|
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
374 |
goalw Integ.thy [znat_def] "$#Suc(0) * z = z"; |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
375 |
by (res_inst_tac [("z","z")] eq_Abs_Integ 1); |
1266 | 376 |
by (asm_simp_tac (!simpset addsimps [zmult]) 1); |
925
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
377 |
qed "zmult_1"; |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
378 |
|
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
379 |
goal Integ.thy "($~ z) * w = $~ (z * w)"; |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
380 |
by (res_inst_tac [("z","z")] eq_Abs_Integ 1); |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
381 |
by (res_inst_tac [("z","w")] eq_Abs_Integ 1); |
1266 | 382 |
by (asm_simp_tac (!simpset addsimps ([zminus, zmult] @ add_ac)) 1); |
925
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
383 |
qed "zmult_zminus"; |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
384 |
|
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
385 |
|
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
386 |
goal Integ.thy "($~ z) * ($~ w) = (z * w)"; |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
387 |
by (res_inst_tac [("z","z")] eq_Abs_Integ 1); |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
388 |
by (res_inst_tac [("z","w")] eq_Abs_Integ 1); |
1266 | 389 |
by (asm_simp_tac (!simpset addsimps ([zminus, zmult] @ add_ac)) 1); |
925
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
390 |
qed "zmult_zminus_zminus"; |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
391 |
|
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
392 |
goal Integ.thy "(z::int) * w = w * z"; |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
393 |
by (res_inst_tac [("z","z")] eq_Abs_Integ 1); |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
394 |
by (res_inst_tac [("z","w")] eq_Abs_Integ 1); |
1266 | 395 |
by (asm_simp_tac (!simpset addsimps ([zmult] @ add_ac @ mult_ac)) 1); |
925
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
396 |
qed "zmult_commute"; |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
397 |
|
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
398 |
goal Integ.thy "z * $# 0 = $#0"; |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
399 |
by (rtac ([zmult_commute, zmult_0] MRS trans) 1); |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
400 |
qed "zmult_0_right"; |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
401 |
|
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
402 |
goal Integ.thy "z * $#Suc(0) = z"; |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
403 |
by (rtac ([zmult_commute, zmult_1] MRS trans) 1); |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
404 |
qed "zmult_1_right"; |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
405 |
|
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
406 |
goal Integ.thy "((z1::int) * z2) * z3 = z1 * (z2 * z3)"; |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
407 |
by (res_inst_tac [("z","z1")] eq_Abs_Integ 1); |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
408 |
by (res_inst_tac [("z","z2")] eq_Abs_Integ 1); |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
409 |
by (res_inst_tac [("z","z3")] eq_Abs_Integ 1); |
1798
c055505f36d1
Explicitly included add_mult_distrib & add_mult_distrib2
paulson
parents:
1619
diff
changeset
|
410 |
by (asm_simp_tac (!simpset addsimps ([add_mult_distrib2,zmult] @ |
2036 | 411 |
add_ac @ mult_ac)) 1); |
925
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
412 |
qed "zmult_assoc"; |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
413 |
|
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
414 |
(*For AC rewriting*) |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
415 |
qed_goal "zmult_left_commute" Integ.thy |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
416 |
"(z1::int)*(z2*z3) = z2*(z1*z3)" |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
417 |
(fn _ => [rtac (zmult_commute RS trans) 1, rtac (zmult_assoc RS trans) 1, |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
418 |
rtac (zmult_commute RS arg_cong) 1]); |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
419 |
|
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
420 |
(*Integer multiplication is an AC operator*) |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
421 |
val zmult_ac = [zmult_assoc, zmult_commute, zmult_left_commute]; |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
422 |
|
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
423 |
goal Integ.thy "((z1::int) + z2) * w = (z1 * w) + (z2 * w)"; |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
424 |
by (res_inst_tac [("z","z1")] eq_Abs_Integ 1); |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
425 |
by (res_inst_tac [("z","z2")] eq_Abs_Integ 1); |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
426 |
by (res_inst_tac [("z","w")] eq_Abs_Integ 1); |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
427 |
by (asm_simp_tac |
1798
c055505f36d1
Explicitly included add_mult_distrib & add_mult_distrib2
paulson
parents:
1619
diff
changeset
|
428 |
(!simpset addsimps ([add_mult_distrib2, zadd, zmult] @ |
2036 | 429 |
add_ac @ mult_ac)) 1); |
925
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
430 |
qed "zadd_zmult_distrib"; |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
431 |
|
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
432 |
val zmult_commute'= read_instantiate [("z","w")] zmult_commute; |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
433 |
|
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
434 |
goal Integ.thy "w * ($~ z) = $~ (w * z)"; |
1266 | 435 |
by (simp_tac (!simpset addsimps [zmult_commute', zmult_zminus]) 1); |
925
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
436 |
qed "zmult_zminus_right"; |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
437 |
|
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
438 |
goal Integ.thy "(w::int) * (z1 + z2) = (w * z1) + (w * z2)"; |
1266 | 439 |
by (simp_tac (!simpset addsimps [zmult_commute',zadd_zmult_distrib]) 1); |
925
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
440 |
qed "zadd_zmult_distrib2"; |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
441 |
|
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
442 |
val zadd_simps = |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
443 |
[zadd_0, zadd_0_right, zadd_zminus_inverse, zadd_zminus_inverse2]; |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
444 |
|
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
445 |
val zminus_simps = [zminus_zminus, zminus_0, zminus_zadd_distrib]; |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
446 |
|
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
447 |
val zmult_simps = [zmult_0, zmult_1, zmult_0_right, zmult_1_right, |
1465 | 448 |
zmult_zminus, zmult_zminus_right]; |
925
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
449 |
|
1266 | 450 |
Addsimps (zadd_simps @ zminus_simps @ zmult_simps @ |
451 |
[zmagnitude_znat, zmagnitude_zminus_znat]); |
|
925
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
452 |
|
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
453 |
|
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
454 |
(**** Additional Theorems (by Mattolini; proofs mainly by lcp) ****) |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
455 |
|
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
456 |
(* Some Theorems about zsuc and zpred *) |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
457 |
goalw Integ.thy [zsuc_def] "$#(Suc(n)) = zsuc($# n)"; |
1266 | 458 |
by (simp_tac (!simpset addsimps [znat_add RS sym]) 1); |
925
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
459 |
qed "znat_Suc"; |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
460 |
|
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
461 |
goalw Integ.thy [zpred_def,zsuc_def,zdiff_def] "$~ zsuc(z) = zpred($~ z)"; |
1266 | 462 |
by (Simp_tac 1); |
925
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
463 |
qed "zminus_zsuc"; |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
464 |
|
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
465 |
goalw Integ.thy [zpred_def,zsuc_def,zdiff_def] "$~ zpred(z) = zsuc($~ z)"; |
1266 | 466 |
by (Simp_tac 1); |
925
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
467 |
qed "zminus_zpred"; |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
468 |
|
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
469 |
goalw Integ.thy [zsuc_def,zpred_def,zdiff_def] |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
470 |
"zpred(zsuc(z)) = z"; |
1266 | 471 |
by (simp_tac (!simpset addsimps [zadd_assoc]) 1); |
925
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
472 |
qed "zpred_zsuc"; |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
473 |
|
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
474 |
goalw Integ.thy [zsuc_def,zpred_def,zdiff_def] |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
475 |
"zsuc(zpred(z)) = z"; |
1266 | 476 |
by (simp_tac (!simpset addsimps [zadd_assoc]) 1); |
925
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
477 |
qed "zsuc_zpred"; |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
478 |
|
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
479 |
goal Integ.thy "(zpred(z)=w) = (z=zsuc(w))"; |
1894 | 480 |
by (safe_tac (!claset)); |
925
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
481 |
by (rtac (zsuc_zpred RS sym) 1); |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
482 |
by (rtac zpred_zsuc 1); |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
483 |
qed "zpred_to_zsuc"; |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
484 |
|
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
485 |
goal Integ.thy "(zsuc(z)=w)=(z=zpred(w))"; |
1894 | 486 |
by (safe_tac (!claset)); |
925
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
487 |
by (rtac (zpred_zsuc RS sym) 1); |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
488 |
by (rtac zsuc_zpred 1); |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
489 |
qed "zsuc_to_zpred"; |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
490 |
|
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
491 |
goal Integ.thy "($~ z = w) = (z = $~ w)"; |
1894 | 492 |
by (safe_tac (!claset)); |
925
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
493 |
by (rtac (zminus_zminus RS sym) 1); |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
494 |
by (rtac zminus_zminus 1); |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
495 |
qed "zminus_exchange"; |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
496 |
|
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
497 |
goal Integ.thy"(zsuc(z)=zsuc(w)) = (z=w)"; |
1894 | 498 |
by (safe_tac (!claset)); |
925
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
499 |
by (dres_inst_tac [("f","zpred")] arg_cong 1); |
1266 | 500 |
by (asm_full_simp_tac (!simpset addsimps [zpred_zsuc]) 1); |
925
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
501 |
qed "bijective_zsuc"; |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
502 |
|
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
503 |
goal Integ.thy"(zpred(z)=zpred(w)) = (z=w)"; |
1894 | 504 |
by (safe_tac (!claset)); |
925
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
505 |
by (dres_inst_tac [("f","zsuc")] arg_cong 1); |
1266 | 506 |
by (asm_full_simp_tac (!simpset addsimps [zsuc_zpred]) 1); |
925
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
507 |
qed "bijective_zpred"; |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
508 |
|
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
509 |
(* Additional Theorems about zadd *) |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
510 |
|
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
511 |
goalw Integ.thy [zsuc_def] "zsuc(z) + w = zsuc(z+w)"; |
1266 | 512 |
by (simp_tac (!simpset addsimps zadd_ac) 1); |
925
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
513 |
qed "zadd_zsuc"; |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
514 |
|
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
515 |
goalw Integ.thy [zsuc_def] "w + zsuc(z) = zsuc(w+z)"; |
1266 | 516 |
by (simp_tac (!simpset addsimps zadd_ac) 1); |
925
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
517 |
qed "zadd_zsuc_right"; |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
518 |
|
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
519 |
goalw Integ.thy [zpred_def,zdiff_def] "zpred(z) + w = zpred(z+w)"; |
1266 | 520 |
by (simp_tac (!simpset addsimps zadd_ac) 1); |
925
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
521 |
qed "zadd_zpred"; |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
522 |
|
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
523 |
goalw Integ.thy [zpred_def,zdiff_def] "w + zpred(z) = zpred(w+z)"; |
1266 | 524 |
by (simp_tac (!simpset addsimps zadd_ac) 1); |
925
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
525 |
qed "zadd_zpred_right"; |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
526 |
|
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
527 |
|
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
528 |
(* Additional Theorems about zmult *) |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
529 |
|
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
530 |
goalw Integ.thy [zsuc_def] "zsuc(w) * z = z + w * z"; |
1266 | 531 |
by (simp_tac (!simpset addsimps [zadd_zmult_distrib, zadd_commute]) 1); |
925
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
532 |
qed "zmult_zsuc"; |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
533 |
|
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
534 |
goalw Integ.thy [zsuc_def] "z * zsuc(w) = z + w * z"; |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
535 |
by (simp_tac |
1266 | 536 |
(!simpset addsimps [zadd_zmult_distrib2, zadd_commute, zmult_commute]) 1); |
925
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
537 |
qed "zmult_zsuc_right"; |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
538 |
|
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
539 |
goalw Integ.thy [zpred_def, zdiff_def] "zpred(w) * z = w * z - z"; |
1266 | 540 |
by (simp_tac (!simpset addsimps [zadd_zmult_distrib]) 1); |
925
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
541 |
qed "zmult_zpred"; |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
542 |
|
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
543 |
goalw Integ.thy [zpred_def, zdiff_def] "z * zpred(w) = w * z - z"; |
1266 | 544 |
by (simp_tac (!simpset addsimps [zadd_zmult_distrib2, zmult_commute]) 1); |
925
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
545 |
qed "zmult_zpred_right"; |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
546 |
|
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
547 |
(* Further Theorems about zsuc and zpred *) |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
548 |
goal Integ.thy "$#Suc(m) ~= $#0"; |
1266 | 549 |
by (simp_tac (!simpset addsimps [inj_znat RS inj_eq]) 1); |
925
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
550 |
qed "znat_Suc_not_znat_Zero"; |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
551 |
|
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
552 |
bind_thm ("znat_Zero_not_znat_Suc", (znat_Suc_not_znat_Zero RS not_sym)); |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
553 |
|
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
554 |
|
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
555 |
goalw Integ.thy [zsuc_def,znat_def] "w ~= zsuc(w)"; |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
556 |
by (res_inst_tac [("z","w")] eq_Abs_Integ 1); |
1266 | 557 |
by (asm_full_simp_tac (!simpset addsimps [zadd]) 1); |
925
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
558 |
qed "n_not_zsuc_n"; |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
559 |
|
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
560 |
val zsuc_n_not_n = n_not_zsuc_n RS not_sym; |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
561 |
|
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
562 |
goal Integ.thy "w ~= zpred(w)"; |
1894 | 563 |
by (safe_tac (!claset)); |
925
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
564 |
by (dres_inst_tac [("x","w"),("f","zsuc")] arg_cong 1); |
1266 | 565 |
by (asm_full_simp_tac (!simpset addsimps [zsuc_zpred,zsuc_n_not_n]) 1); |
925
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
566 |
qed "n_not_zpred_n"; |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
567 |
|
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
568 |
val zpred_n_not_n = n_not_zpred_n RS not_sym; |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
569 |
|
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
570 |
|
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
571 |
(* Theorems about less and less_equal *) |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
572 |
|
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
573 |
goalw Integ.thy [zless_def, znegative_def, zdiff_def, znat_def] |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
574 |
"!!w. w<z ==> ? n. z = w + $#(Suc(n))"; |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
575 |
by (res_inst_tac [("z","z")] eq_Abs_Integ 1); |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
576 |
by (res_inst_tac [("z","w")] eq_Abs_Integ 1); |
1894 | 577 |
by (safe_tac (!claset)); |
1266 | 578 |
by (asm_full_simp_tac (!simpset addsimps [zadd, zminus]) 1); |
1894 | 579 |
by (safe_tac (!claset addSDs [less_eq_Suc_add])); |
2596
3b4ad6c7726f
Modified proofs because of added "triv_forall_equality".
nipkow
parents:
2224
diff
changeset
|
580 |
by (rename_tac "k" 1); |
925
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
581 |
by (res_inst_tac [("x","k")] exI 1); |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
582 |
(*To cancel x2, rename it to be first!*) |
2683 | 583 |
by (rename_tac "a b c" 1); |
584 |
by (Asm_full_simp_tac 1); |
|
1266 | 585 |
by (asm_full_simp_tac (!simpset delsimps [add_Suc_right] |
2683 | 586 |
addsimps ([add_Suc_right RS sym] @ add_ac)) 1); |
925
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
587 |
qed "zless_eq_zadd_Suc"; |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
588 |
|
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
589 |
goalw Integ.thy [zless_def, znegative_def, zdiff_def, znat_def] |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
590 |
"z < z + $#(Suc(n))"; |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
591 |
by (res_inst_tac [("z","z")] eq_Abs_Integ 1); |
1894 | 592 |
by (safe_tac (!claset)); |
1266 | 593 |
by (simp_tac (!simpset addsimps [zadd, zminus]) 1); |
925
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
594 |
by (REPEAT_SOME (ares_tac [refl, exI, singletonI, ImageI, conjI, intrelI])); |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
595 |
by (rtac le_less_trans 1); |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
596 |
by (rtac lessI 2); |
1266 | 597 |
by (asm_simp_tac (!simpset addsimps ([le_add1,add_left_cancel_le]@add_ac)) 1); |
925
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
598 |
qed "zless_zadd_Suc"; |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
599 |
|
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
600 |
goal Integ.thy "!!z1 z2 z3. [| z1<z2; z2<z3 |] ==> z1 < (z3::int)"; |
1894 | 601 |
by (safe_tac (!claset addSDs [zless_eq_zadd_Suc])); |
925
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
602 |
by (simp_tac |
1266 | 603 |
(!simpset addsimps [zadd_assoc, zless_zadd_Suc, znat_add RS sym]) 1); |
925
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
604 |
qed "zless_trans"; |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
605 |
|
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
606 |
goalw Integ.thy [zsuc_def] "z<zsuc(z)"; |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
607 |
by (rtac zless_zadd_Suc 1); |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
608 |
qed "zlessI"; |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
609 |
|
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
610 |
val zless_zsucI = zlessI RSN (2,zless_trans); |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
611 |
|
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
612 |
goal Integ.thy "!!z w::int. z<w ==> ~w<z"; |
1894 | 613 |
by (safe_tac (!claset addSDs [zless_eq_zadd_Suc])); |
925
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
614 |
by (res_inst_tac [("z","z")] eq_Abs_Integ 1); |
1894 | 615 |
by (safe_tac (!claset)); |
1266 | 616 |
by (asm_full_simp_tac (!simpset addsimps ([znat_def, zadd])) 1); |
925
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
617 |
by (asm_full_simp_tac |
1266 | 618 |
(!simpset delsimps [add_Suc_right] addsimps [add_left_cancel, add_assoc, add_Suc_right RS sym]) 1); |
925
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
619 |
by (resolve_tac [less_not_refl2 RS notE] 1); |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
620 |
by (etac sym 2); |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
621 |
by (REPEAT (resolve_tac [lessI, trans_less_add2, less_SucI] 1)); |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
622 |
qed "zless_not_sym"; |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
623 |
|
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
624 |
(* [| n<m; m<n |] ==> R *) |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
625 |
bind_thm ("zless_asym", (zless_not_sym RS notE)); |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
626 |
|
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
627 |
goal Integ.thy "!!z::int. ~ z<z"; |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
628 |
by (resolve_tac [zless_asym RS notI] 1); |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
629 |
by (REPEAT (assume_tac 1)); |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
630 |
qed "zless_not_refl"; |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
631 |
|
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
632 |
(* z<z ==> R *) |
1619 | 633 |
bind_thm ("zless_irrefl", (zless_not_refl RS notE)); |
925
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
634 |
|
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
635 |
goal Integ.thy "!!w. z<w ==> w ~= (z::int)"; |
2036 | 636 |
by (fast_tac (!claset addEs [zless_irrefl]) 1); |
925
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
637 |
qed "zless_not_refl2"; |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
638 |
|
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
639 |
|
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
640 |
(*"Less than" is a linear ordering*) |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
641 |
goalw Integ.thy [zless_def, znegative_def, zdiff_def] |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
642 |
"z<w | z=w | w<(z::int)"; |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
643 |
by (res_inst_tac [("z","z")] eq_Abs_Integ 1); |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
644 |
by (res_inst_tac [("z","w")] eq_Abs_Integ 1); |
1894 | 645 |
by (safe_tac (!claset)); |
925
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
646 |
by (asm_full_simp_tac |
1266 | 647 |
(!simpset addsimps [zadd, zminus, Image_iff, Bex_def]) 1); |
925
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
648 |
by (res_inst_tac [("m1", "x+ya"), ("n1", "xa+y")] (less_linear RS disjE) 1); |
2083
b56425a385b9
Tidied some proofs: changed needed for de Morgan laws
paulson
parents:
2036
diff
changeset
|
649 |
by (REPEAT (fast_tac (!claset addss (!simpset addsimps add_ac)) 1)); |
925
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
650 |
qed "zless_linear"; |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
651 |
|
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
652 |
|
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
653 |
(*** Properties of <= ***) |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
654 |
|
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
655 |
goalw Integ.thy [zless_def, znegative_def, zdiff_def, znat_def] |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
656 |
"($#m < $#n) = (m<n)"; |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
657 |
by (simp_tac |
1266 | 658 |
(!simpset addsimps [zadd, zminus, Image_iff, Bex_def]) 1); |
1894 | 659 |
by (fast_tac (!claset addIs [add_commute] addSEs [less_add_eq_less]) 1); |
925
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
660 |
qed "zless_eq_less"; |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
661 |
|
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
662 |
goalw Integ.thy [zle_def, le_def] "($#m <= $#n) = (m<=n)"; |
1266 | 663 |
by (simp_tac (!simpset addsimps [zless_eq_less]) 1); |
925
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
664 |
qed "zle_eq_le"; |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
665 |
|
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
666 |
goalw Integ.thy [zle_def] "!!w. ~(w<z) ==> z<=(w::int)"; |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
667 |
by (assume_tac 1); |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
668 |
qed "zleI"; |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
669 |
|
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
670 |
goalw Integ.thy [zle_def] "!!w. z<=w ==> ~(w<(z::int))"; |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
671 |
by (assume_tac 1); |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
672 |
qed "zleD"; |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
673 |
|
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
674 |
val zleE = make_elim zleD; |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
675 |
|
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
676 |
goalw Integ.thy [zle_def] "!!z. ~ z <= w ==> w<(z::int)"; |
1894 | 677 |
by (Fast_tac 1); |
925
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
678 |
qed "not_zleE"; |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
679 |
|
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
680 |
goalw Integ.thy [zle_def] "!!z. z < w ==> z <= (w::int)"; |
1894 | 681 |
by (fast_tac (!claset addEs [zless_asym]) 1); |
925
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
682 |
qed "zless_imp_zle"; |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
683 |
|
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
684 |
goalw Integ.thy [zle_def] "!!z. z <= w ==> z < w | z=(w::int)"; |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
685 |
by (cut_facts_tac [zless_linear] 1); |
1894 | 686 |
by (fast_tac (!claset addEs [zless_irrefl,zless_asym]) 1); |
925
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
687 |
qed "zle_imp_zless_or_eq"; |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
688 |
|
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
689 |
goalw Integ.thy [zle_def] "!!z. z<w | z=w ==> z <=(w::int)"; |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
690 |
by (cut_facts_tac [zless_linear] 1); |
1894 | 691 |
by (fast_tac (!claset addEs [zless_irrefl,zless_asym]) 1); |
925
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
692 |
qed "zless_or_eq_imp_zle"; |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
693 |
|
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
694 |
goal Integ.thy "(x <= (y::int)) = (x < y | x=y)"; |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
695 |
by (REPEAT(ares_tac [iffI, zless_or_eq_imp_zle, zle_imp_zless_or_eq] 1)); |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
696 |
qed "zle_eq_zless_or_eq"; |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
697 |
|
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
698 |
goal Integ.thy "w <= (w::int)"; |
1266 | 699 |
by (simp_tac (!simpset addsimps [zle_eq_zless_or_eq]) 1); |
925
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
700 |
qed "zle_refl"; |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
701 |
|
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
702 |
val prems = goal Integ.thy "!!i. [| i <= j; j < k |] ==> i < (k::int)"; |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
703 |
by (dtac zle_imp_zless_or_eq 1); |
1894 | 704 |
by (fast_tac (!claset addIs [zless_trans]) 1); |
925
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
705 |
qed "zle_zless_trans"; |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
706 |
|
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
707 |
goal Integ.thy "!!i. [| i <= j; j <= k |] ==> i <= (k::int)"; |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
708 |
by (EVERY1 [dtac zle_imp_zless_or_eq, dtac zle_imp_zless_or_eq, |
1894 | 709 |
rtac zless_or_eq_imp_zle, fast_tac (!claset addIs [zless_trans])]); |
925
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
710 |
qed "zle_trans"; |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
711 |
|
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
712 |
goal Integ.thy "!!z. [| z <= w; w <= z |] ==> z = (w::int)"; |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
713 |
by (EVERY1 [dtac zle_imp_zless_or_eq, dtac zle_imp_zless_or_eq, |
1894 | 714 |
fast_tac (!claset addEs [zless_irrefl,zless_asym])]); |
925
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
715 |
qed "zle_anti_sym"; |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
716 |
|
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
717 |
|
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
718 |
goal Integ.thy "!!w w' z::int. z + w' = z + w ==> w' = w"; |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
719 |
by (dres_inst_tac [("f", "%x. x + $~z")] arg_cong 1); |
1266 | 720 |
by (asm_full_simp_tac (!simpset addsimps zadd_ac) 1); |
925
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
721 |
qed "zadd_left_cancel"; |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
722 |
|
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
723 |
|
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
724 |
(*** Monotonicity results ***) |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
725 |
|
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
726 |
goal Integ.thy "!!v w z::int. v < w ==> v + z < w + z"; |
1894 | 727 |
by (safe_tac (!claset addSDs [zless_eq_zadd_Suc])); |
1266 | 728 |
by (simp_tac (!simpset addsimps zadd_ac) 1); |
729 |
by (simp_tac (!simpset addsimps [zadd_assoc RS sym, zless_zadd_Suc]) 1); |
|
925
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
730 |
qed "zadd_zless_mono1"; |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
731 |
|
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
732 |
goal Integ.thy "!!v w z::int. (v+z < w+z) = (v < w)"; |
1894 | 733 |
by (safe_tac (!claset addSEs [zadd_zless_mono1])); |
925
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
734 |
by (dres_inst_tac [("z", "$~z")] zadd_zless_mono1 1); |
1266 | 735 |
by (asm_full_simp_tac (!simpset addsimps [zadd_assoc]) 1); |
925
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
736 |
qed "zadd_left_cancel_zless"; |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
737 |
|
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
738 |
goal Integ.thy "!!v w z::int. (v+z <= w+z) = (v <= w)"; |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
739 |
by (asm_full_simp_tac |
1266 | 740 |
(!simpset addsimps [zle_def, zadd_left_cancel_zless]) 1); |
925
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
741 |
qed "zadd_left_cancel_zle"; |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
742 |
|
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
743 |
(*"v<=w ==> v+z <= w+z"*) |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
744 |
bind_thm ("zadd_zle_mono1", zadd_left_cancel_zle RS iffD2); |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
745 |
|
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
746 |
|
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
747 |
goal Integ.thy "!!z' z::int. [| w'<=w; z'<=z |] ==> w' + z' <= w + z"; |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
748 |
by (etac (zadd_zle_mono1 RS zle_trans) 1); |
1266 | 749 |
by (simp_tac (!simpset addsimps [zadd_commute]) 1); |
925
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
750 |
(*w moves to the end because it is free while z', z are bound*) |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
751 |
by (etac zadd_zle_mono1 1); |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
752 |
qed "zadd_zle_mono"; |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
753 |
|
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
754 |
goal Integ.thy "!!w z::int. z<=$#0 ==> w+z <= w"; |
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
755 |
by (dres_inst_tac [("z", "w")] zadd_zle_mono1 1); |
1266 | 756 |
by (asm_full_simp_tac (!simpset addsimps [zadd_commute]) 1); |
925
15539deb6863
new version of HOL/Integ with curried function application
clasohm
parents:
diff
changeset
|
757 |
qed "zadd_zle_self"; |
2224
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
paulson
parents:
2215
diff
changeset
|
758 |
|
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
paulson
parents:
2215
diff
changeset
|
759 |
|
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
paulson
parents:
2215
diff
changeset
|
760 |
(**** Comparisons: lemmas and proofs by Norbert Völker ****) |
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
paulson
parents:
2215
diff
changeset
|
761 |
|
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
paulson
parents:
2215
diff
changeset
|
762 |
(** One auxiliary theorem...**) |
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
paulson
parents:
2215
diff
changeset
|
763 |
|
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
paulson
parents:
2215
diff
changeset
|
764 |
goal HOL.thy "(x = False) = (~ x)"; |
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
paulson
parents:
2215
diff
changeset
|
765 |
by (fast_tac HOL_cs 1); |
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
paulson
parents:
2215
diff
changeset
|
766 |
qed "eq_False_conv"; |
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
paulson
parents:
2215
diff
changeset
|
767 |
|
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
paulson
parents:
2215
diff
changeset
|
768 |
(** Additional theorems for Integ.thy **) |
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
paulson
parents:
2215
diff
changeset
|
769 |
|
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
paulson
parents:
2215
diff
changeset
|
770 |
Addsimps [zless_eq_less, zle_eq_le, |
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
paulson
parents:
2215
diff
changeset
|
771 |
znegative_zminus_znat, not_znegative_znat]; |
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
paulson
parents:
2215
diff
changeset
|
772 |
|
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
paulson
parents:
2215
diff
changeset
|
773 |
goal Integ.thy "!! x. (x::int) = y ==> x <= y"; |
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
paulson
parents:
2215
diff
changeset
|
774 |
by (etac subst 1); by (rtac zle_refl 1); |
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
paulson
parents:
2215
diff
changeset
|
775 |
val zequalD1 = result(); |
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
paulson
parents:
2215
diff
changeset
|
776 |
|
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
paulson
parents:
2215
diff
changeset
|
777 |
goal Integ.thy "($~ x < $~ y) = (y < x)"; |
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
paulson
parents:
2215
diff
changeset
|
778 |
by (rewrite_goals_tac [zless_def,zdiff_def]); |
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
paulson
parents:
2215
diff
changeset
|
779 |
by (simp_tac (!simpset addsimps zadd_ac ) 1); |
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
paulson
parents:
2215
diff
changeset
|
780 |
val zminus_zless_zminus = result(); |
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
paulson
parents:
2215
diff
changeset
|
781 |
|
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
paulson
parents:
2215
diff
changeset
|
782 |
goal Integ.thy "($~ x <= $~ y) = (y <= x)"; |
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
paulson
parents:
2215
diff
changeset
|
783 |
by (simp_tac (HOL_ss addsimps[zle_def, zminus_zless_zminus]) 1); |
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
paulson
parents:
2215
diff
changeset
|
784 |
val zminus_zle_zminus = result(); |
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
paulson
parents:
2215
diff
changeset
|
785 |
|
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
paulson
parents:
2215
diff
changeset
|
786 |
goal Integ.thy "(x < $~ y) = (y < $~ x)"; |
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
paulson
parents:
2215
diff
changeset
|
787 |
by (rewrite_goals_tac [zless_def,zdiff_def]); |
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
paulson
parents:
2215
diff
changeset
|
788 |
by (simp_tac (!simpset addsimps zadd_ac ) 1); |
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
paulson
parents:
2215
diff
changeset
|
789 |
val zless_zminus = result(); |
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
paulson
parents:
2215
diff
changeset
|
790 |
|
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
paulson
parents:
2215
diff
changeset
|
791 |
goal Integ.thy "($~ x < y) = ($~ y < x)"; |
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
paulson
parents:
2215
diff
changeset
|
792 |
by (rewrite_goals_tac [zless_def,zdiff_def]); |
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
paulson
parents:
2215
diff
changeset
|
793 |
by (simp_tac (!simpset addsimps zadd_ac ) 1); |
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
paulson
parents:
2215
diff
changeset
|
794 |
val zminus_zless = result(); |
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
paulson
parents:
2215
diff
changeset
|
795 |
|
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
paulson
parents:
2215
diff
changeset
|
796 |
goal Integ.thy "(x <= $~ y) = (y <= $~ x)"; |
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
paulson
parents:
2215
diff
changeset
|
797 |
by (simp_tac (HOL_ss addsimps[zle_def, zminus_zless]) 1); |
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
paulson
parents:
2215
diff
changeset
|
798 |
val zle_zminus = result(); |
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
paulson
parents:
2215
diff
changeset
|
799 |
|
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
paulson
parents:
2215
diff
changeset
|
800 |
goal Integ.thy "($~ x <= y) = ($~ y <= x)"; |
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
paulson
parents:
2215
diff
changeset
|
801 |
by (simp_tac (HOL_ss addsimps[zle_def, zless_zminus]) 1); |
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
paulson
parents:
2215
diff
changeset
|
802 |
val zminus_zle = result(); |
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
paulson
parents:
2215
diff
changeset
|
803 |
|
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
paulson
parents:
2215
diff
changeset
|
804 |
goal Integ.thy " $#0 < $# Suc n"; |
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
paulson
parents:
2215
diff
changeset
|
805 |
by (rtac (zero_less_Suc RS (zless_eq_less RS iffD2)) 1); |
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
paulson
parents:
2215
diff
changeset
|
806 |
val zero_zless_Suc_pos = result(); |
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
paulson
parents:
2215
diff
changeset
|
807 |
|
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
paulson
parents:
2215
diff
changeset
|
808 |
goal Integ.thy "($# n= $# m) = (n = m)"; |
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
paulson
parents:
2215
diff
changeset
|
809 |
by (fast_tac (HOL_cs addSEs[inj_znat RS injD]) 1); |
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
paulson
parents:
2215
diff
changeset
|
810 |
val znat_znat_eq = result(); |
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
paulson
parents:
2215
diff
changeset
|
811 |
AddIffs[znat_znat_eq]; |
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
paulson
parents:
2215
diff
changeset
|
812 |
|
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
paulson
parents:
2215
diff
changeset
|
813 |
goal Integ.thy "$~ $# Suc n < $#0"; |
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
paulson
parents:
2215
diff
changeset
|
814 |
by (stac (zminus_0 RS sym) 1); |
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
paulson
parents:
2215
diff
changeset
|
815 |
by (rtac (zminus_zless_zminus RS iffD2) 1); |
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
paulson
parents:
2215
diff
changeset
|
816 |
by (rtac (zero_less_Suc RS (zless_eq_less RS iffD2)) 1); |
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
paulson
parents:
2215
diff
changeset
|
817 |
val negative_zless_0 = result(); |
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
paulson
parents:
2215
diff
changeset
|
818 |
Addsimps [zero_zless_Suc_pos, negative_zless_0]; |
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
paulson
parents:
2215
diff
changeset
|
819 |
|
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
paulson
parents:
2215
diff
changeset
|
820 |
goal Integ.thy "$~ $# n <= $#0"; |
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
paulson
parents:
2215
diff
changeset
|
821 |
by (rtac zless_or_eq_imp_zle 1); |
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
paulson
parents:
2215
diff
changeset
|
822 |
by (nat_ind_tac "n" 1); |
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
paulson
parents:
2215
diff
changeset
|
823 |
by (ALLGOALS Asm_simp_tac); |
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
paulson
parents:
2215
diff
changeset
|
824 |
val negative_zle_0 = result(); |
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
paulson
parents:
2215
diff
changeset
|
825 |
Addsimps[negative_zle_0]; |
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
paulson
parents:
2215
diff
changeset
|
826 |
|
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
paulson
parents:
2215
diff
changeset
|
827 |
goal Integ.thy "~($#0 <= $~ $# Suc n)"; |
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
paulson
parents:
2215
diff
changeset
|
828 |
by (stac zle_zminus 1); |
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
paulson
parents:
2215
diff
changeset
|
829 |
by (Simp_tac 1); |
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
paulson
parents:
2215
diff
changeset
|
830 |
val not_zle_0_negative = result(); |
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
paulson
parents:
2215
diff
changeset
|
831 |
Addsimps[not_zle_0_negative]; |
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
paulson
parents:
2215
diff
changeset
|
832 |
|
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
paulson
parents:
2215
diff
changeset
|
833 |
goal Integ.thy "($# n <= $~ $# m) = (n = 0 & m = 0)"; |
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
paulson
parents:
2215
diff
changeset
|
834 |
by (safe_tac HOL_cs); |
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
paulson
parents:
2215
diff
changeset
|
835 |
by (Simp_tac 3); |
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
paulson
parents:
2215
diff
changeset
|
836 |
by (dtac (zle_zminus RS iffD1) 2); |
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
paulson
parents:
2215
diff
changeset
|
837 |
by (ALLGOALS(dtac (negative_zle_0 RSN(2,zle_trans)))); |
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
paulson
parents:
2215
diff
changeset
|
838 |
by (ALLGOALS Asm_full_simp_tac); |
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
paulson
parents:
2215
diff
changeset
|
839 |
val znat_zle_znegative = result(); |
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
paulson
parents:
2215
diff
changeset
|
840 |
|
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
paulson
parents:
2215
diff
changeset
|
841 |
goal Integ.thy "~($# n < $~ $# Suc m)"; |
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
paulson
parents:
2215
diff
changeset
|
842 |
by (rtac notI 1); by (forward_tac [zless_imp_zle] 1); |
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
paulson
parents:
2215
diff
changeset
|
843 |
by (dtac (znat_zle_znegative RS iffD1) 1); |
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
paulson
parents:
2215
diff
changeset
|
844 |
by (safe_tac HOL_cs); |
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
paulson
parents:
2215
diff
changeset
|
845 |
by (dtac (zless_zminus RS iffD1) 1); |
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
paulson
parents:
2215
diff
changeset
|
846 |
by (Asm_full_simp_tac 1); |
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
paulson
parents:
2215
diff
changeset
|
847 |
val not_znat_zless_negative = result(); |
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
paulson
parents:
2215
diff
changeset
|
848 |
|
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
paulson
parents:
2215
diff
changeset
|
849 |
goal Integ.thy "($~ $# n = $# m) = (n = 0 & m = 0)"; |
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
paulson
parents:
2215
diff
changeset
|
850 |
by (rtac iffI 1); |
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
paulson
parents:
2215
diff
changeset
|
851 |
by (rtac (znat_zle_znegative RS iffD1) 1); |
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
paulson
parents:
2215
diff
changeset
|
852 |
by (dtac sym 1); |
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
paulson
parents:
2215
diff
changeset
|
853 |
by (ALLGOALS Asm_simp_tac); |
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
paulson
parents:
2215
diff
changeset
|
854 |
val negative_eq_positive = result(); |
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
paulson
parents:
2215
diff
changeset
|
855 |
|
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
paulson
parents:
2215
diff
changeset
|
856 |
Addsimps [zminus_zless_zminus, zminus_zle_zminus, |
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
paulson
parents:
2215
diff
changeset
|
857 |
negative_eq_positive, not_znat_zless_negative]; |
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
paulson
parents:
2215
diff
changeset
|
858 |
|
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
paulson
parents:
2215
diff
changeset
|
859 |
goalw Integ.thy [zdiff_def,zless_def] "!! x. znegative x = (x < $# 0)"; |
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
paulson
parents:
2215
diff
changeset
|
860 |
by (Auto_tac()); |
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
paulson
parents:
2215
diff
changeset
|
861 |
val znegative_less_0 = result(); |
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
paulson
parents:
2215
diff
changeset
|
862 |
|
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
paulson
parents:
2215
diff
changeset
|
863 |
goalw Integ.thy [zdiff_def,zless_def] "!! x. (~znegative x) = ($# 0 <= x)"; |
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
paulson
parents:
2215
diff
changeset
|
864 |
by (stac znegative_less_0 1); |
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
paulson
parents:
2215
diff
changeset
|
865 |
by (safe_tac (HOL_cs addSDs[zleD,not_zleE,zleI]) ); |
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
paulson
parents:
2215
diff
changeset
|
866 |
val not_znegative_ge_0 = result(); |
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
paulson
parents:
2215
diff
changeset
|
867 |
|
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
paulson
parents:
2215
diff
changeset
|
868 |
goal Integ.thy "!! x. znegative x ==> ? n. x = $~ $# Suc n"; |
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
paulson
parents:
2215
diff
changeset
|
869 |
by (dtac (znegative_less_0 RS iffD1 RS zless_eq_zadd_Suc) 1); |
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
paulson
parents:
2215
diff
changeset
|
870 |
by (etac exE 1); |
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
paulson
parents:
2215
diff
changeset
|
871 |
by (rtac exI 1); |
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
paulson
parents:
2215
diff
changeset
|
872 |
by (dres_inst_tac [("f","(% z. z + $~ $# Suc n )")] arg_cong 1); |
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
paulson
parents:
2215
diff
changeset
|
873 |
by (auto_tac(!claset, !simpset addsimps [zadd_assoc])); |
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
paulson
parents:
2215
diff
changeset
|
874 |
val znegativeD = result(); |
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
paulson
parents:
2215
diff
changeset
|
875 |
|
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
paulson
parents:
2215
diff
changeset
|
876 |
goal Integ.thy "!! x. ~znegative x ==> ? n. x = $# n"; |
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
paulson
parents:
2215
diff
changeset
|
877 |
by (dtac (not_znegative_ge_0 RS iffD1) 1); |
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
paulson
parents:
2215
diff
changeset
|
878 |
by (dtac zle_imp_zless_or_eq 1); |
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
paulson
parents:
2215
diff
changeset
|
879 |
by (etac disjE 1); |
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
paulson
parents:
2215
diff
changeset
|
880 |
by (dtac zless_eq_zadd_Suc 1); |
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
paulson
parents:
2215
diff
changeset
|
881 |
by (Auto_tac()); |
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
paulson
parents:
2215
diff
changeset
|
882 |
val not_znegativeD = result(); |
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
paulson
parents:
2215
diff
changeset
|
883 |
|
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
paulson
parents:
2215
diff
changeset
|
884 |
(* a case theorem distinguishing positive and negative int *) |
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
paulson
parents:
2215
diff
changeset
|
885 |
|
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
paulson
parents:
2215
diff
changeset
|
886 |
val prems = goal Integ.thy |
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
paulson
parents:
2215
diff
changeset
|
887 |
"[|!! n. P ($# n); !! n. P ($~ $# Suc n) |] ==> P z"; |
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
paulson
parents:
2215
diff
changeset
|
888 |
by (cut_inst_tac [("P","znegative z")] excluded_middle 1); |
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
paulson
parents:
2215
diff
changeset
|
889 |
by (fast_tac (HOL_cs addSDs[znegativeD,not_znegativeD] addSIs prems) 1); |
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
paulson
parents:
2215
diff
changeset
|
890 |
val int_cases = result(); |
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
paulson
parents:
2215
diff
changeset
|
891 |
|
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
paulson
parents:
2215
diff
changeset
|
892 |
fun int_case_tac x = res_inst_tac [("z",x)] int_cases; |
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
paulson
parents:
2215
diff
changeset
|
893 |