| author | paulson | 
| Tue, 28 Jul 1998 16:30:56 +0200 | |
| changeset 5204 | 858da18069d7 | 
| parent 5184 | 9b8547a9496a | 
| child 5278 | a903b66822e2 | 
| 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"; | 
| 4473 | 25 | by (res_inst_tac [("k1","x2")] (add_left_cancel RS iffD1) 1);
 | 
| 925 
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: 
925diff
changeset | 37 | \ ((x1,y1),(x2,y2)): intrel"; | 
| 4089 | 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*) | 
| 5069 | 42 | Goalw [intrel_def] | 
| 925 
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: 
925diff
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: 
925diff
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 | |
| 5069 | 59 | Goal "((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 | |
| 5069 | 63 | Goal "(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 | |
| 5069 | 67 | Goalw [equiv_def, refl_def, sym_def, trans_def] | 
| 925 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 68 |     "equiv {x::(nat*nat).True} intrel";
 | 
| 4089 | 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 | |
| 5069 | 78 | Goalw  [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 | |
| 5069 | 82 | Goal "inj_on Abs_Integ Integ"; | 
| 4831 | 83 | by (rtac inj_on_inverseI 1); | 
| 925 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 84 | by (etac Abs_Integ_inverse 1); | 
| 4831 | 85 | qed "inj_on_Abs_Integ"; | 
| 925 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 86 | |
| 4831 | 87 | Addsimps [equiv_intrel_iff, inj_on_Abs_Integ RS inj_on_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 | |
| 5069 | 90 | Goal "inj(Rep_Integ)"; | 
| 925 
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 | |
| 5069 | 100 | Goal "inj(znat)"; | 
| 925 
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); | 
| 4831 | 103 | by (dtac (inj_on_Abs_Integ RS inj_onD) 1); | 
| 925 
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); | 
| 4162 | 108 | by Safe_tac; | 
| 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 | |
| 5069 | 115 | Goalw [congruent_def] | 
| 972 
e61b058d58d2
changed syntax of tuples from <..., ...> to (..., ...)
 clasohm parents: 
925diff
changeset | 116 |   "congruent intrel (%p. split (%x y. intrel^^{(y,x)}) p)";
 | 
| 4162 | 117 | by Safe_tac; | 
| 4089 | 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 | |
| 5069 | 125 | Goalw [zminus_def] | 
| 972 
e61b058d58d2
changed syntax of tuples from <..., ...> to (..., ...)
 clasohm parents: 
925diff
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);
 | 
| 4089 | 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: 
925diff
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); | 
| 4089 | 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 | |
| 5069 | 143 | Goal "$~ ($~ z) = z"; | 
| 925 
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);
 | 
| 4089 | 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 | |
| 5069 | 148 | Goal "inj(zminus)"; | 
| 925 
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);
 | 
| 4089 | 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 | |
| 5069 | 154 | Goalw [znat_def] "$~ ($#0) = $#0"; | 
| 4089 | 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 | |
| 5069 | 162 | Goalw [znegative_def, znat_def] | 
| 925 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 163 | "~ znegative($# n)"; | 
| 1266 | 164 | by (Simp_tac 1); | 
| 4162 | 165 | by Safe_tac; | 
| 925 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 166 | qed "not_znegative_znat"; | 
| 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 167 | |
| 5069 | 168 | Goalw [znegative_def, znat_def] "znegative($~ $# Suc(n))"; | 
| 4089 | 169 | by (simp_tac (simpset() addsimps [zminus]) 1); | 
| 925 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 170 | qed "znegative_zminus_znat"; | 
| 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 171 | |
| 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 172 | |
| 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 173 | (**** zmagnitude: magnitide of an integer, as a natural number ****) | 
| 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 174 | |
| 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 175 | goal Arith.thy "!!n::nat. n - Suc(n+m)=0"; | 
| 5184 | 176 | by (induct_tac "n" 1); | 
| 1266 | 177 | by (ALLGOALS Asm_simp_tac); | 
| 925 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 178 | qed "diff_Suc_add_0"; | 
| 
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 | goal Arith.thy "Suc((n::nat)+m)-n=Suc(m)"; | 
| 5184 | 181 | by (induct_tac "n" 1); | 
| 1266 | 182 | by (ALLGOALS Asm_simp_tac); | 
| 925 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 183 | qed "diff_Suc_add_inverse"; | 
| 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 184 | |
| 5069 | 185 | Goalw [congruent_def] | 
| 972 
e61b058d58d2
changed syntax of tuples from <..., ...> to (..., ...)
 clasohm parents: 
925diff
changeset | 186 |     "congruent intrel (split (%x y. intrel^^{((y-x) + (x-(y::nat)),0)}))";
 | 
| 4162 | 187 | by Safe_tac; | 
| 1266 | 188 | by (Asm_simp_tac 1); | 
| 925 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 189 | by (etac rev_mp 1); | 
| 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 190 | by (res_inst_tac [("m","x1"),("n","y1")] diff_induct 1);
 | 
| 4089 | 191 | by (asm_simp_tac (simpset() addsimps [inj_Suc RS inj_eq]) 3); | 
| 192 | by (asm_simp_tac (simpset() addsimps [diff_add_inverse,diff_add_0]) 2); | |
| 1266 | 193 | by (Asm_simp_tac 1); | 
| 925 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 194 | by (rtac impI 1); | 
| 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 195 | by (etac subst 1); | 
| 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 196 | by (res_inst_tac [("m1","x")] (add_commute RS ssubst) 1);
 | 
| 4089 | 197 | 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 | 198 | qed "zmagnitude_congruent"; | 
| 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 199 | |
| 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 200 | (*Resolve th against the corresponding facts for zmagnitude*) | 
| 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 201 | val zmagnitude_ize = RSLIST [equiv_intrel, zmagnitude_congruent]; | 
| 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 202 | |
| 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 203 | |
| 5069 | 204 | Goalw [zmagnitude_def] | 
| 972 
e61b058d58d2
changed syntax of tuples from <..., ...> to (..., ...)
 clasohm parents: 
925diff
changeset | 205 |     "zmagnitude (Abs_Integ(intrel^^{(x,y)})) = \
 | 
| 
e61b058d58d2
changed syntax of tuples from <..., ...> to (..., ...)
 clasohm parents: 
925diff
changeset | 206 | \    Abs_Integ(intrel^^{((y - x) + (x - y),0)})";
 | 
| 925 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 207 | by (res_inst_tac [("f","Abs_Integ")] arg_cong 1);
 | 
