| author | mueller | 
| Fri, 16 May 1997 16:08:38 +0200 | |
| changeset 3218 | 44f01b718eab | 
| parent 2988 | d38f330e58b3 | 
| child 3919 | c036caebfc75 | 
| permissions | -rw-r--r-- | 
| 2224 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
1894diff
changeset | 1 | (* Title: HOL/Integ/Bin.ML | 
| 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
1894diff
changeset | 2 | Authors: Lawrence C Paulson, Cambridge University Computer Laboratory | 
| 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
1894diff
changeset | 3 | David Spelt, University of Twente | 
| 1632 | 4 | Copyright 1994 University of Cambridge | 
| 5 | Copyright 1996 University of Twente | |
| 6 | ||
| 7 | Arithmetic on binary integers. | |
| 8 | *) | |
| 9 | ||
| 10 | open Bin; | |
| 11 | ||
| 12 | (** extra rules for bin_succ, bin_pred, bin_add, bin_mult **) | |
| 13 | ||
| 2224 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
1894diff
changeset | 14 | qed_goal "norm_Bcons_Plus_0" Bin.thy | 
| 2988 
d38f330e58b3
Renamed sign constructors to eliminate clash with the Plus infix of Sum.thy
 paulson parents: 
2224diff
changeset | 15 | "norm_Bcons PlusSign False = PlusSign" | 
| 2224 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
1894diff
changeset | 16 | (fn _ => [(Simp_tac 1)]); | 
| 1632 | 17 | |
| 2224 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
1894diff
changeset | 18 | qed_goal "norm_Bcons_Plus_1" Bin.thy | 
| 2988 
d38f330e58b3
Renamed sign constructors to eliminate clash with the Plus infix of Sum.thy
 paulson parents: 
2224diff
changeset | 19 | "norm_Bcons PlusSign True = Bcons PlusSign True" | 
| 2224 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
1894diff
changeset | 20 | (fn _ => [(Simp_tac 1)]); | 
| 1632 | 21 | |
| 2224 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
1894diff
changeset | 22 | qed_goal "norm_Bcons_Minus_0" Bin.thy | 
| 2988 
d38f330e58b3
Renamed sign constructors to eliminate clash with the Plus infix of Sum.thy
 paulson parents: 
2224diff
changeset | 23 | "norm_Bcons MinusSign False = Bcons MinusSign False" | 
| 2224 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
1894diff
changeset | 24 | (fn _ => [(Simp_tac 1)]); | 
| 1632 | 25 | |
| 2224 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
1894diff
changeset | 26 | qed_goal "norm_Bcons_Minus_1" Bin.thy | 
| 2988 
d38f330e58b3
Renamed sign constructors to eliminate clash with the Plus infix of Sum.thy
 paulson parents: 
