| author | wenzelm | 
| Fri, 15 Feb 2002 20:43:44 +0100 | |
| changeset 12896 | 4518acda6d93 | 
| parent 12018 | ec054019c910 | 
| child 13462 | 56610e2ba220 | 
| permissions | -rw-r--r-- | 
| 10751 | 1 | (* Title: HOL/Hyperreal/HyperRealArith0.ML | 
| 2 | ID: $Id$ | |
| 3 | Author: Lawrence C Paulson, Cambridge University Computer Laboratory | |
| 4 | Copyright 1999 University of Cambridge | |
| 5 | ||
| 6 | Assorted facts that need binary literals and the arithmetic decision procedure | |
| 7 | ||
| 8 | Also, common factor cancellation | |
| 9 | *) | |
| 10 | ||
| 12018 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 paulson parents: 
11704diff
changeset | 11 | Goal "x - - y = x + (y::hypreal)"; | 
| 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 paulson parents: 
11704diff
changeset | 12 | by (Simp_tac 1); | 
| 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 paulson parents: 
11704diff
changeset | 13 | qed "hypreal_diff_minus_eq"; | 
| 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 paulson parents: 
11704diff
changeset | 14 | Addsimps [hypreal_diff_minus_eq]; | 
| 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 paulson parents: 
11704diff
changeset | 15 | |
| 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 paulson parents: 
11704diff
changeset | 16 | Goal "((x * y = 0) = (x = 0 | y = (0::hypreal)))"; | 
| 10751 | 17 | by Auto_tac; | 
| 18 | by (cut_inst_tac [("x","x"),("y","y")] hypreal_mult_zero_disj 1);
 | |
| 19 | by Auto_tac; | |
| 20 | qed "hypreal_mult_is_0"; | |
| 21 | AddIffs [hypreal_mult_is_0]; | |
| 22 | ||
| 23 | (** Division and inverse **) | |
| 24 | ||
| 12018 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 paulson parents: 
11704diff
changeset | 25 | Goal "0/x = (0::hypreal)"; | 
| 10751 | 26 | by (simp_tac (simpset() addsimps [hypreal_divide_def]) 1); | 
| 27 | qed "hypreal_0_divide"; | |
| 28 | Addsimps [hypreal_0_divide]; | |
| 29 | ||
| 12018 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 paulson parents: 
11704diff
changeset | 30 | Goal "((0::hypreal) < inverse x) = (0 < x)"; | 
| 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 paulson parents: 
11704diff
changeset | 31 | by (case_tac "x=0" 1); | 
| 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 paulson parents: 
11704diff
changeset | 32 | by (asm_simp_tac (HOL_ss addsimps [HYPREAL_INVERSE_ZERO]) 1); | 
| 10751 | 33 | by (auto_tac (claset() addDs [hypreal_inverse_less_0], | 
| 34 | simpset() addsimps [linorder_neq_iff, | |
| 35 | hypreal_inverse_gt_0])); | |
| 36 | qed "hypreal_0_less_inverse_iff"; | |
| 37 | Addsimps [hypreal_0_less_inverse_iff]; | |
| 38 | ||
| 12018 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 paulson parents: 
11704diff
changeset | 39 | Goal "(inverse x < (0::hypreal)) = (x < 0)"; | 
| 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 paulson parents: 
11704diff
changeset | 40 | by (case_tac "x=0" 1); | 
| 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 paulson parents: 
11704diff
changeset | 41 | by (asm_simp_tac (HOL_ss addsimps [HYPREAL_INVERSE_ZERO]) 1); | 
| 10751 | 42 | by (auto_tac (claset() addDs [hypreal_inverse_less_0], | 
| 43 | simpset() addsimps [linorder_neq_iff, | |
| 44 | hypreal_inverse_gt_0])); | |
| 45 | qed "hypreal_inverse_less_0_iff"; | |
| 46 | Addsimps [hypreal_inverse_less_0_iff]; | |
| 47 | ||
| 12018 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 paulson parents: 
11704diff
changeset | 48 | Goal "((0::hypreal) <= inverse x) = (0 <= x)"; | 
| 10751 | 49 | by (simp_tac (simpset() addsimps [linorder_not_less RS sym]) 1); | 
| 50 | qed "hypreal_0_le_inverse_iff"; | |
| 51 | Addsimps [hypreal_0_le_inverse_iff]; | |
| 52 | ||
| 12018 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 paulson parents: 
11704diff
changeset | 53 | Goal "(inverse x <= (0::hypreal)) = (x <= 0)"; | 
| 10751 | 54 | by (simp_tac (simpset() addsimps [linorder_not_less RS sym]) 1); | 
| 55 | qed "hypreal_inverse_le_0_iff"; | |
| 56 | Addsimps [hypreal_inverse_le_0_iff]; | |
| 57 | ||
| 12018 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 paulson parents: 
11704diff
changeset | 58 | Goalw [hypreal_divide_def] "x/(0::hypreal) = 0"; | 
| 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 paulson parents: 
11704diff
changeset | 59 | by (stac (HYPREAL_INVERSE_ZERO) 1); | 
| 10751 | 60 | by (Simp_tac 1); | 
| 61 | qed "HYPREAL_DIVIDE_ZERO"; | |
| 62 | ||
| 12018 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 paulson parents: 
11704diff
changeset | 63 | Goal "inverse (x::hypreal) = 1/x"; | 
| 10751 | 64 | by (simp_tac (simpset() addsimps [hypreal_divide_def]) 1); | 
| 65 | qed "hypreal_inverse_eq_divide"; | |
| 66 | ||
| 12018 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 paulson parents: 
11704diff
changeset | 67 | Goal "((0::hypreal) < x/y) = (0 < x & 0 < y | x < 0 & y < 0)"; | 
| 10751 | 68 | by (simp_tac (simpset() addsimps [hypreal_divide_def, hypreal_0_less_mult_iff]) 1); | 
| 69 | qed "hypreal_0_less_divide_iff"; | |
| 70 | Addsimps [inst "x" "number_of ?w" hypreal_0_less_divide_iff]; | |
| 71 | ||
| 12018 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 paulson parents: 
11704diff
changeset | 72 | Goal "(x/y < (0::hypreal)) = (0 < x & y < 0 | x < 0 & 0 < y)"; | 
| 10751 | 73 | by (simp_tac (simpset() addsimps [hypreal_divide_def, hypreal_mult_less_0_iff]) 1); | 
| 74 | qed "hypreal_divide_less_0_iff"; | |
| 75 | Addsimps [inst "x" "number_of ?w" hypreal_divide_less_0_iff]; | |
| 76 | ||
| 12018 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 paulson parents: 
11704diff
changeset | 77 | Goal "((0::hypreal) <= x/y) = ((x <= 0 | 0 <= y) & (0 <= x | y <= 0))"; | 
| 10751 | 78 | by (simp_tac (simpset() addsimps [hypreal_divide_def, hypreal_0_le_mult_iff]) 1); | 
| 79 | by Auto_tac; | |
| 80 | qed "hypreal_0_le_divide_iff"; | |
| 81 | Addsimps [inst "x" "number_of ?w" hypreal_0_le_divide_iff]; | |
| 82 | ||
| 12018 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 paulson parents: 
11704diff
changeset | 83 | Goal "(x/y <= (0::hypreal)) = ((x <= 0 | y <= 0) & (0 <= x | 0 <= y))"; | 
| 10751 | 84 | by (simp_tac (simpset() addsimps [hypreal_divide_def, | 
| 85 | hypreal_mult_le_0_iff]) 1); | |
| 86 | by Auto_tac; | |
| 87 | qed "hypreal_divide_le_0_iff"; | |
| 88 | Addsimps [inst "x" "number_of ?w" hypreal_divide_le_0_iff]; | |
| 89 | ||
| 12018 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 paulson parents: 
11704diff
changeset | 90 | Goal "(inverse(x::hypreal) = 0) = (x = 0)"; | 
| 10751 | 91 | by (auto_tac (claset(), | 
| 12018 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 paulson parents: 
11704diff
changeset | 92 | simpset() addsimps [HYPREAL_INVERSE_ZERO])); | 
| 10751 | 93 | by (rtac ccontr 1); | 
| 12018 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 paulson parents: 
11704diff
changeset | 94 | by (blast_tac (claset() addDs [hypreal_inverse_not_zero]) 1); | 
| 10751 | 95 | qed "hypreal_inverse_zero_iff"; | 
| 96 | Addsimps [hypreal_inverse_zero_iff]; | |
| 97 | ||
| 12018 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 paulson parents: 
11704diff
changeset | 98 | Goal "(x/y = 0) = (x=0 | y=(0::hypreal))"; | 
| 10751 | 99 | by (auto_tac (claset(), simpset() addsimps [hypreal_divide_def])); | 
| 100 | qed "hypreal_divide_eq_0_iff"; | |
| 101 | Addsimps [hypreal_divide_eq_0_iff]; | |
| 102 | ||
| 12018 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 paulson parents: 
11704diff
changeset | 103 | Goal "h ~= (0::hypreal) ==> h/h = 1"; | 
| 10751 | 104 | by (asm_simp_tac | 
| 105 | (simpset() addsimps [hypreal_divide_def, hypreal_mult_inverse_left]) 1); | |
| 106 | qed "hypreal_divide_self_eq"; | |
| 107 | Addsimps [hypreal_divide_self_eq]; | |
| 108 | ||
| 109 | ||
| 110 | (**** Factor cancellation theorems for "hypreal" ****) | |
| 111 | ||
| 112 | (** Cancellation laws for k*m < k*n and m*k < n*k, also for <= and =, | |
| 113 | but not (yet?) for k*m < n*k. **) | |
| 114 | ||
| 115 | bind_thm ("hypreal_mult_minus_right", hypreal_minus_mult_eq2 RS sym);
 | |