| 4089 | 208 | 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 | 209 | qed "zmagnitude"; | 
| 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 210 | |
| 5069 | 211 | Goalw [znat_def] "zmagnitude($# n) = $#n"; | 
| 4089 | 212 | by (asm_simp_tac (simpset() addsimps [zmagnitude]) 1); | 
| 925 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 213 | qed "zmagnitude_znat"; | 
| 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 214 | |
| 5069 | 215 | Goalw [znat_def] "zmagnitude($~ $# n) = $#n"; | 
| 4089 | 216 | by (asm_simp_tac (simpset() addsimps [zmagnitude, zminus]) 1); | 
| 925 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 217 | qed "zmagnitude_zminus_znat"; | 
| 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 218 | |
| 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 219 | |
| 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 220 | (**** zadd: addition on Integ ****) | 
| 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 221 | |
| 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 222 | (** Congruence property for addition **) | 
| 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 223 | |
| 5069 | 224 | Goalw [congruent2_def] | 
| 925 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 225 | "congruent2 intrel (%p1 p2. \ | 
| 972 
e61b058d58d2
changed syntax of tuples from <..., ...> to (..., ...)
 clasohm parents: 
925diff
changeset | 226 | \         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 | 227 | (*Proof via congruent2_commuteI seems longer*) | 
| 4162 | 228 | by Safe_tac; | 
| 4089 | 229 | by (asm_simp_tac (simpset() addsimps [add_assoc]) 1); | 
| 925 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 230 | (*The rest should be trivial, but rearranging terms is hard*) | 
| 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 231 | by (res_inst_tac [("x1","x1a")] (add_left_commute RS ssubst) 1);
 | 
| 4089 | 232 | by (asm_simp_tac (simpset() addsimps [add_assoc RS sym]) 1); | 
| 925 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 233 | qed "zadd_congruent2"; | 
| 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 234 | |
| 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 235 | (*Resolve th against the corresponding facts for zadd*) | 
| 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 236 | val zadd_ize = RSLIST [equiv_intrel, zadd_congruent2]; | 
| 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 237 | |
| 5069 | 238 | Goalw [zadd_def] | 
| 972 
e61b058d58d2
changed syntax of tuples from <..., ...> to (..., ...)
 clasohm parents: 
925diff
changeset | 239 |   "Abs_Integ(intrel^^{(x1,y1)}) + Abs_Integ(intrel^^{(x2,y2)}) = \
 | 
| 
e61b058d58d2
changed syntax of tuples from <..., ...> to (..., ...)
 clasohm parents: 
925diff
changeset | 240 | \  Abs_Integ(intrel^^{(x1+x2, y1+y2)})";
 | 
| 925 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 241 | by (asm_simp_tac | 
| 4089 | 242 | (simpset() addsimps [zadd_ize UN_equiv_class2]) 1); | 
| 925 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 243 | qed "zadd"; | 
| 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 244 | |
| 5069 | 245 | Goalw [znat_def] "$#0 + z = z"; | 
| 925 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 246 | by (res_inst_tac [("z","z")] eq_Abs_Integ 1);
 | 
| 4089 | 247 | by (asm_simp_tac (simpset() addsimps [zadd]) 1); | 
| 925 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 248 | qed "zadd_0"; | 
| 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 249 | |
| 5069 | 250 | Goal "$~ (z + w) = $~ z + $~ w"; | 
| 925 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 251 | by (res_inst_tac [("z","z")] eq_Abs_Integ 1);
 | 
| 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 252 | by (res_inst_tac [("z","w")] eq_Abs_Integ 1);
 | 
| 4089 | 253 | by (asm_simp_tac (simpset() addsimps [zminus,zadd]) 1); | 
| 925 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 254 | qed "zminus_zadd_distrib"; | 
| 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 255 | |
| 5069 | 256 | Goal "(z::int) + w = w + z"; | 
| 925 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 257 | by (res_inst_tac [("z","z")] eq_Abs_Integ 1);
 | 
| 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 258 | by (res_inst_tac [("z","w")] eq_Abs_Integ 1);
 | 
| 4089 | 259 | 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 | 260 | qed "zadd_commute"; | 
| 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 261 | |
| 5069 | 262 | Goal "((z1::int) + z2) + z3 = z1 + (z2 + z3)"; | 
| 925 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 263 | by (res_inst_tac [("z","z1")] eq_Abs_Integ 1);
 | 
| 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 264 | by (res_inst_tac [("z","z2")] eq_Abs_Integ 1);
 | 
| 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 265 | by (res_inst_tac [("z","z3")] eq_Abs_Integ 1);
 | 
| 4089 | 266 | 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 | 267 | qed "zadd_assoc"; | 
| 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 268 | |
| 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 269 | (*For AC rewriting*) | 
| 5069 | 270 | Goal "(x::int)+(y+z)=y+(x+z)"; | 
| 925 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 271 | by (rtac (zadd_commute RS trans) 1); | 
| 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 272 | by (rtac (zadd_assoc RS trans) 1); | 
| 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 273 | by (rtac (zadd_commute RS arg_cong) 1); | 
| 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 274 | qed "zadd_left_commute"; | 
| 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 275 | |
| 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 276 | (*Integer addition is an AC operator*) | 
| 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 277 | val zadd_ac = [zadd_assoc,zadd_commute,zadd_left_commute]; | 
| 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 278 | |
| 5069 | 279 | Goalw [znat_def] "$# (m + n) = ($#m) + ($#n)"; | 
| 4089 | 280 | by (asm_simp_tac (simpset() addsimps [zadd]) 1); | 
| 925 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 281 | qed "znat_add"; | 
| 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 282 | |
| 5069 | 283 | Goalw [znat_def] "z + ($~ z) = $#0"; | 
| 925 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 284 | by (res_inst_tac [("z","z")] eq_Abs_Integ 1);
 | 