2224diff
changeset | 27 | "norm_Bcons MinusSign True = MinusSign" | 
| 2224 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
1894diff
changeset | 28 | (fn _ => [(Simp_tac 1)]); | 
| 1632 | 29 | |
| 2224 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
1894diff
changeset | 30 | qed_goal "norm_Bcons_Bcons" Bin.thy | 
| 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
1894diff
changeset | 31 | "norm_Bcons (Bcons w x) b = Bcons (Bcons w x) b" | 
| 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
1894diff
changeset | 32 | (fn _ => [(Simp_tac 1)]); | 
| 1632 | 33 | |
| 2224 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
1894diff
changeset | 34 | qed_goal "bin_succ_Bcons1" Bin.thy | 
| 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
1894diff
changeset | 35 | "bin_succ(Bcons w True) = Bcons (bin_succ w) False" | 
| 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
1894diff
changeset | 36 | (fn _ => [(Simp_tac 1)]); | 
| 1632 | 37 | |
| 2224 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
1894diff
changeset | 38 | qed_goal "bin_succ_Bcons0" Bin.thy | 
| 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
1894diff
changeset | 39 | "bin_succ(Bcons w False) = norm_Bcons w True" | 
| 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
1894diff
changeset | 40 | (fn _ => [(Simp_tac 1)]); | 
| 1632 | 41 | |
| 2224 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
1894diff
changeset | 42 | qed_goal "bin_pred_Bcons1" Bin.thy | 
| 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
1894diff
changeset | 43 | "bin_pred(Bcons w True) = norm_Bcons w False" | 
| 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
1894diff
changeset | 44 | (fn _ => [(Simp_tac 1)]); | 
| 1632 | 45 | |
| 2224 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
1894diff
changeset | 46 | qed_goal "bin_pred_Bcons0" Bin.thy | 
| 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
1894diff
changeset | 47 | "bin_pred(Bcons w False) = Bcons (bin_pred w) True" | 
| 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
1894diff
changeset | 48 | (fn _ => [(Simp_tac 1)]); | 
| 1632 | 49 | |
| 2224 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
1894diff
changeset | 50 | qed_goal "bin_minus_Bcons1" Bin.thy | 
| 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
1894diff
changeset | 51 | "bin_minus(Bcons w True) = bin_pred (Bcons(bin_minus w) False)" | 
| 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
1894diff
changeset | 52 | (fn _ => [(Simp_tac 1)]); | 
| 1632 | 53 | |
| 2224 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
1894diff
changeset | 54 | qed_goal "bin_minus_Bcons0" Bin.thy | 
| 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
1894diff
changeset | 55 | "bin_minus(Bcons w False) = Bcons (bin_minus w) False" | 
| 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
1894diff
changeset | 56 | (fn _ => [(Simp_tac 1)]); | 
| 1632 | 57 | |
| 58 | (*** bin_add: binary addition ***) | |
| 59 | ||
| 2224 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
1894diff
changeset | 60 | qed_goal "bin_add_Bcons_Bcons11" Bin.thy | 
| 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
1894diff
changeset | 61 | "bin_add (Bcons v True) (Bcons w True) = \ | 
| 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
1894diff
changeset | 62 | \ norm_Bcons (bin_add v (bin_succ w)) False" | 
| 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
1894diff
changeset | 63 | (fn _ => [(Simp_tac 1)]); | 
| 1632 | 64 | |
| 2224 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
1894diff
changeset | 65 | qed_goal "bin_add_Bcons_Bcons10" Bin.thy | 
| 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
1894diff
changeset | 66 | "bin_add (Bcons v True) (Bcons w False) = norm_Bcons (bin_add v w) True" | 
| 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
1894diff
changeset | 67 | (fn _ => [(Simp_tac 1)]); | 
| 1632 | 68 | |
| 2224 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
1894diff
changeset | 69 | val lemma = prove_goal HOL.thy "(False = (~P)) = P" | 
| 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
1894diff
changeset | 70 | (fn _ => [(Fast_tac 1)]); | 
| 1632 | 71 | |
| 2224 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
1894diff
changeset | 72 | qed_goal "bin_add_Bcons_Bcons0" Bin.thy | 
| 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
1894diff
changeset | 73 | "bin_add (Bcons v False) (Bcons w y) = norm_Bcons (bin_add v w) y" | 
| 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
1894diff
changeset | 74 | (fn _ => [(simp_tac (!simpset addsimps [lemma]) 1)]); | 
| 1632 | 75 | |
| 2224 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
1894diff
changeset | 76 | qed_goal "bin_add_Bcons_Plus" Bin.thy | 
| 2988 
d38f330e58b3
Renamed sign constructors to eliminate clash with the Plus infix of Sum.thy
 paulson parents: 
2224diff
changeset | 77 | "bin_add (Bcons v x) PlusSign = Bcons v x" | 
| 2224 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
1894diff
changeset | 78 | (fn _ => [(Simp_tac 1)]); | 
| 1632 | 79 | |
| 2224 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
1894diff
changeset | 80 | qed_goal "bin_add_Bcons_Minus" Bin.thy | 
| 2988 
d38f330e58b3
Renamed sign constructors to eliminate clash with the Plus infix of Sum.thy
 paulson parents: 