| 116 | ||
| 117 | Goal "(-y < -x) = ((x::hypreal) < y)"; | |
| 118 | by (arith_tac 1); | |
| 119 | qed "hypreal_minus_less_minus"; | |
| 120 | Addsimps [hypreal_minus_less_minus]; | |
| 121 | ||
| 122 | Goal "[| i<j; k < (0::hypreal) |] ==> j*k < i*k"; | |
| 123 | by (rtac (hypreal_minus_less_minus RS iffD1) 1); | |
| 124 | by (auto_tac (claset(), | |
| 125 | simpset() delsimps [hypreal_minus_mult_eq2 RS sym] | |
| 126 | addsimps [hypreal_minus_mult_eq2, | |
| 127 | hypreal_mult_less_mono1])); | |
| 128 | qed "hypreal_mult_less_mono1_neg"; | |
| 129 | ||
| 130 | Goal "[| i<j; k < (0::hypreal) |] ==> k*j < k*i"; | |
| 131 | by (rtac (hypreal_minus_less_minus RS iffD1) 1); | |
| 132 | by (auto_tac (claset(), | |
| 133 | simpset() delsimps [hypreal_minus_mult_eq1 RS sym] | |
| 134 | addsimps [hypreal_minus_mult_eq1, | |
| 135 | hypreal_mult_less_mono2])); | |
| 136 | qed "hypreal_mult_less_mono2_neg"; | |
| 137 | ||
| 138 | Goal "[| i <= j; k <= (0::hypreal) |] ==> j*k <= i*k"; | |
| 139 | by (auto_tac (claset(), | |
| 140 | simpset() addsimps [order_le_less, hypreal_mult_less_mono1_neg])); | |
| 141 | qed "hypreal_mult_le_mono1_neg"; | |
| 142 | ||
| 143 | Goal "[| i <= j; k <= (0::hypreal) |] ==> k*j <= k*i"; | |
| 144 | by (dtac hypreal_mult_le_mono1_neg 1); | |
| 145 | by (ALLGOALS (asm_full_simp_tac (simpset() addsimps [hypreal_mult_commute]))); | |
| 146 | qed "hypreal_mult_le_mono2_neg"; | |
| 147 | ||
| 12018 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 paulson parents: 
11704diff
changeset | 148 | Goal "(m*k < n*k) = (((0::hypreal) < k & m<n) | (k < 0 & n<m))"; | 
| 10751 | 149 | by (case_tac "k = (0::hypreal)" 1); | 
| 150 | by (auto_tac (claset(), | |
| 151 | simpset() addsimps [linorder_neq_iff, | |
| 152 | hypreal_mult_less_mono1, hypreal_mult_less_mono1_neg])); | |
| 153 | by (auto_tac (claset(), | |
| 154 | simpset() addsimps [linorder_not_less, | |
| 155 | inst "y1" "m*k" (linorder_not_le RS sym), | |
| 156 | inst "y1" "m" (linorder_not_le RS sym)])); | |
| 157 | by (TRYALL (etac notE)); | |
| 158 | by (auto_tac (claset(), | |
| 159 | simpset() addsimps [order_less_imp_le, hypreal_mult_le_mono1, | |
| 160 | hypreal_mult_le_mono1_neg])); | |
| 161 | qed "hypreal_mult_less_cancel2"; | |
| 162 | ||
| 12018 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 paulson parents: 
11704diff
changeset | 163 | Goal "(m*k <= n*k) = (((0::hypreal) < k --> m<=n) & (k < 0 --> n<=m))"; | 
| 10751 | 164 | by (simp_tac (simpset() addsimps [linorder_not_less RS sym, | 
| 165 | hypreal_mult_less_cancel2]) 1); | |
| 166 | qed "hypreal_mult_le_cancel2"; | |
| 167 | ||
| 12018 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 paulson parents: 
11704diff
changeset | 168 | Goal "(k*m < k*n) = (((0::hypreal) < k & m<n) | (k < 0 & n<m))"; | 
| 10751 | 169 | by (simp_tac (simpset() addsimps [inst "z" "k" hypreal_mult_commute, | 
| 170 | hypreal_mult_less_cancel2]) 1); | |
| 171 | qed "hypreal_mult_less_cancel1"; | |
| 172 | ||
| 12018 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 paulson parents: 
11704diff
changeset | 173 | Goal "!!k::hypreal. (k*m <= k*n) = ((0 < k --> m<=n) & (k < 0 --> n<=m))"; | 
| 10751 | 174 | by (simp_tac (simpset() addsimps [linorder_not_less RS sym, | 
| 175 | hypreal_mult_less_cancel1]) 1); | |
| 176 | qed "hypreal_mult_le_cancel1"; | |
| 177 | ||
| 12018 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 paulson parents: 
11704diff
changeset | 178 | Goal "!!k::hypreal. (k*m = k*n) = (k = 0 | m=n)"; | 
| 10751 | 179 | by (case_tac "k=0" 1); | 
| 180 | by (auto_tac (claset(), simpset() addsimps [hypreal_mult_left_cancel])); | |
| 181 | qed "hypreal_mult_eq_cancel1"; | |
| 182 | ||
| 12018 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 paulson parents: 
11704diff
changeset | 183 | Goal "!!k::hypreal. (m*k = n*k) = (k = 0 | m=n)"; | 
| 10751 | 184 | by (case_tac "k=0" 1); | 
| 185 | by (auto_tac (claset(), simpset() addsimps [hypreal_mult_right_cancel])); | |
| 186 | qed "hypreal_mult_eq_cancel2"; | |
| 187 | ||
| 12018 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 paulson parents: 
11704diff
changeset | 188 | Goal "!!k::hypreal. k~=0 ==> (k*m) / (k*n) = (m/n)"; | 
| 10751 | 189 | by (asm_simp_tac | 
| 190 | (simpset() addsimps [hypreal_divide_def, hypreal_inverse_distrib]) 1); | |
| 191 | by (subgoal_tac "k * m * (inverse k * inverse n) = \ | |
| 192 | \ (k * inverse k) * (m * inverse n)" 1); | |
| 193 | by (asm_full_simp_tac (simpset() addsimps []) 1); | |
| 194 | by (asm_full_simp_tac (HOL_ss addsimps hypreal_mult_ac) 1); | |
| 195 | qed "hypreal_mult_div_cancel1"; | |
| 196 | ||
| 197 | (*For ExtractCommonTerm*) | |
| 12018 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 paulson parents: 
11704diff
changeset | 198 | Goal "(k*m) / (k*n) = (if k = (0::hypreal) then 0 else m/n)"; | 
| 10751 | 199 | by (simp_tac (simpset() addsimps [hypreal_mult_div_cancel1]) 1); | 
| 200 | qed "hypreal_mult_div_cancel_disj"; | |
| 201 | ||
| 202 | ||
| 203 | local | |
| 204 | open Hyperreal_Numeral_Simprocs | |
| 205 | in | |
| 206 | ||
| 207 | val rel_hypreal_number_of = [eq_hypreal_number_of, less_hypreal_number_of, | |
| 208 | le_hypreal_number_of_eq_not_less]; | |
| 209 | ||
| 210 | structure CancelNumeralFactorCommon = | |
| 211 | struct | |
| 212 | val mk_coeff = mk_coeff | |
| 213 | val dest_coeff = dest_coeff 1 | |
| 12018 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 paulson parents: 
11704diff
changeset | 214 | val trans_tac = Real_Numeral_Simprocs.trans_tac | 
| 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 paulson parents: 
11704diff
changeset | 215 | val norm_tac = | 
| 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 paulson parents: 
11704diff
changeset | 216 | ALLGOALS (simp_tac (HOL_ss addsimps hypreal_minus_from_mult_simps @ mult_1s)) | 
| 10751 | 217 | THEN ALLGOALS (simp_tac (HOL_ss addsimps bin_simps@hypreal_mult_minus_simps)) | 
| 12018 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 paulson parents: 
11704diff
changeset | 218 | THEN ALLGOALS (simp_tac (HOL_ss addsimps hypreal_mult_ac)) | 
| 10751 | 219 | val numeral_simp_tac = | 
| 220 | ALLGOALS (simp_tac (HOL_ss addsimps rel_hypreal_number_of@bin_simps)) | |
| 221 | val simplify_meta_eq = simplify_meta_eq | |
| 222 | end | |
| 223 | ||
| 224 | structure DivCancelNumeralFactor = CancelNumeralFactorFun | |
| 225 | (open CancelNumeralFactorCommon | |
| 12018 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 paulson parents: 
11704diff
changeset | 226 | val prove_conv = Real_Numeral_Simprocs.prove_conv | 
| 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 paulson parents: 
11704diff
changeset | 227 | "hyprealdiv_cancel_numeral_factor" | 
| 10751 | 228 | val mk_bal = HOLogic.mk_binop "HOL.divide" | 
| 229 | val dest_bal = HOLogic.dest_bin "HOL.divide" hyprealT | |
| 230 | val cancel = hypreal_mult_div_cancel1 RS trans | |
| 231 | val neg_exchanges = false | |
| 232 | ) | |
| 233 | ||
| 234 | structure EqCancelNumeralFactor = CancelNumeralFactorFun | |
| 235 | (open CancelNumeralFactorCommon | |
| 12018 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 paulson parents: 
11704diff
changeset | 236 | val prove_conv = Real_Numeral_Simprocs.prove_conv | 
| 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 paulson parents: 
11704diff
changeset | 237 | "hyprealeq_cancel_numeral_factor" | 
| 10751 | 238 | val mk_bal = HOLogic.mk_eq | 
| 239 | val dest_bal = HOLogic.dest_bin "op =" hyprealT | |
| 240 | val cancel = hypreal_mult_eq_cancel1 RS trans | |
| 241 | val neg_exchanges = false | |
| 242 | ) | |
| 243 | ||
| 244 | structure LessCancelNumeralFactor = CancelNumeralFactorFun | |
| 245 | (open CancelNumeralFactorCommon | |
| 12018 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 paulson parents: 
11704diff
changeset | 246 | val prove_conv = Real_Numeral_Simprocs.prove_conv | 
| 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 paulson parents: 
11704diff
changeset | 247 | "hyprealless_cancel_numeral_factor" | 
| 10751 | 248 | val mk_bal = HOLogic.mk_binrel "op <" | 
| 249 | val dest_bal = HOLogic.dest_bin "op <" hyprealT | |
| 250 | val cancel = hypreal_mult_less_cancel1 RS trans | |
| 251 | val neg_exchanges = true | |
| 252 | ) | |
| 253 | ||
| 254 | structure LeCancelNumeralFactor = CancelNumeralFactorFun | |
| 255 | (open CancelNumeralFactorCommon | |
| 12018 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 paulson parents: 
11704diff
changeset | 256 | val prove_conv = Real_Numeral_Simprocs.prove_conv | 
| 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 paulson parents: 
11704diff
changeset | 257 | "hyprealle_cancel_numeral_factor" | 
| 10751 | 258 | val mk_bal = HOLogic.mk_binrel "op <=" | 
| 259 | val dest_bal = HOLogic.dest_bin "op <=" hyprealT | |
| 260 | val cancel = hypreal_mult_le_cancel1 RS trans | |
| 261 | val neg_exchanges = true | |
| 262 | ) | |
| 263 | ||
| 264 | val hypreal_cancel_numeral_factors_relations = | |
| 265 | map prep_simproc | |
| 266 |    [("hyprealeq_cancel_numeral_factor",
 | |
| 267 | prep_pats ["(l::hypreal) * m = n", "(l::hypreal) = m * n"], | |
| 268 | EqCancelNumeralFactor.proc), | |
| 269 |     ("hyprealless_cancel_numeral_factor", 
 | |
| 270 | prep_pats ["(l::hypreal) * m < n", "(l::hypreal) < m * n"], | |
| 271 | LessCancelNumeralFactor.proc), | |
| 272 |     ("hyprealle_cancel_numeral_factor", 
 | |
| 273 | prep_pats ["(l::hypreal) * m <= n", "(l::hypreal) <= m * n"], | |
| 274 | LeCancelNumeralFactor.proc)]; | |
| 275 | ||
| 276 | val hypreal_cancel_numeral_factors_divide = prep_simproc | |
| 277 | 	("hyprealdiv_cancel_numeral_factor", 
 | |
| 10825 
47c4a76b0c7a
additional pattern allows reduction of fractions to lowest terms
 paulson parents: 
10784diff
changeset | 278 | prep_pats ["((l::hypreal) * m) / n", "(l::hypreal) / (m * n)", | 
| 
47c4a76b0c7a
additional pattern allows reduction of fractions to lowest terms
 paulson parents: 
10784diff
changeset | 279 | "((number_of v)::hypreal) / (number_of w)"], | 
| 10751 | 280 | DivCancelNumeralFactor.proc); | 
| 281 | ||
| 282 | val hypreal_cancel_numeral_factors = | |
| 283 | hypreal_cancel_numeral_factors_relations @ | |
| 284 | [hypreal_cancel_numeral_factors_divide]; | |
| 285 | ||
| 286 | end; | |
| 287 | ||
| 288 | Addsimprocs hypreal_cancel_numeral_factors; | |
| 289 | ||
| 290 | ||
| 291 | (*examples: | |
| 292 | print_depth 22; | |
| 293 | set timing; | |
| 294 | set trace_simp; | |
| 295 | fun test s = (Goal s; by (Simp_tac 1)); | |
| 296 | ||
| 12018 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 paulson parents: 
11704diff
changeset | 297 | test "0 <= (y::hypreal) * -2"; | 
| 11704 
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
 wenzelm parents: 
11701diff
changeset | 298 | test "9*x = 12 * (y::hypreal)"; | 
| 
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
 wenzelm parents: 
11701diff
changeset | 299 | test "(9*x) / (12 * (y::hypreal)) = z"; | 
| 
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
 wenzelm parents: 
11701diff
changeset | 300 | test "9*x < 12 * (y::hypreal)"; | 
| 
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
 wenzelm parents: 
11701diff
changeset | 301 | test "9*x <= 12 * (y::hypreal)"; | 
| 10751 | 302 | |
| 11704 
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
 wenzelm parents: 
11701diff
changeset | 303 | test "-99*x = 123 * (y::hypreal)"; | 
| 
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
 wenzelm parents: 
11701diff
changeset | 304 | test "(-99*x) / (123 * (y::hypreal)) = z"; | 
| 
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
 wenzelm parents: 
11701diff
changeset | 305 | test "-99*x < 123 * (y::hypreal)"; | 
| 
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
 wenzelm parents: 
11701diff
changeset | 306 | test "-99*x <= 123 * (y::hypreal)"; | 
| 10751 | 307 | |
| 11704 
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
 wenzelm parents: 
11701diff
changeset | 308 | test "999*x = -396 * (y::hypreal)"; | 
| 
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
 wenzelm parents: 
11701diff
changeset | 309 | test "(999*x) / (-396 * (y::hypreal)) = z"; | 
| 
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
 wenzelm parents: 
11701diff
changeset | 310 | test "999*x < -396 * (y::hypreal)"; | 
| 
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
 wenzelm parents: 
11701diff
changeset | 311 | test "999*x <= -396 * (y::hypreal)"; | 
| 10751 | 312 | |
| 11704 
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
 wenzelm parents: 
11701diff
changeset | 313 | test "-99*x = -81 * (y::hypreal)"; | 
| 
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
 wenzelm parents: 
11701diff
changeset | 314 | test "(-99*x) / (-81 * (y::hypreal)) = z"; | 
| 
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
 wenzelm parents: 
11701diff
changeset | 315 | test "-99*x <= -81 * (y::hypreal)"; | 
| 
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
 wenzelm parents: 
11701diff
changeset | 316 | test "-99*x < -81 * (y::hypreal)"; | 
| 10751 | 317 | |
| 11704 
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
 wenzelm parents: 
11701diff
changeset | 318 | test "-2 * x = -1 * (y::hypreal)"; | 
| 
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
 wenzelm parents: 
11701diff
changeset | 319 | test "-2 * x = -(y::hypreal)"; | 
| 
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
 wenzelm parents: 
11701diff
changeset | 320 | test "(-2 * x) / (-1 * (y::hypreal)) = z"; | 
| 
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
 wenzelm parents: 
11701diff
changeset | 321 | test "-2 * x < -(y::hypreal)"; | 
| 
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
 wenzelm parents: 
11701diff
changeset | 322 | test "-2 * x <= -1 * (y::hypreal)"; | 
| 
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
 wenzelm parents: 
11701diff
changeset | 323 | test "-x < -23 * (y::hypreal)"; | 
| 
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
 wenzelm parents: 
11701diff
changeset | 324 | test "-x <= -23 * (y::hypreal)"; | 
| 10751 | 325 | *) | 
| 326 | ||
| 327 | ||
| 328 | (** Declarations for ExtractCommonTerm **) | |
| 329 | ||
| 330 | local | |
| 331 | open Hyperreal_Numeral_Simprocs | |
| 332 | in | |
| 333 | ||
| 334 | structure CancelFactorCommon = | |
| 335 | struct | |
| 336 | val mk_sum = long_mk_prod | |
| 337 | val dest_sum = dest_prod | |
| 338 | val mk_coeff = mk_coeff | |
| 339 | val dest_coeff = dest_coeff | |
| 340 | val find_first = find_first [] | |
| 12018 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 paulson parents: 
11704diff
changeset | 341 | val trans_tac = Real_Numeral_Simprocs.trans_tac | 
| 10751 | 342 | val norm_tac = ALLGOALS (simp_tac (HOL_ss addsimps mult_1s@hypreal_mult_ac)) | 
| 343 | end; | |
| 344 | ||
| 345 | structure EqCancelFactor = ExtractCommonTermFun | |
| 346 | (open CancelFactorCommon | |
| 12018 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 paulson parents: 
11704diff
changeset | 347 | val prove_conv = Real_Numeral_Simprocs.prove_conv | 
| 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 paulson parents: 
11704diff
changeset | 348 | "hypreal_eq_cancel_factor" | 
| 10751 | 349 | val mk_bal = HOLogic.mk_eq | 
| 350 | val dest_bal = HOLogic.dest_bin "op =" hyprealT | |
| 351 | val simplify_meta_eq = cancel_simplify_meta_eq hypreal_mult_eq_cancel1 | |
| 352 | ); | |
| 353 | ||
| 354 | ||
| 355 | structure DivideCancelFactor = ExtractCommonTermFun | |
| 356 | (open CancelFactorCommon | |
| 12018 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 paulson parents: 
11704diff
changeset | 357 | val prove_conv = Real_Numeral_Simprocs.prove_conv | 
| 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 paulson parents: 
11704diff
changeset | 358 | "hypreal_divide_cancel_factor" | 
| 10751 | 359 | val mk_bal = HOLogic.mk_binop "HOL.divide" | 
| 360 | val dest_bal = HOLogic.dest_bin "HOL.divide" hyprealT | |
| 361 | val simplify_meta_eq = cancel_simplify_meta_eq hypreal_mult_div_cancel_disj | |
| 362 | ); | |
| 363 | ||
| 364 | val hypreal_cancel_factor = | |
| 365 | map prep_simproc | |
| 366 |    [("hypreal_eq_cancel_factor",
 | |
| 367 | prep_pats ["(l::hypreal) * m = n", "(l::hypreal) = m * n"], | |
| 368 | EqCancelFactor.proc), | |
| 369 |     ("hypreal_divide_cancel_factor", 
 | |
| 370 | prep_pats ["((l::hypreal) * m) / n", "(l::hypreal) / (m * n)"], | |
| 371 | DivideCancelFactor.proc)]; | |
| 372 | ||
| 373 | end; | |
| 374 | ||
| 375 | Addsimprocs hypreal_cancel_factor; | |
| 376 | ||
| 377 | ||
| 378 | (*examples: | |
| 379 | print_depth 22; | |
| 380 | set timing; | |
| 381 | set trace_simp; | |
| 382 | fun test s = (Goal s; by (Asm_simp_tac 1)); | |
| 383 | ||
| 384 | test "x*k = k*(y::hypreal)"; | |
| 385 | test "k = k*(y::hypreal)"; | |
| 386 | test "a*(b*c) = (b::hypreal)"; | |
| 387 | test "a*(b*c) = d*(b::hypreal)*(x*a)"; | |
| 388 | ||
| 389 | ||
| 390 | test "(x*k) / (k*(y::hypreal)) = (uu::hypreal)"; | |
| 391 | test "(k) / (k*(y::hypreal)) = (uu::hypreal)"; | |
| 392 | test "(a*(b*c)) / ((b::hypreal)) = (uu::hypreal)"; | |
| 393 | test "(a*(b*c)) / (d*(b::hypreal)*(x*a)) = (uu::hypreal)"; | |
| 394 | ||
| 395 | (*FIXME: what do we do about this?*) | |
| 396 | test "a*(b*c)/(y*z) = d*(b::hypreal)*(x*a)/z"; | |
| 397 | *) | |
| 398 | ||
| 399 | ||
| 400 | (*** Simplification of inequalities involving literal divisors ***) | |
| 401 | ||
| 12018 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 paulson parents: 
11704diff
changeset | 402 | Goal "0<z ==> ((x::hypreal) <= y/z) = (x*z <= y)"; | 
| 10751 | 403 | by (subgoal_tac "(x*z <= y) = (x*z <= (y/z)*z)" 1); | 
| 404 | by (asm_simp_tac (simpset() addsimps [hypreal_divide_def, hypreal_mult_assoc]) 2); | |
| 405 | by (etac ssubst 1); | |
| 406 | by (stac hypreal_mult_le_cancel2 1); | |
| 407 | by (Asm_simp_tac 1); | |
| 408 | qed "pos_hypreal_le_divide_eq"; | |
| 409 | Addsimps [inst "z" "number_of ?w" pos_hypreal_le_divide_eq]; | |
| 410 | ||
| 12018 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 paulson parents: 
11704diff
changeset | 411 | Goal "z<0 ==> ((x::hypreal) <= y/z) = (y <= x*z)"; | 
| 10751 | 412 | by (subgoal_tac "(y <= x*z) = ((y/z)*z <= x*z)" 1); | 
| 413 | by (asm_simp_tac (simpset() addsimps [hypreal_divide_def, hypreal_mult_assoc]) 2); | |
| 414 | by (etac ssubst 1); | |
| 415 | by (stac hypreal_mult_le_cancel2 1); | |
| 416 | by (Asm_simp_tac 1); | |
| 417 | qed "neg_hypreal_le_divide_eq"; | |
| 418 | Addsimps [inst "z" "number_of ?w" neg_hypreal_le_divide_eq]; | |
| 419 | ||
| 12018 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 paulson parents: 
11704diff
changeset | 420 | Goal "0<z ==> (y/z <= (x::hypreal)) = (y <= x*z)"; | 
| 10751 | 421 | by (subgoal_tac "(y <= x*z) = ((y/z)*z <= x*z)" 1); | 
| 422 | by (asm_simp_tac (simpset() addsimps [hypreal_divide_def, hypreal_mult_assoc]) 2); | |
| 423 | by (etac ssubst 1); | |
| 424 | by (stac hypreal_mult_le_cancel2 1); | |
| 425 | by (Asm_simp_tac 1); | |
| 426 | qed "pos_hypreal_divide_le_eq"; | |
| 427 | Addsimps [inst "z" "number_of ?w" pos_hypreal_divide_le_eq]; | |
| 428 | ||
| 12018 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 paulson parents: 
11704diff
changeset | 429 | Goal "z<0 ==> (y/z <= (x::hypreal)) = (x*z <= y)"; | 
| 10751 | 430 | by (subgoal_tac "(x*z <= y) = (x*z <= (y/z)*z)" 1); | 
| 431 | by (asm_simp_tac (simpset() addsimps [hypreal_divide_def, hypreal_mult_assoc]) 2); | |
| 432 | by (etac ssubst 1); | |
| 433 | by (stac hypreal_mult_le_cancel2 1); | |
| 434 | by (Asm_simp_tac 1); | |
| 435 | qed "neg_hypreal_divide_le_eq"; | |
| 436 | Addsimps [inst "z" "number_of ?w" neg_hypreal_divide_le_eq]; | |
| 437 | ||
| 12018 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 paulson parents: 
11704diff
changeset | 438 | Goal "0<z ==> ((x::hypreal) < y/z) = (x*z < y)"; | 
| 10751 | 439 | by (subgoal_tac "(x*z < y) = (x*z < (y/z)*z)" 1); | 
| 440 | by (asm_simp_tac (simpset() addsimps [hypreal_divide_def, hypreal_mult_assoc]) 2); | |
| 441 | by (etac ssubst 1); | |
| 442 | by (stac hypreal_mult_less_cancel2 1); | |
| 443 | by (Asm_simp_tac 1); | |
| 444 | qed "pos_hypreal_less_divide_eq"; | |
| 445 | Addsimps [inst "z" "number_of ?w" pos_hypreal_less_divide_eq]; | |
| 446 | ||
| 12018 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 paulson parents: 
11704diff
changeset | 447 | Goal "z<0 ==> ((x::hypreal) < y/z) = (y < x*z)"; | 
| 10751 | 448 | by (subgoal_tac "(y < x*z) = ((y/z)*z < x*z)" 1); | 
| 449 | by (asm_simp_tac (simpset() addsimps [hypreal_divide_def, hypreal_mult_assoc]) 2); | |
| 450 | by (etac ssubst 1); | |
| 451 | by (stac hypreal_mult_less_cancel2 1); | |
| 452 | by (Asm_simp_tac 1); | |
| 453 | qed "neg_hypreal_less_divide_eq"; | |
| 454 | Addsimps [inst "z" "number_of ?w" neg_hypreal_less_divide_eq]; | |
| 455 | ||
| 12018 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 paulson parents: 
11704diff
changeset | 456 | Goal "0<z ==> (y/z < (x::hypreal)) = (y < x*z)"; | 
| 10751 | 457 | by (subgoal_tac "(y < x*z) = ((y/z)*z < x*z)" 1); | 
| 458 | by (asm_simp_tac (simpset() addsimps [hypreal_divide_def, hypreal_mult_assoc]) 2); | |
| 459 | by (etac ssubst 1); | |
| 460 | by (stac hypreal_mult_less_cancel2 1); | |
| 461 | by (Asm_simp_tac 1); | |
| 462 | qed "pos_hypreal_divide_less_eq"; | |
| 463 | Addsimps [inst "z" "number_of ?w" pos_hypreal_divide_less_eq]; | |
| 464 | ||
| 12018 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 paulson parents: 
11704diff
changeset | 465 | Goal "z<0 ==> (y/z < (x::hypreal)) = (x*z < y)"; | 
| 10751 | 466 | by (subgoal_tac "(x*z < y) = (x*z < (y/z)*z)" 1); | 
| 467 | by (asm_simp_tac (simpset() addsimps [hypreal_divide_def, hypreal_mult_assoc]) 2); | |
| 468 | by (etac ssubst 1); | |
| 469 | by (stac hypreal_mult_less_cancel2 1); | |
| 470 | by (Asm_simp_tac 1); | |
| 471 | qed "neg_hypreal_divide_less_eq"; | |
| 472 | Addsimps [inst "z" "number_of ?w" neg_hypreal_divide_less_eq]; | |
| 473 | ||
| 12018 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 paulson parents: 
11704diff
changeset | 474 | Goal "z~=0 ==> ((x::hypreal) = y/z) = (x*z = y)"; | 
| 10751 | 475 | by (subgoal_tac "(x*z = y) = (x*z = (y/z)*z)" 1); | 
| 476 | by (asm_simp_tac (simpset() addsimps [hypreal_divide_def, hypreal_mult_assoc]) 2); | |
| 477 | by (etac ssubst 1); | |
| 478 | by (stac hypreal_mult_eq_cancel2 1); | |
| 479 | by (Asm_simp_tac 1); | |
| 480 | qed "hypreal_eq_divide_eq"; | |
| 481 | Addsimps [inst "z" "number_of ?w" hypreal_eq_divide_eq]; | |
| 482 | ||
| 12018 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 paulson parents: 
11704diff
changeset | 483 | Goal "z~=0 ==> (y/z = (x::hypreal)) = (y = x*z)"; | 
| 10751 | 484 | by (subgoal_tac "(y = x*z) = ((y/z)*z = x*z)" 1); | 
| 485 | by (asm_simp_tac (simpset() addsimps [hypreal_divide_def, hypreal_mult_assoc]) 2); | |
| 486 | by (etac ssubst 1); | |
| 487 | by (stac hypreal_mult_eq_cancel2 1); | |
| 488 | by (Asm_simp_tac 1); | |
| 489 | qed "hypreal_divide_eq_eq"; | |
| 490 | Addsimps [inst "z" "number_of ?w" hypreal_divide_eq_eq]; | |
| 491 | ||
| 12018 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 paulson parents: 
11704diff
changeset | 492 | Goal "(m/k = n/k) = (k = 0 | m = (n::hypreal))"; | 
| 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 paulson parents: 
11704diff
changeset | 493 | by (case_tac "k=0" 1); | 
| 10751 | 494 | by (asm_simp_tac (simpset() addsimps [HYPREAL_DIVIDE_ZERO]) 1); | 
| 495 | by (asm_simp_tac (simpset() addsimps [hypreal_divide_eq_eq, hypreal_eq_divide_eq, | |
| 496 | hypreal_mult_eq_cancel2]) 1); | |
| 497 | qed "hypreal_divide_eq_cancel2"; | |
| 498 | ||
| 12018 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 paulson parents: 
11704diff
changeset | 499 | Goal "(k/m = k/n) = (k = 0 | m = (n::hypreal))"; | 
| 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 paulson parents: 
11704diff
changeset | 500 | by (case_tac "m=0 | n = 0" 1); | 
| 10751 | 501 | by (auto_tac (claset(), | 
| 502 | simpset() addsimps [HYPREAL_DIVIDE_ZERO, hypreal_divide_eq_eq, | |
| 503 | hypreal_eq_divide_eq, hypreal_mult_eq_cancel1])); | |
| 504 | qed "hypreal_divide_eq_cancel1"; | |
| 505 | ||
| 12018 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 paulson parents: 
11704diff
changeset | 506 | Goal "[| 0 < r; 0 < x|] ==> (inverse x < inverse (r::hypreal)) = (r < x)"; | 
| 10751 | 507 | by (auto_tac (claset() addIs [hypreal_inverse_less_swap], simpset())); | 
| 508 | by (res_inst_tac [("t","r")] (hypreal_inverse_inverse RS subst) 1);
 | |
| 509 | by (res_inst_tac [("t","x")] (hypreal_inverse_inverse RS subst) 1);
 | |
| 510 | by (auto_tac (claset() addIs [hypreal_inverse_less_swap], | |
| 511 | simpset() delsimps [hypreal_inverse_inverse] | |
| 12018 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 paulson parents: 
11704diff
changeset | 512 | addsimps [hypreal_inverse_gt_0])); | 
| 10751 | 513 | qed "hypreal_inverse_less_iff"; | 
| 514 | ||
| 12018 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 paulson parents: 
11704diff
changeset | 515 | Goal "[| 0 < r; 0 < x|] ==> (inverse x <= inverse r) = (r <= (x::hypreal))"; | 
| 10751 | 516 | by (asm_simp_tac (simpset() addsimps [linorder_not_less RS sym, | 
| 517 | hypreal_inverse_less_iff]) 1); | |
| 518 | qed "hypreal_inverse_le_iff"; | |
| 519 | ||
| 520 | (** Division by 1, -1 **) | |
| 521 | ||
| 12018 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 paulson parents: 
11704diff
changeset | 522 | Goal "(x::hypreal)/1 = x"; | 
| 10751 | 523 | by (simp_tac (simpset() addsimps [hypreal_divide_def]) 1); | 
| 524 | qed "hypreal_divide_1"; | |
| 525 | Addsimps [hypreal_divide_1]; | |
| 526 | ||
| 11704 
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
 wenzelm parents: 
11701diff
changeset | 527 | Goal "x/-1 = -(x::hypreal)"; | 
| 10751 | 528 | by (Simp_tac 1); | 
| 529 | qed "hypreal_divide_minus1"; | |
| 530 | Addsimps [hypreal_divide_minus1]; | |
| 531 | ||
| 12018 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 paulson parents: 
11704diff
changeset | 532 | Goal "-1/(x::hypreal) = - (1/x)"; | 
| 10751 | 533 | by (simp_tac (simpset() addsimps [hypreal_divide_def, hypreal_minus_inverse]) 1); | 
| 534 | qed "hypreal_minus1_divide"; | |
| 535 | Addsimps [hypreal_minus1_divide]; | |
| 536 | ||
| 12018 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 paulson parents: 
11704diff
changeset | 537 | Goal "[| (0::hypreal) < d1; 0 < d2 |] ==> EX e. 0 < e & e < d1 & e < d2"; | 
| 11704 
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
 wenzelm parents: 
11701diff
changeset | 538 | by (res_inst_tac [("x","(min d1 d2)/2")] exI 1); 
 | 
| 10751 | 539 | by (asm_simp_tac (simpset() addsimps [min_def]) 1); | 
| 540 | qed "hypreal_lbound_gt_zero"; | |
| 541 | ||
| 542 | ||
| 543 | (*** General rewrites to improve automation, like those for type "int" ***) | |
| 544 | ||
| 545 | (** The next several equations can make the simplifier loop! **) | |
| 546 | ||
| 547 | Goal "(x < - y) = (y < - (x::hypreal))"; | |
| 548 | by Auto_tac; | |
| 549 | qed "hypreal_less_minus"; | |
| 550 | ||
| 551 | Goal "(- x < y) = (- y < (x::hypreal))"; | |
| 552 | by Auto_tac; | |
| 553 | qed "hypreal_minus_less"; | |
| 554 | ||
| 555 | Goal "(x <= - y) = (y <= - (x::hypreal))"; | |
| 556 | by Auto_tac; | |
| 557 | qed "hypreal_le_minus"; | |
| 558 | ||
| 559 | Goal "(- x <= y) = (- y <= (x::hypreal))"; | |
| 560 | by Auto_tac; | |
| 561 | qed "hypreal_minus_le"; | |
| 562 | ||
| 563 | Goal "(x = - y) = (y = - (x::hypreal))"; | |
| 564 | by Auto_tac; | |
| 565 | qed "hypreal_equation_minus"; | |
| 566 | ||
| 567 | Goal "(- x = y) = (- (y::hypreal) = x)"; | |
| 568 | by Auto_tac; | |
| 569 | qed "hypreal_minus_equation"; | |
| 570 | ||
| 12018 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 paulson parents: 
11704diff
changeset | 571 | Goal "(x + - a = (0::hypreal)) = (x=a)"; | 
| 10751 | 572 | by (arith_tac 1); | 
| 573 | qed "hypreal_add_minus_iff"; | |
| 574 | Addsimps [hypreal_add_minus_iff]; | |
| 575 | ||
| 576 | Goal "(-b = -a) = (b = (a::hypreal))"; | |
| 577 | by (arith_tac 1); | |
| 578 | qed "hypreal_minus_eq_cancel"; | |
| 579 | Addsimps [hypreal_minus_eq_cancel]; | |
| 580 | ||
| 581 | Goal "(-s <= -r) = ((r::hypreal) <= s)"; | |
| 582 | by (stac hypreal_minus_le 1); | |
| 583 | by (Simp_tac 1); | |
| 584 | qed "hypreal_le_minus_iff"; | |
| 585 | Addsimps [hypreal_le_minus_iff]; | |
| 586 | ||
| 587 | ||
| 588 | (*Distributive laws for literals*) | |
| 589 | Addsimps (map (inst "w" "number_of ?v") | |
| 590 | [hypreal_add_mult_distrib, hypreal_add_mult_distrib2, | |
| 591 | hypreal_diff_mult_distrib, hypreal_diff_mult_distrib2]); | |
| 592 | ||
| 593 | Addsimps (map (inst "x" "number_of ?v") | |
| 594 | [hypreal_less_minus, hypreal_le_minus, hypreal_equation_minus]); | |
| 595 | Addsimps (map (inst "y" "number_of ?v") | |
| 596 | [hypreal_minus_less, hypreal_minus_le, hypreal_minus_equation]); | |
| 597 | ||
| 12018 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 paulson parents: 
11704diff
changeset | 598 | Addsimps (map (simplify (simpset()) o inst "x" "1") | 
| 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 paulson parents: 
11704diff
changeset | 599 | [hypreal_less_minus, hypreal_le_minus, hypreal_equation_minus]); | 
| 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 paulson parents: 
11704diff
changeset | 600 | Addsimps (map (simplify (simpset()) o inst "y" "1") | 
| 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 paulson parents: 
11704diff
changeset | 601 | [hypreal_minus_less, hypreal_minus_le, hypreal_minus_equation]); | 
| 10751 | 602 | |
| 12018 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 paulson parents: 
11704diff
changeset | 603 | (*** Simprules combining x+y and 0 ***) | 
| 10751 | 604 | |
| 12018 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 paulson parents: 
11704diff
changeset | 605 | Goal "(x+y = (0::hypreal)) = (y = -x)"; | 
| 10751 | 606 | by Auto_tac; | 
| 607 | qed "hypreal_add_eq_0_iff"; | |
| 608 | AddIffs [hypreal_add_eq_0_iff]; | |
| 609 | ||
| 12018 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 paulson parents: 
11704diff
changeset | 610 | Goal "(x+y < (0::hypreal)) = (y < -x)"; | 
| 10751 | 611 | by Auto_tac; | 
| 612 | qed "hypreal_add_less_0_iff"; | |
| 613 | AddIffs [hypreal_add_less_0_iff]; | |
| 614 | ||
| 12018 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 paulson parents: 
11704diff
changeset | 615 | Goal "((0::hypreal) < x+y) = (-x < y)"; | 
| 10751 | 616 | by Auto_tac; | 
| 617 | qed "hypreal_0_less_add_iff"; | |
| 618 | AddIffs [hypreal_0_less_add_iff]; | |
| 619 | ||
| 12018 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 paulson parents: 
11704diff
changeset | 620 | Goal "(x+y <= (0::hypreal)) = (y <= -x)"; | 
| 10751 | 621 | by Auto_tac; | 
| 622 | qed "hypreal_add_le_0_iff"; | |
| 623 | AddIffs [hypreal_add_le_0_iff]; | |
| 624 | ||
| 12018 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 paulson parents: 
11704diff
changeset | 625 | Goal "((0::hypreal) <= x+y) = (-x <= y)"; | 
| 10751 | 626 | by Auto_tac; | 
| 627 | qed "hypreal_0_le_add_iff"; | |
| 628 | AddIffs [hypreal_0_le_add_iff]; | |
| 629 | ||
| 630 | ||
| 12018 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 paulson parents: 
11704diff
changeset | 631 | (** Simprules combining x-y and 0; see also hypreal_less_iff_diff_less_0 etc | 
| 10751 | 632 | in HyperBin | 
| 633 | **) | |
| 634 | ||
| 12018 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 paulson parents: 
11704diff
changeset | 635 | Goal "((0::hypreal) < x-y) = (y < x)"; | 
| 10751 | 636 | by Auto_tac; | 
| 637 | qed "hypreal_0_less_diff_iff"; | |
| 638 | AddIffs [hypreal_0_less_diff_iff]; | |
| 639 | ||
| 12018 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 paulson parents: 
11704diff
changeset | 640 | Goal "((0::hypreal) <= x-y) = (y <= x)"; | 
| 10751 | 641 | by Auto_tac; | 
| 642 | qed "hypreal_0_le_diff_iff"; | |
| 643 | AddIffs [hypreal_0_le_diff_iff]; | |
| 644 | ||
| 645 | (* | |
| 646 | FIXME: we should have this, as for type int, but many proofs would break. | |
| 647 | It replaces x+-y by x-y. | |
| 648 | Addsimps [symmetric hypreal_diff_def]; | |
| 649 | *) | |
| 650 | ||
| 651 | Goal "-(x-y) = y - (x::hypreal)"; | |
| 652 | by (arith_tac 1); | |
| 653 | qed "hypreal_minus_diff_eq"; | |
| 654 | Addsimps [hypreal_minus_diff_eq]; | |
| 655 | ||
| 656 | ||
| 657 | (*** Density of the Hyperreals ***) | |
| 658 | ||
| 11704 
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
 wenzelm parents: 
11701diff
changeset | 659 | Goal "x < y ==> x < (x+y) / (2::hypreal)"; | 
| 10751 | 660 | by Auto_tac; | 
| 661 | qed "hypreal_less_half_sum"; | |
| 662 | ||
| 11704 
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
 wenzelm parents: 
11701diff
changeset | 663 | Goal "x < y ==> (x+y)/(2::hypreal) < y"; | 
| 10751 | 664 | by Auto_tac; | 
| 665 | qed "hypreal_gt_half_sum"; | |
| 666 | ||
| 667 | Goal "x < y ==> EX r::hypreal. x < r & r < y"; | |
| 668 | by (blast_tac (claset() addSIs [hypreal_less_half_sum, hypreal_gt_half_sum]) 1); | |
| 669 | qed "hypreal_dense"; | |
| 670 | ||
| 671 | ||
| 12018 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 paulson parents: 
11704diff
changeset | 672 | (*Replaces "inverse #nn" by 1/#nn *) | 
| 10751 | 673 | Addsimps [inst "x" "number_of ?w" hypreal_inverse_eq_divide]; | 
| 674 | ||
| 675 |