| 4089 | 285 | 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 | 286 | qed "zadd_zminus_inverse"; | 
| 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 287 | |
| 5069 | 288 | Goal "($~ z) + z = $#0"; | 
| 925 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 289 | by (rtac (zadd_commute RS trans) 1); | 
| 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 290 | by (rtac zadd_zminus_inverse 1); | 
| 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 291 | qed "zadd_zminus_inverse2"; | 
| 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 292 | |
| 5069 | 293 | Goal "z + $#0 = z"; | 
| 925 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 294 | by (rtac (zadd_commute RS trans) 1); | 
| 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 295 | by (rtac zadd_0 1); | 
| 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 296 | qed "zadd_0_right"; | 
| 
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 | |
| 2224 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
2215diff
changeset | 299 | (** Lemmas **) | 
| 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
2215diff
changeset | 300 | |
| 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
2215diff
changeset | 301 | qed_goal "zadd_assoc_cong" Integ.thy | 
| 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
2215diff
changeset | 302 | "!!z. (z::int) + v = z' + v' ==> z + (v + w) = z' + (v' + w)" | 
| 4089 | 303 | (fn _ => [(asm_simp_tac (simpset() addsimps [zadd_assoc RS sym]) 1)]); | 
| 2224 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
2215diff
changeset | 304 | |
| 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
2215diff
changeset | 305 | 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: 
2215diff
changeset | 306 | (fn _ => [(REPEAT (ares_tac [zadd_commute RS zadd_assoc_cong] 1))]); | 
| 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
2215diff
changeset | 307 | |
| 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
2215diff
changeset | 308 | |
| 925 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 309 | (*Need properties of subtraction? Or use $- just as an abbreviation!*) | 
| 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 310 | |
| 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 311 | (**** zmult: multiplication on Integ ****) | 
| 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 312 | |
| 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 313 | (** Congruence property for multiplication **) | 
| 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 314 | |
| 5069 | 315 | Goal "((k::nat) + l) + (m + n) = (k + m) + (n + l)"; | 
| 4089 | 316 | by (simp_tac (simpset() addsimps add_ac) 1); | 
| 925 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 317 | qed "zmult_congruent_lemma"; | 
| 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 318 | |
| 5069 | 319 | Goal | 
| 1465 | 320 | "congruent2 intrel (%p1 p2. \ | 
| 321 | \ split (%x1 y1. split (%x2 y2. \ | |
| 972 
e61b058d58d2
changed syntax of tuples from <..., ...> to (..., ...)
 clasohm parents: 
925diff
changeset | 322 | \                   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 | 323 | by (rtac (equiv_intrel RS congruent2_commuteI) 1); | 
| 4817 | 324 | by (pair_tac "w" 2); | 
| 325 | by (rename_tac "z1 z2" 2); | |
| 4772 
8c7e7eaffbdf
split_all_tac now fails if there is nothing to split
 oheimb parents: 
4676diff
changeset | 326 | by Safe_tac; | 
| 
8c7e7eaffbdf
split_all_tac now fails if there is nothing to split
 oheimb parents: 
4676diff
changeset | 327 | by (rewtac split_def); | 
| 
8c7e7eaffbdf
split_all_tac now fails if there is nothing to split
 oheimb parents: 
4676diff
changeset | 328 | by (simp_tac (simpset() addsimps add_ac@mult_ac) 1); | 
| 4089 | 329 | by (asm_simp_tac (simpset() delsimps [equiv_intrel_iff] | 
| 1266 | 330 | addsimps add_ac@mult_ac) 1); | 
| 925 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 331 | 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 | 332 | by (rtac (zmult_congruent_lemma RS trans) 1); | 
| 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 333 | by (rtac (zmult_congruent_lemma RS trans RS sym) 1); | 
| 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 334 | by (rtac (zmult_congruent_lemma RS trans RS sym) 1); | 
| 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 335 | by (rtac (zmult_congruent_lemma RS trans RS sym) 1); | 
| 4089 | 336 | by (asm_simp_tac (simpset() addsimps [add_mult_distrib RS sym]) 1); | 
| 337 | 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 | 338 | qed "zmult_congruent2"; | 
| 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 339 | |
| 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 340 | (*Resolve th against the corresponding facts for zmult*) | 
| 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 341 | val zmult_ize = RSLIST [equiv_intrel, zmult_congruent2]; | 
| 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 342 | |
| 5069 | 343 | Goalw [zmult_def] | 
| 1465 | 344 |    "Abs_Integ((intrel^^{(x1,y1)})) * Abs_Integ((intrel^^{(x2,y2)})) =   \
 | 
| 972 
e61b058d58d2
changed syntax of tuples from <..., ...> to (..., ...)
 clasohm parents: 
925diff
changeset | 345 | \   Abs_Integ(intrel ^^ {(x1*x2 + y1*y2, x1*y2 + y1*x2)})";
 | 
| 4089 | 346 | 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 | 347 | qed "zmult"; | 
| 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 348 | |
| 5069 | 349 | Goalw [znat_def] "$#0 * z = $#0"; | 
| 925 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 350 | by (res_inst_tac [("z","z")] eq_Abs_Integ 1);
 | 
| 4089 | 351 | by (asm_simp_tac (simpset() addsimps [zmult]) 1); | 
| 925 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 352 | qed "zmult_0"; | 
| 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 353 | |
| 5069 | 354 | Goalw [znat_def] "$#Suc(0) * z = z"; | 
| 925 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 355 | by (res_inst_tac [("z","z")] eq_Abs_Integ 1);
 | 
| 4089 | 356 | by (asm_simp_tac (simpset() addsimps [zmult]) 1); | 
| 925 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 357 | qed "zmult_1"; | 
| 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 358 | |
| 5069 | 359 | Goal "($~ z) * w = $~ (z * w)"; | 
| 925 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 360 | by (res_inst_tac [("z","z")] eq_Abs_Integ 1);
 | 
| 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 361 | by (res_inst_tac [("z","w")] eq_Abs_Integ 1);
 | 
| 4089 | 362 | 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 | 363 | qed "zmult_zminus"; | 
| 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 364 | |
| 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 365 | |
| 5069 | 366 | Goal "($~ z) * ($~ w) = (z * w)"; | 
| 925 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 367 | by (res_inst_tac [("z","z")] eq_Abs_Integ 1);
 | 
| 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 368 | by (res_inst_tac [("z","w")] eq_Abs_Integ 1);
 | 
| 4089 | 369 | 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 | 370 | qed "zmult_zminus_zminus"; | 
| 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 371 | |
| 5069 | 372 | Goal "(z::int) * w = w * z"; | 
| 925 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 373 | by (res_inst_tac [("z","z")] eq_Abs_Integ 1);
 | 
| 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 374 | by (res_inst_tac [("z","w")] eq_Abs_Integ 1);
 | 
| 4089 | 375 | 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 | 376 | qed "zmult_commute"; | 
| 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 377 | |
| 5069 | 378 | Goal "z * $# 0 = $#0"; | 
| 925 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 379 | by (rtac ([zmult_commute, zmult_0] MRS trans) 1); | 
| 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 380 | qed "zmult_0_right"; | 
| 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 381 | |
| 5069 | 382 | Goal "z * $#Suc(0) = z"; | 
| 925 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 383 | by (rtac ([zmult_commute, zmult_1] MRS trans) 1); | 
| 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 384 | qed "zmult_1_right"; | 
| 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 385 | |
| 5069 | 386 | Goal "((z1::int) * z2) * z3 = z1 * (z2 * z3)"; | 
| 925 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 387 | by (res_inst_tac [("z","z1")] eq_Abs_Integ 1);
 | 
| 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 388 | by (res_inst_tac [("z","z2")] eq_Abs_Integ 1);
 | 
| 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 389 | by (res_inst_tac [("z","z3")] eq_Abs_Integ 1);
 | 
| 4089 | 390 | by (asm_simp_tac (simpset() addsimps ([add_mult_distrib2,zmult] @ | 
| 2036 | 391 | add_ac @ mult_ac)) 1); | 
| 925 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 392 | qed "zmult_assoc"; | 
| 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 393 | |
| 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 394 | (*For AC rewriting*) | 
| 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 395 | qed_goal "zmult_left_commute" Integ.thy | 
| 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 396 | "(z1::int)*(z2*z3) = z2*(z1*z3)" | 
| 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 397 | (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 | 398 | rtac (zmult_commute RS arg_cong) 1]); | 
| 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 399 | |
| 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 400 | (*Integer multiplication is an AC operator*) | 
| 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 401 | val zmult_ac = [zmult_assoc, zmult_commute, zmult_left_commute]; | 
| 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 402 | |
| 5069 | 403 | Goal "((z1::int) + z2) * w = (z1 * w) + (z2 * w)"; | 
| 925 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 404 | by (res_inst_tac [("z","z1")] eq_Abs_Integ 1);
 | 
| 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 405 | by (res_inst_tac [("z","z2")] eq_Abs_Integ 1);
 | 
| 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 406 | by (res_inst_tac [("z","w")] eq_Abs_Integ 1);
 | 
| 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 407 | by (asm_simp_tac | 
| 4089 | 408 | (simpset() addsimps ([add_mult_distrib2, zadd, zmult] @ | 
| 2036 | 409 | add_ac @ mult_ac)) 1); | 
| 925 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 410 | qed "zadd_zmult_distrib"; | 
| 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 411 | |
| 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 412 | val zmult_commute'= read_instantiate [("z","w")] zmult_commute;
 | 
| 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 413 | |
| 5069 | 414 | Goal "w * ($~ z) = $~ (w * z)"; | 
| 4089 | 415 | 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 | 416 | qed "zmult_zminus_right"; | 
| 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 417 | |
| 5069 | 418 | Goal "(w::int) * (z1 + z2) = (w * z1) + (w * z2)"; | 
| 4089 | 419 | 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 | 420 | qed "zadd_zmult_distrib2"; | 
| 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 421 | |
| 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 422 | val zadd_simps = | 
| 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 423 | [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 | 424 | |
| 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 425 | val zminus_simps = [zminus_zminus, zminus_0, zminus_zadd_distrib]; | 
| 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 426 | |
| 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 427 | val zmult_simps = [zmult_0, zmult_1, zmult_0_right, zmult_1_right, | 
| 1465 | 428 | zmult_zminus, zmult_zminus_right]; | 
| 925 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 429 | |
| 1266 | 430 | Addsimps (zadd_simps @ zminus_simps @ zmult_simps @ | 
| 431 | [zmagnitude_znat, zmagnitude_zminus_znat]); | |
| 925 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 432 | |
| 
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 | (**** Additional Theorems (by Mattolini; proofs mainly by lcp) ****) | 
| 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 435 | |
| 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 436 | (* Some Theorems about zsuc and zpred *) | 
| 5069 | 437 | Goalw [zsuc_def] "$#(Suc(n)) = zsuc($# n)"; | 
| 4089 | 438 | 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 | 439 | qed "znat_Suc"; | 
| 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 440 | |
| 5069 | 441 | Goalw [zpred_def,zsuc_def,zdiff_def] "$~ zsuc(z) = zpred($~ z)"; | 
| 1266 | 442 | by (Simp_tac 1); | 
| 925 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 443 | qed "zminus_zsuc"; | 
| 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 444 | |
| 5069 | 445 | Goalw [zpred_def,zsuc_def,zdiff_def] "$~ zpred(z) = zsuc($~ z)"; | 
| 1266 | 446 | by (Simp_tac 1); | 
| 925 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 447 | qed "zminus_zpred"; | 
| 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 448 | |
| 5069 | 449 | Goalw [zsuc_def,zpred_def,zdiff_def] | 
| 925 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 450 | "zpred(zsuc(z)) = z"; | 
| 4089 | 451 | by (simp_tac (simpset() addsimps [zadd_assoc]) 1); | 
| 925 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 452 | qed "zpred_zsuc"; | 
| 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 453 | |
| 5069 | 454 | Goalw [zsuc_def,zpred_def,zdiff_def] | 
| 925 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 455 | "zsuc(zpred(z)) = z"; | 
| 4089 | 456 | by (simp_tac (simpset() addsimps [zadd_assoc]) 1); | 
| 925 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 457 | qed "zsuc_zpred"; | 
| 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 458 | |
| 5069 | 459 | Goal "(zpred(z)=w) = (z=zsuc(w))"; | 
| 4162 | 460 | by Safe_tac; | 
| 925 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 461 | by (rtac (zsuc_zpred RS sym) 1); | 
| 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 462 | by (rtac zpred_zsuc 1); | 
| 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 463 | qed "zpred_to_zsuc"; | 
| 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 464 | |
| 5069 | 465 | Goal "(zsuc(z)=w)=(z=zpred(w))"; | 
| 4162 | 466 | by Safe_tac; | 
| 925 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 467 | by (rtac (zpred_zsuc RS sym) 1); | 
| 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 468 | by (rtac zsuc_zpred 1); | 
| 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 469 | qed "zsuc_to_zpred"; | 
| 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 470 | |
| 5069 | 471 | Goal "($~ z = w) = (z = $~ w)"; | 
| 4162 | 472 | by Safe_tac; | 
| 925 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 473 | by (rtac (zminus_zminus RS sym) 1); | 
| 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 474 | by (rtac zminus_zminus 1); | 
| 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 475 | qed "zminus_exchange"; | 
| 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 476 | |
| 5069 | 477 | Goal"(zsuc(z)=zsuc(w)) = (z=w)"; | 
| 4162 | 478 | by Safe_tac; | 
| 925 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 479 | by (dres_inst_tac [("f","zpred")] arg_cong 1);
 | 
| 4089 | 480 | 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 | 481 | qed "bijective_zsuc"; | 
| 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 482 | |
| 5069 | 483 | Goal"(zpred(z)=zpred(w)) = (z=w)"; | 
| 4162 | 484 | by Safe_tac; | 
| 925 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 485 | by (dres_inst_tac [("f","zsuc")] arg_cong 1);
 | 
| 4089 | 486 | 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 | 487 | qed "bijective_zpred"; | 
| 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 488 | |
| 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 489 | (* Additional Theorems about zadd *) | 
| 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 490 | |
| 5069 | 491 | Goalw [zsuc_def] "zsuc(z) + w = zsuc(z+w)"; | 
| 4089 | 492 | by (simp_tac (simpset() addsimps zadd_ac) 1); | 
| 925 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 493 | qed "zadd_zsuc"; | 
| 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 494 | |
| 5069 | 495 | Goalw [zsuc_def] "w + zsuc(z) = zsuc(w+z)"; | 
| 4089 | 496 | by (simp_tac (simpset() addsimps zadd_ac) 1); | 
| 925 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 497 | qed "zadd_zsuc_right"; | 
| 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 498 | |
| 5069 | 499 | Goalw [zpred_def,zdiff_def] "zpred(z) + w = zpred(z+w)"; | 
| 4089 | 500 | by (simp_tac (simpset() addsimps zadd_ac) 1); | 
| 925 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 501 | qed "zadd_zpred"; | 
| 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 502 | |
| 5069 | 503 | Goalw [zpred_def,zdiff_def] "w + zpred(z) = zpred(w+z)"; | 
| 4089 | 504 | by (simp_tac (simpset() addsimps zadd_ac) 1); | 
| 925 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 505 | qed "zadd_zpred_right"; | 
| 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 506 | |
| 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 507 | |
| 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 508 | (* Additional Theorems about zmult *) | 
| 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 509 | |
| 5069 | 510 | Goalw [zsuc_def] "zsuc(w) * z = z + w * z"; | 
| 4089 | 511 | 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 | 512 | qed "zmult_zsuc"; | 
| 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 513 | |
| 5069 | 514 | Goalw [zsuc_def] "z * zsuc(w) = z + w * z"; | 
| 925 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 515 | by (simp_tac | 
| 4089 | 516 | (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 | 517 | qed "zmult_zsuc_right"; | 
| 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 518 | |
| 5069 | 519 | Goalw [zpred_def, zdiff_def] "zpred(w) * z = w * z - z"; | 
| 4089 | 520 | by (simp_tac (simpset() addsimps [zadd_zmult_distrib]) 1); | 
| 925 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 521 | qed "zmult_zpred"; | 
| 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 522 | |
| 5069 | 523 | Goalw [zpred_def, zdiff_def] "z * zpred(w) = w * z - z"; | 
| 4089 | 524 | 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 | 525 | qed "zmult_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 | (* Further Theorems about zsuc and zpred *) | 
| 5069 | 528 | Goal "$#Suc(m) ~= $#0"; | 
| 4089 | 529 | 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 | 530 | qed "znat_Suc_not_znat_Zero"; | 
| 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 531 | |
| 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 532 | 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 | 533 | |
| 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 534 | |
| 5069 | 535 | Goalw [zsuc_def,znat_def] "w ~= zsuc(w)"; | 
| 925 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 536 | by (res_inst_tac [("z","w")] eq_Abs_Integ 1);
 | 
| 4089 | 537 | by (asm_full_simp_tac (simpset() addsimps [zadd]) 1); | 
| 925 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 538 | qed "n_not_zsuc_n"; | 
| 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 539 | |
| 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 540 | 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 | 541 | |
| 5069 | 542 | Goal "w ~= zpred(w)"; | 
| 4162 | 543 | by Safe_tac; | 
| 925 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 544 | by (dres_inst_tac [("x","w"),("f","zsuc")] arg_cong 1);
 | 
| 4089 | 545 | 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 | 546 | qed "n_not_zpred_n"; | 
| 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 547 | |
| 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 548 | 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 | 549 | |
| 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 550 | |
| 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 551 | (* Theorems about less and less_equal *) | 
| 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 552 | |
| 5069 | 553 | Goalw [zless_def, znegative_def, zdiff_def, znat_def] | 
| 5148 
74919e8f221c
More tidying and removal of "\!\!... from Goal commands
 paulson parents: 
5143diff
changeset | 554 | "w<z ==> ? n. z = w + $#(Suc(n))"; | 
| 925 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 555 | by (res_inst_tac [("z","z")] eq_Abs_Integ 1);
 | 
| 
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);
 | 
| 4162 | 557 | by Safe_tac; | 
| 4089 | 558 | by (asm_full_simp_tac (simpset() addsimps [zadd, zminus]) 1); | 
| 559 | by (safe_tac (claset() addSDs [less_eq_Suc_add])); | |
| 925 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 560 | by (res_inst_tac [("x","k")] exI 1);
 | 
| 4676 
335dfafb1816
Better simplification allows deletion of parts of proofs
 paulson parents: 
4477diff
changeset | 561 | 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 | 562 | qed "zless_eq_zadd_Suc"; | 
| 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 563 | |
| 5069 | 564 | Goalw [zless_def, znegative_def, zdiff_def, znat_def] | 
| 925 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 565 | "z < z + $#(Suc(n))"; | 
| 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 566 | by (res_inst_tac [("z","z")] eq_Abs_Integ 1);
 | 
| 4676 
335dfafb1816
Better simplification allows deletion of parts of proofs
 paulson parents: 
4477diff
changeset | 567 | by (Clarify_tac 1); | 
| 4089 | 568 | by (simp_tac (simpset() addsimps [zadd, zminus]) 1); | 
| 925 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 569 | qed "zless_zadd_Suc"; | 
| 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 570 | |
| 5143 
b94cd208f073
Removal of leading "\!\!..." from most Goal commands
 paulson parents: 
5069diff
changeset | 571 | Goal "[| z1<z2; z2<z3 |] ==> z1 < (z3::int)"; | 
| 4089 | 572 | by (safe_tac (claset() addSDs [zless_eq_zadd_Suc])); | 
| 925 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 573 | by (simp_tac | 
| 4089 | 574 | (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 | 575 | qed "zless_trans"; | 
| 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 576 | |
| 5069 | 577 | Goalw [zsuc_def] "z<zsuc(z)"; | 
| 925 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 578 | by (rtac zless_zadd_Suc 1); | 
| 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 579 | qed "zlessI"; | 
| 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 580 | |
| 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 581 | val zless_zsucI = zlessI RSN (2,zless_trans); | 
| 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 582 | |
| 5069 | 583 | Goal "!!z w::int. z<w ==> ~w<z"; | 
| 4089 | 584 | by (safe_tac (claset() addSDs [zless_eq_zadd_Suc])); | 
| 925 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 585 | by (res_inst_tac [("z","z")] eq_Abs_Integ 1);
 | 
| 4162 | 586 | by Safe_tac; | 
| 4089 | 587 | 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 | 588 | qed "zless_not_sym"; | 
| 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 589 | |
| 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 590 | (* [| n<m; m<n |] ==> R *) | 
| 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 591 | bind_thm ("zless_asym", (zless_not_sym RS notE));
 | 
| 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 592 | |
| 5069 | 593 | Goal "!!z::int. ~ z<z"; | 
| 925 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 594 | by (resolve_tac [zless_asym RS notI] 1); | 
| 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 595 | by (REPEAT (assume_tac 1)); | 
| 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 596 | qed "zless_not_refl"; | 
| 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 597 | |
| 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 598 | (* z<z ==> R *) | 
| 1619 | 599 | bind_thm ("zless_irrefl", (zless_not_refl RS notE));
 | 
| 925 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 600 | |
| 5143 
b94cd208f073
Removal of leading "\!\!..." from most Goal commands
 paulson parents: 
5069diff
changeset | 601 | Goal "z<w ==> w ~= (z::int)"; | 
| 4089 | 602 | by (fast_tac (claset() addEs [zless_irrefl]) 1); | 
| 925 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 603 | qed "zless_not_refl2"; | 
| 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 604 | |
| 
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 | (*"Less than" is a linear ordering*) | 
| 5069 | 607 | Goalw [zless_def, znegative_def, zdiff_def] | 
| 925 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 608 | "z<w | z=w | w<(z::int)"; | 
| 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 609 | by (res_inst_tac [("z","z")] eq_Abs_Integ 1);
 | 
| 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 610 | by (res_inst_tac [("z","w")] eq_Abs_Integ 1);
 | 
| 4162 | 611 | by Safe_tac; | 
| 925 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 612 | by (asm_full_simp_tac | 
| 4089 | 613 | (simpset() addsimps [zadd, zminus, Image_iff, Bex_def]) 1); | 
| 925 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 614 | by (res_inst_tac [("m1", "x+ya"), ("n1", "xa+y")] (less_linear RS disjE) 1);
 | 
| 4089 | 615 | 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 | 616 | qed "zless_linear"; | 
| 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 617 | |
| 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 618 | |
| 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 619 | (*** Properties of <= ***) | 
| 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 620 | |
| 5069 | 621 | Goalw [zless_def, znegative_def, zdiff_def, znat_def] | 
| 925 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 622 | "($#m < $#n) = (m<n)"; | 
| 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 623 | by (simp_tac | 
| 4089 | 624 | (simpset() addsimps [zadd, zminus, Image_iff, Bex_def]) 1); | 
| 625 | 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 | 626 | qed "zless_eq_less"; | 
| 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 627 | |
| 5069 | 628 | Goalw [zle_def, le_def] "($#m <= $#n) = (m<=n)"; | 
| 4089 | 629 | by (simp_tac (simpset() addsimps [zless_eq_less]) 1); | 
| 925 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 630 | qed "zle_eq_le"; | 
| 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 631 | |
| 5143 
b94cd208f073
Removal of leading "\!\!..." from most Goal commands
 paulson parents: 
5069diff
changeset | 632 | Goalw [zle_def] "~(w<z) ==> z<=(w::int)"; | 
| 925 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 633 | by (assume_tac 1); | 
| 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 634 | qed "zleI"; | 
| 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 635 | |
| 5143 
b94cd208f073
Removal of leading "\!\!..." from most Goal commands
 paulson parents: 
5069diff
changeset | 636 | Goalw [zle_def] "z<=w ==> ~(w<(z::int))"; | 
| 925 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 637 | by (assume_tac 1); | 
| 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 638 | qed "zleD"; | 
| 
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 | val zleE = make_elim zleD; | 
| 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 641 | |
| 5143 
b94cd208f073
Removal of leading "\!\!..." from most Goal commands
 paulson parents: 
5069diff
changeset | 642 | Goalw [zle_def] "~ z <= w ==> w<(z::int)"; | 
| 1894 | 643 | by (Fast_tac 1); | 
| 925 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 644 | qed "not_zleE"; | 
| 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 645 | |
| 5143 
b94cd208f073
Removal of leading "\!\!..." from most Goal commands
 paulson parents: 
5069diff
changeset | 646 | Goalw [zle_def] "z < w ==> z <= (w::int)"; | 
| 4089 | 647 | by (fast_tac (claset() addEs [zless_asym]) 1); | 
| 925 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 648 | qed "zless_imp_zle"; | 
| 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 649 | |
| 5143 
b94cd208f073
Removal of leading "\!\!..." from most Goal commands
 paulson parents: 
5069diff
changeset | 650 | Goalw [zle_def] "z <= w ==> z < w | z=(w::int)"; | 
| 925 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 651 | by (cut_facts_tac [zless_linear] 1); | 
| 4089 | 652 | 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 | 653 | qed "zle_imp_zless_or_eq"; | 
| 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 654 | |
| 5143 
b94cd208f073
Removal of leading "\!\!..." from most Goal commands
 paulson parents: 
5069diff
changeset | 655 | Goalw [zle_def] "z<w | z=w ==> z <=(w::int)"; | 
| 925 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 656 | by (cut_facts_tac [zless_linear] 1); | 
| 4089 | 657 | 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 | 658 | qed "zless_or_eq_imp_zle"; | 
| 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 659 | |
| 5069 | 660 | Goal "(x <= (y::int)) = (x < y | x=y)"; | 
| 925 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 661 | 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 | 662 | qed "zle_eq_zless_or_eq"; | 
| 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 663 | |
| 5069 | 664 | Goal "w <= (w::int)"; | 
| 4089 | 665 | 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 | 666 | qed "zle_refl"; | 
| 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 667 | |
| 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 668 | 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 | 669 | by (dtac zle_imp_zless_or_eq 1); | 
| 4089 | 670 | by (fast_tac (claset() addIs [zless_trans]) 1); | 
| 925 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 671 | qed "zle_zless_trans"; | 
| 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 672 | |
| 5143 
b94cd208f073
Removal of leading "\!\!..." from most Goal commands
 paulson parents: 
5069diff
changeset | 673 | Goal "[| i <= j; j <= k |] ==> i <= (k::int)"; | 
| 925 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 674 | by (EVERY1 [dtac zle_imp_zless_or_eq, dtac zle_imp_zless_or_eq, | 
| 4089 | 675 | 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 | 676 | qed "zle_trans"; | 
| 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 677 | |
| 5143 
b94cd208f073
Removal of leading "\!\!..." from most Goal commands
 paulson parents: 
5069diff
changeset | 678 | Goal "[| z <= w; w <= z |] ==> z = (w::int)"; | 
| 925 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 679 | by (EVERY1 [dtac zle_imp_zless_or_eq, dtac zle_imp_zless_or_eq, | 
| 4089 | 680 | fast_tac (claset() addEs [zless_irrefl,zless_asym])]); | 
| 925 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 681 | qed "zle_anti_sym"; | 
| 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 682 | |
| 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 683 | |
| 5069 | 684 | Goal "!!w w' z::int. z + w' = z + w ==> w' = w"; | 
| 925 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 685 | by (dres_inst_tac [("f", "%x. x + $~z")] arg_cong 1);
 | 
| 4089 | 686 | 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 | 687 | qed "zadd_left_cancel"; | 
| 
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 | |
| 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 690 | (*** Monotonicity results ***) | 
| 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 691 | |
| 5069 | 692 | Goal "!!v w z::int. v < w ==> v + z < w + z"; | 
| 4089 | 693 | by (safe_tac (claset() addSDs [zless_eq_zadd_Suc])); | 
| 694 | by (simp_tac (simpset() addsimps zadd_ac) 1); | |
| 695 | 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 | 696 | qed "zadd_zless_mono1"; | 
| 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 697 | |
| 5069 | 698 | Goal "!!v w z::int. (v+z < w+z) = (v < w)"; | 
| 4089 | 699 | by (safe_tac (claset() addSEs [zadd_zless_mono1])); | 
| 925 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 700 | by (dres_inst_tac [("z", "$~z")] zadd_zless_mono1 1);
 | 
| 4089 | 701 | 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 | 702 | qed "zadd_left_cancel_zless"; | 
| 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 703 | |
| 5069 | 704 | Goal "!!v w z::int. (v+z <= w+z) = (v <= w)"; | 
| 925 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 705 | by (asm_full_simp_tac | 
| 4089 | 706 | (simpset() addsimps [zle_def, zadd_left_cancel_zless]) 1); | 
| 925 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 707 | qed "zadd_left_cancel_zle"; | 
| 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 708 | |
| 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 709 | (*"v<=w ==> v+z <= w+z"*) | 
| 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 710 | 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 | 711 | |
| 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 712 | |
| 5069 | 713 | Goal "!!z' z::int. [| w'<=w; z'<=z |] ==> w' + z' <= w + z"; | 
| 925 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 714 | by (etac (zadd_zle_mono1 RS zle_trans) 1); | 
| 4089 | 715 | by (simp_tac (simpset() addsimps [zadd_commute]) 1); | 
| 925 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 716 | (*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 | 717 | by (etac zadd_zle_mono1 1); | 
| 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 718 | qed "zadd_zle_mono"; | 
| 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 719 | |
| 5069 | 720 | Goal "!!w z::int. z<=$#0 ==> w+z <= w"; | 
| 925 
15539deb6863
new version of HOL/Integ with curried function application
 clasohm parents: diff
changeset | 721 | by (dres_inst_tac [("z", "w")] zadd_zle_mono1 1);
 | 
| 4089 | 722 | 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 | 723 | qed "zadd_zle_self"; | 
| 2224 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
2215diff
changeset | 724 | |
| 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
2215diff
changeset | 725 | |
| 4195 | 726 | (**** Comparisons: lemmas and proofs by Norbert Voelker ****) | 
| 2224 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
2215diff
changeset | 727 | |
| 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
2215diff
changeset | 728 | (** One auxiliary theorem...**) | 
| 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
2215diff
changeset | 729 | |
| 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
2215diff
changeset | 730 | goal HOL.thy "(x = False) = (~ x)"; | 
| 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
2215diff
changeset | 731 | by (fast_tac HOL_cs 1); | 
| 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
2215diff
changeset | 732 | qed "eq_False_conv"; | 
| 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
2215diff
changeset | 733 | |
| 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
2215diff
changeset | 734 | (** Additional theorems for Integ.thy **) | 
| 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
2215diff
changeset | 735 | |
| 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
2215diff
changeset | 736 | Addsimps [zless_eq_less, zle_eq_le, | 
| 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
2215diff
changeset | 737 | znegative_zminus_znat, not_znegative_znat]; | 
| 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
2215diff
changeset | 738 | |
| 5143 
b94cd208f073
Removal of leading "\!\!..." from most Goal commands
 paulson parents: 
5069diff
changeset | 739 | Goal "(x::int) = y ==> x <= y"; | 
| 2224 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
2215diff
changeset | 740 | by (etac subst 1); by (rtac zle_refl 1); | 
| 3725 | 741 | qed "zequalD1"; | 
| 2224 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
2215diff
changeset | 742 | |
| 5069 | 743 | Goal "($~ x < $~ y) = (y < x)"; | 
| 2224 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
2215diff
changeset | 744 | by (rewrite_goals_tac [zless_def,zdiff_def]); | 
| 4089 | 745 | by (simp_tac (simpset() addsimps zadd_ac ) 1); | 
| 3725 | 746 | qed "zminus_zless_zminus"; | 
| 2224 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
2215diff
changeset | 747 | |
| 5069 | 748 | Goal "($~ x <= $~ y) = (y <= x)"; | 
| 2224 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
2215diff
changeset | 749 | by (simp_tac (HOL_ss addsimps[zle_def, zminus_zless_zminus]) 1); | 
| 3725 | 750 | qed "zminus_zle_zminus"; | 
| 2224 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
2215diff
changeset | 751 | |
| 5069 | 752 | Goal "(x < $~ y) = (y < $~ x)"; | 
| 2224 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
2215diff
changeset | 753 | by (rewrite_goals_tac [zless_def,zdiff_def]); | 
| 4089 | 754 | by (simp_tac (simpset() addsimps zadd_ac ) 1); | 
| 3725 | 755 | qed "zless_zminus"; | 
| 2224 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
2215diff
changeset | 756 | |
| 5069 | 757 | Goal "($~ x < y) = ($~ y < x)"; | 
| 2224 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
2215diff
changeset | 758 | by (rewrite_goals_tac [zless_def,zdiff_def]); | 
| 4089 | 759 | by (simp_tac (simpset() addsimps zadd_ac ) 1); | 
| 3725 | 760 | qed "zminus_zless"; | 
| 2224 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
2215diff
changeset | 761 | |
| 5069 | 762 | Goal "(x <= $~ y) = (y <= $~ x)"; | 
| 2224 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
2215diff
changeset | 763 | by (simp_tac (HOL_ss addsimps[zle_def, zminus_zless]) 1); | 
| 3725 | 764 | qed "zle_zminus"; | 
| 2224 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
2215diff
changeset | 765 | |
| 5069 | 766 | Goal "($~ x <= y) = ($~ y <= x)"; | 
| 2224 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
2215diff
changeset | 767 | by (simp_tac (HOL_ss addsimps[zle_def, zless_zminus]) 1); | 
| 3725 | 768 | qed "zminus_zle"; | 
| 2224 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
2215diff
changeset | 769 | |
| 5069 | 770 | Goal " $#0 < $# Suc n"; | 
| 2224 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
2215diff
changeset | 771 | by (rtac (zero_less_Suc RS (zless_eq_less RS iffD2)) 1); | 
| 3725 | 772 | qed "zero_zless_Suc_pos"; | 
| 2224 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
2215diff
changeset | 773 | |
| 5069 | 774 | Goal "($# n= $# m) = (n = m)"; | 
| 2224 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
2215diff
changeset | 775 | by (fast_tac (HOL_cs addSEs[inj_znat RS injD]) 1); | 
| 3725 | 776 | qed "znat_znat_eq"; | 
| 2224 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
2215diff
changeset | 777 | AddIffs[znat_znat_eq]; | 
| 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
2215diff
changeset | 778 | |
| 5069 | 779 | Goal "$~ $# Suc n < $#0"; | 
| 2224 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
2215diff
changeset | 780 | by (stac (zminus_0 RS sym) 1); | 
| 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
2215diff
changeset | 781 | by (rtac (zminus_zless_zminus RS iffD2) 1); | 
| 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
2215diff
changeset | 782 | by (rtac (zero_less_Suc RS (zless_eq_less RS iffD2)) 1); | 
| 3725 | 783 | qed "negative_zless_0"; | 
| 2224 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
2215diff
changeset | 784 | Addsimps [zero_zless_Suc_pos, negative_zless_0]; | 
| 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
2215diff
changeset | 785 | |
| 5069 | 786 | Goal "$~ $# n <= $#0"; | 
| 2224 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
2215diff
changeset | 787 | by (rtac zless_or_eq_imp_zle 1); | 
| 5184 | 788 | by (induct_tac "n" 1); | 
| 2224 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
2215diff
changeset | 789 | by (ALLGOALS Asm_simp_tac); | 
| 3725 | 790 | qed "negative_zle_0"; | 
| 2224 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
2215diff
changeset | 791 | Addsimps[negative_zle_0]; | 
| 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
2215diff
changeset | 792 | |
| 5069 | 793 | Goal "~($#0 <= $~ $# Suc n)"; | 
| 2224 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
2215diff
changeset | 794 | by (stac zle_zminus 1); | 
| 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
2215diff
changeset | 795 | by (Simp_tac 1); | 
| 3725 | 796 | qed "not_zle_0_negative"; | 
| 2224 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
2215diff
changeset | 797 | Addsimps[not_zle_0_negative]; | 
| 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
2215diff
changeset | 798 | |
| 5069 | 799 | Goal "($# n <= $~ $# m) = (n = 0 & m = 0)"; | 
| 2224 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
2215diff
changeset | 800 | by (safe_tac HOL_cs); | 
| 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
2215diff
changeset | 801 | by (Simp_tac 3); | 
| 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
2215diff
changeset | 802 | by (dtac (zle_zminus RS iffD1) 2); | 
| 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
2215diff
changeset | 803 | by (ALLGOALS(dtac (negative_zle_0 RSN(2,zle_trans)))); | 
| 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
2215diff
changeset | 804 | by (ALLGOALS Asm_full_simp_tac); | 
| 3725 | 805 | qed "znat_zle_znegative"; | 
| 2224 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
2215diff
changeset | 806 | |
| 5069 | 807 | Goal "~($# n < $~ $# Suc m)"; | 
| 2224 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
2215diff
changeset | 808 | by (rtac notI 1); by (forward_tac [zless_imp_zle] 1); | 
| 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
2215diff
changeset | 809 | by (dtac (znat_zle_znegative RS iffD1) 1); | 
| 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
2215diff
changeset | 810 | by (safe_tac HOL_cs); | 
| 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
2215diff
changeset | 811 | by (dtac (zless_zminus RS iffD1) 1); | 
| 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
2215diff
changeset | 812 | by (Asm_full_simp_tac 1); | 
| 3725 | 813 | qed "not_znat_zless_negative"; | 
| 2224 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
2215diff
changeset | 814 | |
| 5069 | 815 | Goal "($~ $# n = $# m) = (n = 0 & m = 0)"; | 
| 2224 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
2215diff
changeset | 816 | by (rtac iffI 1); | 
| 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
2215diff
changeset | 817 | by (rtac (znat_zle_znegative RS iffD1) 1); | 
| 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
2215diff
changeset | 818 | by (dtac sym 1); | 
| 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
2215diff
changeset | 819 | by (ALLGOALS Asm_simp_tac); | 
| 3725 | 820 | qed "negative_eq_positive"; | 
| 2224 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
2215diff
changeset | 821 | |
| 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
2215diff
changeset | 822 | Addsimps [zminus_zless_zminus, zminus_zle_zminus, | 
| 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
2215diff
changeset | 823 | negative_eq_positive, not_znat_zless_negative]; | 
| 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
2215diff
changeset | 824 | |
| 5143 
b94cd208f073
Removal of leading "\!\!..." from most Goal commands
 paulson parents: 
5069diff
changeset | 825 | Goalw [zdiff_def,zless_def] "znegative x = (x < $# 0)"; | 
| 4477 
b3e5857d8d99
New Auto_tac (by Oheimb), and new syntax (without parens), and expandshort
 paulson parents: 
4473diff
changeset | 826 | by Auto_tac; | 
| 3725 | 827 | qed "znegative_less_0"; | 
| 2224 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
2215diff
changeset | 828 | |
| 5143 
b94cd208f073
Removal of leading "\!\!..." from most Goal commands
 paulson parents: 
5069diff
changeset | 829 | Goalw [zdiff_def,zless_def] "(~znegative x) = ($# 0 <= x)"; | 
| 2224 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
2215diff
changeset | 830 | by (stac znegative_less_0 1); | 
| 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
2215diff
changeset | 831 | by (safe_tac (HOL_cs addSDs[zleD,not_zleE,zleI]) ); | 
| 3725 | 832 | qed "not_znegative_ge_0"; | 
| 2224 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
2215diff
changeset | 833 | |
| 5143 
b94cd208f073
Removal of leading "\!\!..." from most Goal commands
 paulson parents: 
5069diff
changeset | 834 | Goal "znegative x ==> ? n. x = $~ $# Suc n"; | 
| 2224 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
2215diff
changeset | 835 | 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: 
2215diff
changeset | 836 | by (etac exE 1); | 
| 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
2215diff
changeset | 837 | by (rtac exI 1); | 
| 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
2215diff
changeset | 838 |   by (dres_inst_tac [("f","(% z. z + $~ $# Suc n )")] arg_cong 1); 
 | 
| 4089 | 839 | by (auto_tac(claset(), simpset() addsimps [zadd_assoc])); | 
| 3725 | 840 | qed "znegativeD"; | 
| 2224 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
2215diff
changeset | 841 | |
| 5143 
b94cd208f073
Removal of leading "\!\!..." from most Goal commands
 paulson parents: 
5069diff
changeset | 842 | Goal "~znegative x ==> ? n. x = $# n"; | 
| 2224 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
2215diff
changeset | 843 | by (dtac (not_znegative_ge_0 RS iffD1) 1); | 
| 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
2215diff
changeset | 844 | by (dtac zle_imp_zless_or_eq 1); | 
| 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
2215diff
changeset | 845 | by (etac disjE 1); | 
| 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
2215diff
changeset | 846 | by (dtac zless_eq_zadd_Suc 1); | 
| 4477 
b3e5857d8d99
New Auto_tac (by Oheimb), and new syntax (without parens), and expandshort
 paulson parents: 
4473diff
changeset | 847 | by Auto_tac; | 
| 3725 | 848 | qed "not_znegativeD"; | 
| 2224 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
2215diff
changeset | 849 | |
| 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
2215diff
changeset | 850 | (* a case theorem distinguishing positive and negative int *) | 
| 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
2215diff
changeset | 851 | |
| 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
2215diff
changeset | 852 | val prems = goal Integ.thy | 
| 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
2215diff
changeset | 853 | "[|!! n. P ($# n); !! n. P ($~ $# Suc n) |] ==> P z"; | 
| 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
2215diff
changeset | 854 |   by (cut_inst_tac [("P","znegative z")] excluded_middle 1); 
 | 
| 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
2215diff
changeset | 855 | by (fast_tac (HOL_cs addSDs[znegativeD,not_znegativeD] addSIs prems) 1); | 
| 3725 | 856 | qed "int_cases"; | 
| 2224 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
2215diff
changeset | 857 | |
| 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
2215diff
changeset | 858 | fun int_case_tac x = res_inst_tac [("z",x)] int_cases; 
 | 
| 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
2215diff
changeset | 859 |