2224diff
changeset | 81 | "bin_add (Bcons v x) MinusSign = bin_pred (Bcons v x)" | 
| 2224 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
1894diff
changeset | 82 | (fn _ => [(Simp_tac 1)]); | 
| 1632 | 83 | |
| 2224 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
1894diff
changeset | 84 | qed_goal "bin_add_Bcons_Bcons" Bin.thy | 
| 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
1894diff
changeset | 85 | "bin_add (Bcons v x) (Bcons w y) = \ | 
| 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
1894diff
changeset | 86 | \ norm_Bcons(bin_add v (if x & y then (bin_succ w) else w)) (x~= y)" | 
| 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
1894diff
changeset | 87 | (fn _ => [(Simp_tac 1)]); | 
| 1632 | 88 | |
| 89 | ||
| 90 | (*** bin_add: binary multiplication ***) | |
| 91 | ||
| 2224 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
1894diff
changeset | 92 | qed_goal "bin_mult_Bcons1" Bin.thy | 
| 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
1894diff
changeset | 93 | "bin_mult (Bcons v True) w = bin_add (norm_Bcons (bin_mult v w) False) w" | 
| 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
1894diff
changeset | 94 | (fn _ => [(Simp_tac 1)]); | 
| 1632 | 95 | |
| 2224 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
1894diff
changeset | 96 | qed_goal "bin_mult_Bcons0" Bin.thy | 
| 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
1894diff
changeset | 97 | "bin_mult (Bcons v False) w = norm_Bcons (bin_mult v w) False" | 
| 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
1894diff
changeset | 98 | (fn _ => [(Simp_tac 1)]); | 
| 1632 | 99 | |
| 100 | ||
| 101 | (**** The carry/borrow functions, bin_succ and bin_pred ****) | |
| 102 | ||
| 2224 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
1894diff
changeset | 103 | val if_ss = !simpset setloop (split_tac [expand_if]) ; | 
| 1632 | 104 | |
| 105 | ||
| 106 | (**** integ_of_bin ****) | |
| 107 | ||
| 108 | ||
| 2224 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
1894diff
changeset | 109 | qed_goal "integ_of_bin_norm_Bcons" Bin.thy | 
| 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
1894diff
changeset | 110 | "integ_of_bin(norm_Bcons w b) = integ_of_bin(Bcons w b)" | 
| 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
1894diff
changeset | 111 | (fn _ =>[(bin.induct_tac "w" 1), | 
| 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
1894diff
changeset | 112 | (ALLGOALS(simp_tac if_ss)) ]); | 
| 1632 | 113 | |
| 2224 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
1894diff
changeset | 114 | qed_goal "integ_of_bin_succ" Bin.thy | 
| 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
1894diff
changeset | 115 | "integ_of_bin(bin_succ w) = $#1 + integ_of_bin w" | 
| 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
1894diff
changeset | 116 | (fn _ =>[(rtac bin.induct 1), | 
| 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
1894diff
changeset | 117 | (ALLGOALS(asm_simp_tac | 
| 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
1894diff
changeset | 118 | (if_ss addsimps (integ_of_bin_norm_Bcons::zadd_ac)))) ]); | 
| 1632 | 119 | |
| 2224 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
1894diff
changeset | 120 | qed_goal "integ_of_bin_pred" Bin.thy | 
| 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
1894diff
changeset | 121 | "integ_of_bin(bin_pred w) = $~ ($#1) + integ_of_bin w" | 
| 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
1894diff
changeset | 122 | (fn _ =>[(rtac bin.induct 1), | 
| 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
1894diff
changeset | 123 | (ALLGOALS(asm_simp_tac | 
| 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
1894diff
changeset | 124 | (if_ss addsimps (integ_of_bin_norm_Bcons::zadd_ac)))) ]); | 
| 1632 | 125 | |
| 2224 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
1894diff
changeset | 126 | qed_goal "integ_of_bin_minus" Bin.thy | 
| 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
1894diff
changeset | 127 | "integ_of_bin(bin_minus w) = $~ (integ_of_bin w)" | 
| 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
1894diff
changeset | 128 | (fn _ =>[(rtac bin.induct 1), | 
| 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
1894diff
changeset | 129 | (Simp_tac 1), | 
| 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
1894diff
changeset | 130 | (Simp_tac 1), | 
| 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
1894diff
changeset | 131 | (asm_simp_tac (if_ss | 
| 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
1894diff
changeset | 132 | delsimps [pred_Plus,pred_Minus,pred_Bcons] | 
| 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
1894diff
changeset | 133 | addsimps [integ_of_bin_succ,integ_of_bin_pred, | 
| 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
1894diff
changeset | 134 | zadd_assoc]) 1)]); | 
| 1632 | 135 | |
| 136 | ||
| 2224 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
1894diff
changeset | 137 | val bin_add_simps = [add_Plus,add_Minus,bin_add_Bcons_Plus, | 
| 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
1894diff
changeset | 138 | bin_add_Bcons_Minus,bin_add_Bcons_Bcons, | 
| 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
1894diff
changeset | 139 | integ_of_bin_succ, integ_of_bin_pred, | 
| 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
1894diff
changeset | 140 | integ_of_bin_norm_Bcons]; | 
| 1632 | 141 | val bin_simps = [iob_Plus,iob_Minus,iob_Bcons]; | 
| 142 | ||
| 2224 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
1894diff
changeset | 143 | goal Bin.thy | 
| 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
1894diff
changeset | 144 | "! w. integ_of_bin(bin_add v w) = integ_of_bin v + integ_of_bin w"; | 
| 1632 | 145 | by (bin.induct_tac "v" 1); | 
| 2224 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
1894diff
changeset | 146 | by (simp_tac (if_ss addsimps bin_add_simps) 1); | 
| 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
1894diff
changeset | 147 | by (simp_tac (if_ss addsimps bin_add_simps) 1); | 
| 1632 | 148 | by (rtac allI 1); | 
| 149 | by (bin.induct_tac "w" 1); | |
| 2224 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
1894diff
changeset | 150 | by (asm_simp_tac (if_ss addsimps (bin_add_simps)) 1); | 
| 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
1894diff
changeset | 151 | by (asm_simp_tac (if_ss addsimps (bin_add_simps @ zadd_ac)) 1); | 
| 1632 | 152 | by (cut_inst_tac [("P","bool")] True_or_False 1);
 | 
| 153 | by (etac disjE 1); | |
| 2224 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
1894diff
changeset | 154 | by (asm_simp_tac (if_ss addsimps (bin_add_simps @ zadd_ac)) 1); | 
| 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
1894diff
changeset | 155 | by (asm_simp_tac (if_ss addsimps (bin_add_simps @ zadd_ac)) 1); | 
| 1632 | 156 | val integ_of_bin_add_lemma = result(); | 
| 157 | ||
| 158 | goal Bin.thy "integ_of_bin(bin_add v w) = integ_of_bin v + integ_of_bin w"; | |
| 159 | by (cut_facts_tac [integ_of_bin_add_lemma] 1); | |
| 1894 | 160 | by (Fast_tac 1); | 
| 1632 | 161 | qed "integ_of_bin_add"; | 
| 162 | ||
| 2224 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
1894diff
changeset | 163 | val bin_mult_simps = [integ_of_bin_minus, integ_of_bin_add, | 
| 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
1894diff
changeset | 164 | integ_of_bin_norm_Bcons]; | 
| 1632 | 165 | |
| 166 | goal Bin.thy "integ_of_bin(bin_mult v w) = integ_of_bin v * integ_of_bin w"; | |
| 167 | by (bin.induct_tac "v" 1); | |
| 2224 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
1894diff
changeset | 168 | by (simp_tac (if_ss addsimps bin_mult_simps) 1); | 
| 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
1894diff
changeset | 169 | by (simp_tac (if_ss addsimps bin_mult_simps) 1); | 
| 1632 | 170 | by (cut_inst_tac [("P","bool")] True_or_False 1);
 | 
| 171 | by (etac disjE 1); | |
| 2224 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
1894diff
changeset | 172 | by (asm_simp_tac (if_ss addsimps (bin_mult_simps @ [zadd_zmult_distrib])) 2); | 
| 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
1894diff
changeset | 173 | by (asm_simp_tac (if_ss addsimps (bin_mult_simps @ [zadd_zmult_distrib] @ | 
| 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
1894diff
changeset | 174 | zadd_ac)) 1); | 
| 1632 | 175 | qed "integ_of_bin_mult"; | 
| 176 | ||
| 177 | ||
| 178 | Delsimps [succ_Bcons,pred_Bcons,min_Bcons,add_Bcons,mult_Bcons, | |
| 2224 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
1894diff
changeset | 179 | iob_Plus,iob_Minus,iob_Bcons, | 
| 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
1894diff
changeset | 180 | norm_Plus,norm_Minus,norm_Bcons]; | 
| 1632 | 181 | |
| 182 | Addsimps [integ_of_bin_add RS sym, | |
| 2224 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
1894diff
changeset | 183 | integ_of_bin_minus RS sym, | 
| 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
1894diff
changeset | 184 | integ_of_bin_mult RS sym, | 
| 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
1894diff
changeset | 185 | bin_succ_Bcons1,bin_succ_Bcons0, | 
| 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
1894diff
changeset | 186 | bin_pred_Bcons1,bin_pred_Bcons0, | 
| 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
1894diff
changeset | 187 | bin_minus_Bcons1,bin_minus_Bcons0, | 
| 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
1894diff
changeset | 188 | bin_add_Bcons_Plus,bin_add_Bcons_Minus, | 
| 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
1894diff
changeset | 189 | bin_add_Bcons_Bcons0,bin_add_Bcons_Bcons10,bin_add_Bcons_Bcons11, | 
| 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
1894diff
changeset | 190 | bin_mult_Bcons1,bin_mult_Bcons0, | 
| 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
1894diff
changeset | 191 | norm_Bcons_Plus_0,norm_Bcons_Plus_1, | 
| 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
1894diff
changeset | 192 | norm_Bcons_Minus_0,norm_Bcons_Minus_1, | 
| 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
1894diff
changeset | 193 | norm_Bcons_Bcons]; | 
| 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
1894diff
changeset | 194 | |
| 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
1894diff
changeset | 195 | |
| 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
1894diff
changeset | 196 | (** Simplification rules for comparison of binary numbers (Norbert Völker) **) | 
| 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
1894diff
changeset | 197 | |
| 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
1894diff
changeset | 198 | Addsimps [zadd_assoc]; | 
| 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
1894diff
changeset | 199 | |
| 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
1894diff
changeset | 200 | goal Bin.thy | 
| 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
1894diff
changeset | 201 | "(integ_of_bin x = integ_of_bin y) \ | 
| 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
1894diff
changeset | 202 | \ = (integ_of_bin (bin_add x (bin_minus y)) = $# 0)"; | 
| 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
1894diff
changeset | 203 | by (simp_tac (HOL_ss addsimps | 
| 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
1894diff
changeset | 204 | [integ_of_bin_add, integ_of_bin_minus,zdiff_def]) 1); | 
| 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
1894diff
changeset | 205 | by (rtac iffI 1); | 
| 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
1894diff
changeset | 206 | by (etac ssubst 1); | 
| 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
1894diff
changeset | 207 | by (rtac zadd_zminus_inverse 1); | 
| 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
1894diff
changeset | 208 |   by (dres_inst_tac [("f","(% z. z + integ_of_bin y)")] arg_cong 1); 
 | 
| 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
1894diff
changeset | 209 | by (asm_full_simp_tac | 
| 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
1894diff
changeset | 210 | (HOL_ss addsimps[zadd_assoc,zadd_0,zadd_0_right, | 
| 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
1894diff
changeset | 211 | zadd_zminus_inverse2]) 1); | 
| 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
1894diff
changeset | 212 | val iob_eq_eq_diff_0 = result(); | 
| 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
1894diff
changeset | 213 | |
| 2988 
d38f330e58b3
Renamed sign constructors to eliminate clash with the Plus infix of Sum.thy
 paulson parents: 
2224diff
changeset | 214 | goal Bin.thy "(integ_of_bin PlusSign = $# 0) = True"; | 
| 2224 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
1894diff
changeset | 215 | by (stac iob_Plus 1); by (Simp_tac 1); | 
| 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
1894diff
changeset | 216 | val iob_Plus_eq_0 = result(); | 
| 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
1894diff
changeset | 217 | |
| 2988 
d38f330e58b3
Renamed sign constructors to eliminate clash with the Plus infix of Sum.thy
 paulson parents: 
2224diff
changeset | 218 | goal Bin.thy "(integ_of_bin MinusSign = $# 0) = False"; | 
| 2224 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
1894diff
changeset | 219 | by (stac iob_Minus 1); | 
| 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
1894diff
changeset | 220 | by (Simp_tac 1); | 
| 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
1894diff
changeset | 221 | val iob_Minus_not_eq_0 = result(); | 
| 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
1894diff
changeset | 222 | |
| 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
1894diff
changeset | 223 | goal Bin.thy | 
| 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
1894diff
changeset | 224 | "(integ_of_bin (Bcons w x) = $# 0) = (~x & integ_of_bin w = $# 0)"; | 
| 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
1894diff
changeset | 225 | by (stac iob_Bcons 1); | 
| 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
1894diff
changeset | 226 | by (case_tac "x" 1); | 
| 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
1894diff
changeset | 227 | by (ALLGOALS Asm_simp_tac); | 
| 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
1894diff
changeset | 228 | by (ALLGOALS(asm_simp_tac (HOL_ss addsimps [integ_of_bin_add]))); | 
| 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
1894diff
changeset | 229 | by (ALLGOALS(int_case_tac "integ_of_bin w")); | 
| 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
1894diff
changeset | 230 | by (ALLGOALS(asm_simp_tac | 
| 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
1894diff
changeset | 231 | (!simpset addsimps[zminus_zadd_distrib RS sym, | 
| 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
1894diff
changeset | 232 | znat_add RS sym]))); | 
| 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
1894diff
changeset | 233 | by (stac eq_False_conv 1); | 
| 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
1894diff
changeset | 234 | by (rtac notI 1); | 
| 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
1894diff
changeset | 235 |   by (dres_inst_tac [("f","(% z. z +  $# Suc (Suc (n + n)))")] arg_cong 1); 
 | 
| 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
1894diff
changeset | 236 | by (Asm_full_simp_tac 1); | 
| 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
1894diff
changeset | 237 | val iob_Bcons_eq_0 = result(); | 
| 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
1894diff
changeset | 238 | |
| 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
1894diff
changeset | 239 | goalw Bin.thy [zless_def,zdiff_def] | 
| 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
1894diff
changeset | 240 | "integ_of_bin x < integ_of_bin y \ | 
| 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
1894diff
changeset | 241 | \ = (integ_of_bin (bin_add x (bin_minus y)) < $# 0)"; | 
| 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
1894diff
changeset | 242 | by (Simp_tac 1); | 
| 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
1894diff
changeset | 243 | val iob_less_eq_diff_lt_0 = result(); | 
| 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
1894diff
changeset | 244 | |
| 2988 
d38f330e58b3
Renamed sign constructors to eliminate clash with the Plus infix of Sum.thy
 paulson parents: 
2224diff
changeset | 245 | goal Bin.thy "(integ_of_bin PlusSign < $# 0) = False"; | 
| 2224 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
1894diff
changeset | 246 | by (stac iob_Plus 1); by (Simp_tac 1); | 
| 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
1894diff
changeset | 247 | val iob_Plus_not_lt_0 = result(); | 
| 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
1894diff
changeset | 248 | |
| 2988 
d38f330e58b3
Renamed sign constructors to eliminate clash with the Plus infix of Sum.thy
 paulson parents: 
2224diff
changeset | 249 | goal Bin.thy "(integ_of_bin MinusSign < $# 0) = True"; | 
| 2224 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
1894diff
changeset | 250 | by (stac iob_Minus 1); by (Simp_tac 1); | 
| 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
1894diff
changeset | 251 | val iob_Minus_lt_0 = result(); | 
| 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
1894diff
changeset | 252 | |
| 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
1894diff
changeset | 253 | goal Bin.thy | 
| 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
1894diff
changeset | 254 | "(integ_of_bin (Bcons w x) < $# 0) = (integ_of_bin w < $# 0)"; | 
| 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
1894diff
changeset | 255 | by (stac iob_Bcons 1); | 
| 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
1894diff
changeset | 256 | by (case_tac "x" 1); | 
| 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
1894diff
changeset | 257 | by (ALLGOALS Asm_simp_tac); | 
| 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
1894diff
changeset | 258 | by (ALLGOALS(asm_simp_tac (HOL_ss addsimps [integ_of_bin_add]))); | 
| 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
1894diff
changeset | 259 | by (ALLGOALS(int_case_tac "integ_of_bin w")); | 
| 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
1894diff
changeset | 260 | by (ALLGOALS(asm_simp_tac | 
| 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
1894diff
changeset | 261 | (!simpset addsimps[zminus_zadd_distrib RS sym, | 
| 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
1894diff
changeset | 262 | znat_add RS sym]))); | 
| 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
1894diff
changeset | 263 | by (stac (zadd_zminus_inverse RS sym) 1); | 
| 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
1894diff
changeset | 264 | by (rtac zadd_zless_mono1 1); | 
| 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
1894diff
changeset | 265 | by (Simp_tac 1); | 
| 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
1894diff
changeset | 266 | val iob_Bcons_lt_0 = result(); | 
| 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
1894diff
changeset | 267 | |
| 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
1894diff
changeset | 268 | goal Bin.thy | 
| 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
1894diff
changeset | 269 | "integ_of_bin x <= integ_of_bin y \ | 
| 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
1894diff
changeset | 270 | \ = ( integ_of_bin (bin_add x (bin_minus y)) < $# 0 \ | 
| 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
1894diff
changeset | 271 | \ | integ_of_bin (bin_add x (bin_minus y)) = $# 0)"; | 
| 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
1894diff
changeset | 272 | by (simp_tac (HOL_ss addsimps | 
| 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
1894diff
changeset | 273 | [zle_eq_zless_or_eq,iob_less_eq_diff_lt_0,zdiff_def | 
| 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
1894diff
changeset | 274 | ,iob_eq_eq_diff_0,integ_of_bin_minus,integ_of_bin_add]) 1); | 
| 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
1894diff
changeset | 275 | val iob_le_diff_lt_0_or_eq_0 = result(); | 
| 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
1894diff
changeset | 276 | |
| 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
1894diff
changeset | 277 | Delsimps [zless_eq_less, zle_eq_le, zadd_assoc, znegative_zminus_znat, | 
| 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
1894diff
changeset | 278 | not_znegative_znat, zero_zless_Suc_pos, negative_zless_0, | 
| 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
1894diff
changeset | 279 | negative_zle_0, not_zle_0_negative, not_znat_zless_negative, | 
| 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
1894diff
changeset | 280 | zminus_zless_zminus, zminus_zle_zminus, negative_eq_positive]; | 
| 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
1894diff
changeset | 281 | |
| 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
1894diff
changeset | 282 | Addsimps [zdiff_def, iob_eq_eq_diff_0, iob_less_eq_diff_lt_0, | 
| 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
1894diff
changeset | 283 | iob_le_diff_lt_0_or_eq_0, iob_Plus_eq_0, iob_Minus_not_eq_0, | 
| 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
1894diff
changeset | 284 | iob_Bcons_eq_0, iob_Plus_not_lt_0, iob_Minus_lt_0, iob_Bcons_lt_0]; | 
| 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
1894diff
changeset | 285 | |
| 1632 | 286 | |
| 287 | (*** Examples of performing binary arithmetic by simplification ***) | |
| 288 | ||
| 289 | goal Bin.thy "#13 + #19 = #32"; | |
| 290 | by (Simp_tac 1); | |
| 291 | result(); | |
| 292 | ||
| 293 | goal Bin.thy "#1234 + #5678 = #6912"; | |
| 294 | by (Simp_tac 1); | |
| 295 | result(); | |
| 296 | ||
| 297 | goal Bin.thy "#1359 + #~2468 = #~1109"; | |
| 298 | by (Simp_tac 1); | |
| 299 | result(); | |
| 300 | ||
| 301 | goal Bin.thy "#93746 + #~46375 = #47371"; | |
| 302 | by (Simp_tac 1); | |
| 303 | result(); | |
| 304 | ||
| 305 | goal Bin.thy "$~ #65745 = #~65745"; | |
| 306 | by (Simp_tac 1); | |
| 307 | result(); | |
| 308 | ||
| 309 | goal Bin.thy "$~ #~54321 = #54321"; | |
| 310 | by (Simp_tac 1); | |
| 311 | result(); | |
| 312 | ||
| 313 | goal Bin.thy "#13 * #19 = #247"; | |
| 314 | by (Simp_tac 1); | |
| 315 | result(); | |
| 316 | ||
| 317 | goal Bin.thy "#~84 * #51 = #~4284"; | |
| 318 | by (Simp_tac 1); | |
| 319 | result(); | |
| 320 | ||
| 321 | goal Bin.thy "#255 * #255 = #65025"; | |
| 322 | by (Simp_tac 1); | |
| 323 | result(); | |
| 324 | ||
| 325 | goal Bin.thy "#1359 * #~2468 = #~3354012"; | |
| 326 | by (Simp_tac 1); | |
| 327 | result(); | |
| 2224 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
1894diff
changeset | 328 | |
| 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
1894diff
changeset | 329 | goal Bin.thy "#89 * #10 ~= #889"; | 
| 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
1894diff
changeset | 330 | by (Simp_tac 1); | 
| 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
1894diff
changeset | 331 | result(); | 
| 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
1894diff
changeset | 332 | |
| 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
1894diff
changeset | 333 | goal Bin.thy "#13 < #18 - #4"; | 
| 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
1894diff
changeset | 334 | by (Simp_tac 1); | 
| 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
1894diff
changeset | 335 | result(); | 
| 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
1894diff
changeset | 336 | |
| 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
1894diff
changeset | 337 | goal Bin.thy "#~345 < #~242 + #~100"; | 
| 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
1894diff
changeset | 338 | by (Simp_tac 1); | 
| 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
1894diff
changeset | 339 | result(); | 
| 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
1894diff
changeset | 340 | |
| 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
1894diff
changeset | 341 | goal Bin.thy "#13557456 < #18678654"; | 
| 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
1894diff
changeset | 342 | by (Simp_tac 1); | 
| 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
1894diff
changeset | 343 | result(); | 
| 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
1894diff
changeset | 344 | |
| 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
1894diff
changeset | 345 | goal Bin.thy "#999999 <= (#1000001 + #1)-#2"; | 
| 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
1894diff
changeset | 346 | by (Simp_tac 1); | 
| 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
1894diff
changeset | 347 | result(); | 
| 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
1894diff
changeset | 348 | |
| 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
1894diff
changeset | 349 | goal Bin.thy "#1234567 <= #1234567"; | 
| 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
1894diff
changeset | 350 | by (Simp_tac 1); | 
| 
4fc4b465be5b
New material from Norbert Voelker for efficient binary comparisons
 paulson parents: 
1894diff
changeset | 351 | result(); |