| author | paulson | 
| Wed, 26 Jun 2002 10:26:46 +0200 | |
| changeset 13248 | ae66c22ed52e | 
| parent 12018 | ec054019c910 | 
| child 14299 | 0b5c0b0a3eba | 
| permissions | -rw-r--r-- | 
| 10750 | 1 | (* Title : HRealAbs.ML | 
| 2 | Author : Jacques D. Fleuriot | |
| 3 | Copyright : 1998 University of Cambridge | |
| 4 | Description : Absolute value function for the hyperreals | |
| 5 | Similar to RealAbs.thy | |
| 6 | *) | |
| 7 | ||
| 8 | (*------------------------------------------------------------ | |
| 9 | absolute value on hyperreals as pointwise operation on | |
| 10 | equivalence class representative | |
| 11 | ------------------------------------------------------------*) | |
| 12 | ||
| 13 | Goalw [hrabs_def] | |
| 14 | "abs (number_of v :: hypreal) = \ | |
| 15 | \ (if neg (number_of v) then number_of (bin_minus v) \ | |
| 16 | \ else number_of v)"; | |
| 17 | by (Simp_tac 1); | |
| 18 | qed "hrabs_number_of"; | |
| 19 | Addsimps [hrabs_number_of]; | |
| 20 | ||
| 21 | Goalw [hrabs_def] | |
| 10834 | 22 |      "abs (Abs_hypreal (hyprel `` {X})) = \
 | 
| 23 | \     Abs_hypreal(hyprel `` {%n. abs (X n)})";
 | |
| 10750 | 24 | by (auto_tac (claset(), | 
| 25 | simpset_of HyperDef.thy | |
| 26 | addsimps [hypreal_zero_def, hypreal_le,hypreal_minus])); | |
| 27 | by (ALLGOALS(Ultra_tac THEN' arith_tac )); | |
| 28 | qed "hypreal_hrabs"; | |
| 29 | ||
| 30 | (*------------------------------------------------------------ | |
| 31 | Properties of the absolute value function over the reals | |
| 32 | (adapted version of previously proved theorems about abs) | |
| 33 | ------------------------------------------------------------*) | |
| 34 | ||
| 12018 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 paulson parents: 
11713diff
changeset | 35 | Goal "abs (0::hypreal) = 0"; | 
| 10750 | 36 | by (simp_tac (simpset() addsimps [hrabs_def]) 1); | 
| 37 | qed "hrabs_zero"; | |
| 38 | Addsimps [hrabs_zero]; | |
| 39 | ||
| 12018 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 paulson parents: 
11713diff
changeset | 40 | Goal "abs (1::hypreal) = 1"; | 
| 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 paulson parents: 
11713diff
changeset | 41 | by (simp_tac (simpset() addsimps [hrabs_def]) 1); | 
| 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 paulson parents: 
11713diff
changeset | 42 | qed "hrabs_one"; | 
| 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 paulson parents: 
11713diff
changeset | 43 | Addsimps [hrabs_one]; | 
| 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 paulson parents: 
11713diff
changeset | 44 | |
| 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 paulson parents: 
11713diff
changeset | 45 | Goal "(0::hypreal)<=x ==> abs x = x"; | 
| 10750 | 46 | by (asm_simp_tac (simpset() addsimps [hrabs_def]) 1); | 
| 47 | qed "hrabs_eqI1"; | |
| 48 | ||
| 12018 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 paulson parents: 
11713diff
changeset | 49 | Goal "(0::hypreal)<x ==> abs x = x"; | 
| 10750 | 50 | by (asm_simp_tac (simpset() addsimps [order_less_imp_le, hrabs_eqI1]) 1); | 
| 51 | qed "hrabs_eqI2"; | |
| 52 | ||
| 12018 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 paulson parents: 
11713diff
changeset | 53 | Goal "x<(0::hypreal) ==> abs x = -x"; | 
| 10750 | 54 | by (asm_simp_tac (simpset() addsimps [hypreal_le_def, hrabs_def]) 1); | 
| 55 | qed "hrabs_minus_eqI2"; | |
| 56 | ||
| 12018 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 paulson parents: 
11713diff
changeset | 57 | Goal "x<=(0::hypreal) ==> abs x = -x"; | 
| 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 paulson parents: 
11713diff
changeset | 58 | by (auto_tac (claset() addDs [order_antisym], simpset() addsimps [hrabs_def]));qed "hrabs_minus_eqI1"; | 
| 10750 | 59 | |
| 12018 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 paulson parents: 
11713diff
changeset | 60 | Goal "(0::hypreal)<= abs x"; | 
| 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 paulson parents: 
11713diff
changeset | 61 | by (simp_tac (simpset() addsimps [hrabs_def]) 1); | 
| 10750 | 62 | qed "hrabs_ge_zero"; | 
| 63 | ||
| 64 | Goal "abs(abs x) = abs (x::hypreal)"; | |
| 12018 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 paulson parents: 
11713diff
changeset | 65 | by (simp_tac (simpset() addsimps [hrabs_def]) 1); | 
| 10750 | 66 | qed "hrabs_idempotent"; | 
| 67 | Addsimps [hrabs_idempotent]; | |
| 68 | ||
| 12018 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 paulson parents: 
11713diff
changeset | 69 | Goalw [hrabs_def] "(abs x = (0::hypreal)) = (x=0)"; | 
| 10750 | 70 | by (Simp_tac 1); | 
| 71 | qed "hrabs_zero_iff"; | |
| 72 | AddIffs [hrabs_zero_iff]; | |
| 73 | ||
| 74 | Goalw [hrabs_def] "(x::hypreal) <= abs x"; | |
| 12018 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 paulson parents: 
11713diff
changeset | 75 | by (Simp_tac 1); | 
| 10750 | 76 | qed "hrabs_ge_self"; | 
| 77 | ||
| 78 | Goalw [hrabs_def] "-(x::hypreal) <= abs x"; | |
| 12018 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 paulson parents: 
11713diff
changeset | 79 | by (Simp_tac 1); | 
| 10750 | 80 | qed "hrabs_ge_minus_self"; | 
| 81 | ||
| 12018 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 paulson parents: 
11713diff
changeset | 82 | (* proof by "transfer" *) | 
| 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 paulson parents: 
11713diff
changeset | 83 | Goal "abs(x*(y::hypreal)) = (abs x)*(abs y)"; | 
| 10750 | 84 | by (res_inst_tac [("z","x")] eq_Abs_hypreal 1);
 | 
| 85 | by (res_inst_tac [("z","y")] eq_Abs_hypreal 1);
 | |
| 86 | by (auto_tac (claset(), | |
| 87 | simpset() addsimps [hypreal_hrabs, hypreal_mult,abs_mult])); | |
| 88 | qed "hrabs_mult"; | |
| 89 | Addsimps [hrabs_mult]; | |
| 90 | ||
| 91 | Goal "abs(inverse(x)) = inverse(abs(x::hypreal))"; | |
| 12018 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 paulson parents: 
11713diff
changeset | 92 | by (simp_tac (simpset() addsimps [hypreal_minus_inverse, hrabs_def]) 1); | 
| 10750 | 93 | qed "hrabs_inverse"; | 
| 94 | ||
| 12018 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 paulson parents: 
11713diff
changeset | 95 | Goalw [hrabs_def] "abs(x+(y::hypreal)) <= abs x + abs y"; | 
| 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 paulson parents: 
11713diff
changeset | 96 | by (Simp_tac 1); | 
| 10750 | 97 | qed "hrabs_triangle_ineq"; | 
| 98 | ||
| 99 | Goal "abs((w::hypreal) + x + y) <= abs(w) + abs(x) + abs(y)"; | |
| 12018 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 paulson parents: 
11713diff
changeset | 100 | by (simp_tac (simpset() addsimps [hrabs_triangle_ineq RS order_trans]) 1); | 
| 10750 | 101 | qed "hrabs_triangle_ineq_three"; | 
| 102 | ||
| 12018 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 paulson parents: 
11713diff
changeset | 103 | Goalw [hrabs_def] "abs(-x) = abs((x::hypreal))"; | 
| 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 paulson parents: 
11713diff
changeset | 104 | by (Simp_tac 1); | 
| 10750 | 105 | qed "hrabs_minus_cancel"; | 
| 106 | Addsimps [hrabs_minus_cancel]; | |
| 107 | ||
| 12018 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 paulson parents: 
11713diff
changeset | 108 | Goalw [hrabs_def] "[| abs x < r; abs y < s |] ==> abs(x+y) < r + (s::hypreal)"; | 
| 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 paulson parents: 
11713diff
changeset | 109 | by (asm_full_simp_tac (simpset() addsplits [split_if_asm]) 1); | 
| 10750 | 110 | qed "hrabs_add_less"; | 
| 111 | ||
| 112 | Goal "[| abs x<r; abs y<s |] ==> abs x * abs y < r * (s::hypreal)"; | |
| 12018 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 paulson parents: 
11713diff
changeset | 113 | by (subgoal_tac "0 < r" 1); | 
| 10750 | 114 | by (asm_full_simp_tac (simpset() addsimps [hrabs_def] | 
| 115 | addsplits [split_if_asm]) 2); | |
| 12018 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 paulson parents: 
11713diff
changeset | 116 | by (case_tac "y = 0" 1); | 
| 10750 | 117 | by (asm_full_simp_tac (simpset() addsimps [hypreal_0_less_mult_iff]) 1); | 
| 118 | by (rtac hypreal_mult_less_mono 1); | |
| 119 | by (auto_tac (claset(), | |
| 120 | simpset() addsimps [hrabs_def, linorder_neq_iff] | |
| 121 | addsplits [split_if_asm])); | |
| 122 | qed "hrabs_mult_less"; | |
| 123 | ||
| 12018 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 paulson parents: 
11713diff
changeset | 124 | Goal "((0::hypreal) < abs x) = (x ~= 0)"; | 
| 10750 | 125 | by (simp_tac (simpset() addsimps [hrabs_def]) 1); | 
| 126 | by (arith_tac 1); | |
| 127 | qed "hypreal_0_less_abs_iff"; | |
| 128 | Addsimps [hypreal_0_less_abs_iff]; | |
| 129 | ||
| 12018 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 paulson parents: 
11713diff
changeset | 130 | Goal "abs x < r ==> (0::hypreal) < r"; | 
| 10750 | 131 | by (blast_tac (claset() addSIs [order_le_less_trans, hrabs_ge_zero]) 1); | 
| 132 | qed "hrabs_less_gt_zero"; | |
| 133 | ||
| 134 | Goal "abs x = (x::hypreal) | abs x = -x"; | |
| 12018 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 paulson parents: 
11713diff
changeset | 135 | by (simp_tac (simpset() addsimps [hrabs_def]) 1); | 
| 10750 | 136 | qed "hrabs_disj"; | 
| 137 | ||
| 138 | Goal "abs x = (y::hypreal) ==> x = y | -x = y"; | |
| 12018 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 paulson parents: 
11713diff
changeset | 139 | by (asm_full_simp_tac (simpset() addsimps [hrabs_def] | 
| 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 paulson parents: 
11713diff
changeset | 140 | addsplits [split_if_asm]) 1); | 
| 10750 | 141 | qed "hrabs_eq_disj"; | 
| 142 | ||
| 12018 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 paulson parents: 
11713diff
changeset | 143 | Goalw [hrabs_def] "(abs x < (r::hypreal)) = (-r < x & x < r)"; | 
| 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 paulson parents: 
11713diff
changeset | 144 | by Auto_tac; | 
| 10750 | 145 | qed "hrabs_interval_iff"; | 
| 146 | ||
| 147 | Goal "(abs x < (r::hypreal)) = (- x < r & x < r)"; | |
| 148 | by (auto_tac (claset(), simpset() addsimps [hrabs_interval_iff])); | |
| 149 | qed "hrabs_interval_iff2"; | |
| 150 | ||
| 151 | ||
| 152 | (* Needed in Geom.ML *) | |
| 153 | Goal "(y::hypreal) + - x + (y + - z) = abs (x + - z) ==> y = z | x = y"; | |
| 154 | by (asm_full_simp_tac (simpset() addsimps [hrabs_def] | |
| 155 | addsplits [split_if_asm]) 1); | |
| 156 | qed "hrabs_add_lemma_disj"; | |
| 157 | ||
| 158 | Goal "abs((x::hypreal) + -y) = abs (y + -x)"; | |
| 159 | by (simp_tac (simpset() addsimps [hrabs_def]) 1); | |
| 160 | qed "hrabs_minus_add_cancel"; | |
| 161 | ||
| 162 | (* Needed in Geom.ML?? *) | |
| 163 | Goal "(x::hypreal) + - y + (z + - y) = abs (x + - z) ==> y = z | x = y"; | |
| 164 | by (asm_full_simp_tac (simpset() addsimps [hrabs_def] | |
| 165 | addsplits [split_if_asm]) 1); | |
| 166 | qed "hrabs_add_lemma_disj2"; | |
| 167 | ||
| 168 | ||
| 169 | (*---------------------------------------------------------- | |
| 170 | Relating hrabs to abs through embedding of IR into IR* | |
| 171 | ----------------------------------------------------------*) | |
| 172 | Goalw [hypreal_of_real_def] | |
| 173 | "abs (hypreal_of_real r) = hypreal_of_real (abs r)"; | |
| 174 | by (auto_tac (claset(), simpset() addsimps [hypreal_hrabs])); | |
| 175 | qed "hypreal_of_real_hrabs"; | |
| 10778 
2c6605049646
more tidying, especially to remove real_of_posnat
 paulson parents: 
10750diff
changeset | 176 | |
| 
2c6605049646
more tidying, especially to remove real_of_posnat
 paulson parents: 
10750diff
changeset | 177 | |
| 
2c6605049646
more tidying, especially to remove real_of_posnat
 paulson parents: 
10750diff
changeset | 178 | (*---------------------------------------------------------------------------- | 
| 
2c6605049646
more tidying, especially to remove real_of_posnat
 paulson parents: 
10750diff
changeset | 179 | Embedding of the naturals in the hyperreals | 
| 
2c6605049646
more tidying, especially to remove real_of_posnat
 paulson parents: 
10750diff
changeset | 180 | ----------------------------------------------------------------------------*) | 
| 
2c6605049646
more tidying, especially to remove real_of_posnat
 paulson parents: 
10750diff
changeset | 181 | |
| 10784 | 182 | Goal "hypreal_of_nat (m + n) = hypreal_of_nat m + hypreal_of_nat n"; | 
| 183 | by (simp_tac (simpset() addsimps [hypreal_of_nat_def]) 1); | |
| 10778 
2c6605049646
more tidying, especially to remove real_of_posnat
 paulson parents: 
10750diff
changeset | 184 | qed "hypreal_of_nat_add"; | 
| 10784 | 185 | Addsimps [hypreal_of_nat_add]; | 
| 186 | ||
| 187 | Goal "hypreal_of_nat (m * n) = hypreal_of_nat m * hypreal_of_nat n"; | |
| 188 | by (simp_tac (simpset() addsimps [hypreal_of_nat_def]) 1); | |
| 189 | qed "hypreal_of_nat_mult"; | |
| 190 | Addsimps [hypreal_of_nat_mult]; | |
| 10778 
2c6605049646
more tidying, especially to remove real_of_posnat
 paulson parents: 
10750diff
changeset | 191 | |
| 
2c6605049646
more tidying, especially to remove real_of_posnat
 paulson parents: 
10750diff
changeset | 192 | Goalw [hypreal_of_nat_def] | 
| 
2c6605049646
more tidying, especially to remove real_of_posnat
 paulson parents: 
10750diff
changeset | 193 | "(n < m) = (hypreal_of_nat n < hypreal_of_nat m)"; | 
| 
2c6605049646
more tidying, especially to remove real_of_posnat
 paulson parents: 
10750diff
changeset | 194 | by (auto_tac (claset() addIs [hypreal_add_less_mono1], simpset())); | 
| 
2c6605049646
more tidying, especially to remove real_of_posnat
 paulson parents: 
10750diff
changeset | 195 | qed "hypreal_of_nat_less_iff"; | 
| 
2c6605049646
more tidying, especially to remove real_of_posnat
 paulson parents: 
10750diff
changeset | 196 | Addsimps [hypreal_of_nat_less_iff RS sym]; | 
| 
2c6605049646
more tidying, especially to remove real_of_posnat
 paulson parents: 
10750diff
changeset | 197 | |
| 
2c6605049646
more tidying, especially to remove real_of_posnat
 paulson parents: 
10750diff
changeset | 198 | (*------------------------------------------------------------*) | 
| 
2c6605049646
more tidying, especially to remove real_of_posnat
 paulson parents: 
10750diff
changeset | 199 | (* naturals embedded in hyperreals *) | 
| 
2c6605049646
more tidying, especially to remove real_of_posnat
 paulson parents: 
10750diff
changeset | 200 | (* is a hyperreal c.f. NS extension *) | 
| 
2c6605049646
more tidying, especially to remove real_of_posnat
 paulson parents: 
10750diff
changeset | 201 | (*------------------------------------------------------------*) | 
| 
2c6605049646
more tidying, especially to remove real_of_posnat
 paulson parents: 
10750diff
changeset | 202 | |
| 
2c6605049646
more tidying, especially to remove real_of_posnat
 paulson parents: 
10750diff
changeset | 203 | Goalw [hypreal_of_nat_def, hypreal_of_real_def, real_of_nat_def] | 
| 10919 
144ede948e58
renamings: real_of_nat, real_of_int -> (overloaded) real
 paulson parents: 
10834diff
changeset | 204 |      "hypreal_of_nat  m = Abs_hypreal(hyprel``{%n. real m})";
 | 
| 10778 
2c6605049646
more tidying, especially to remove real_of_posnat
 paulson parents: 
10750diff
changeset | 205 | by Auto_tac; | 
| 
2c6605049646
more tidying, especially to remove real_of_posnat
 paulson parents: 
10750diff
changeset | 206 | qed "hypreal_of_nat_iff"; | 
| 
2c6605049646
more tidying, especially to remove real_of_posnat
 paulson parents: 
10750diff
changeset | 207 | |
| 
2c6605049646
more tidying, especially to remove real_of_posnat
 paulson parents: 
10750diff
changeset | 208 | Goal "inj hypreal_of_nat"; | 
| 
2c6605049646
more tidying, especially to remove real_of_posnat
 paulson parents: 
10750diff
changeset | 209 | by (simp_tac (simpset() addsimps [inj_on_def, hypreal_of_nat_def]) 1); | 
| 
2c6605049646
more tidying, especially to remove real_of_posnat
 paulson parents: 
10750diff
changeset | 210 | qed "inj_hypreal_of_nat"; | 
| 
2c6605049646
more tidying, especially to remove real_of_posnat
 paulson parents: 
10750diff
changeset | 211 | |
| 
2c6605049646
more tidying, especially to remove real_of_posnat
 paulson parents: 
10750diff
changeset | 212 | Goalw [hypreal_of_nat_def] | 
| 11713 
883d559b0b8c
sane numerals (stage 3): provide generic "1" on all number types;
 wenzelm parents: 
11701diff
changeset | 213 | "hypreal_of_nat (Suc n) = hypreal_of_nat n + (1::hypreal)"; | 
| 10778 
2c6605049646
more tidying, especially to remove real_of_posnat
 paulson parents: 
10750diff
changeset | 214 | by (simp_tac (simpset() addsimps [real_of_nat_Suc]) 1); | 
| 
2c6605049646
more tidying, especially to remove real_of_posnat
 paulson parents: 
10750diff
changeset | 215 | qed "hypreal_of_nat_Suc"; | 
| 
2c6605049646
more tidying, especially to remove real_of_posnat
 paulson parents: 
10750diff
changeset | 216 | |
| 
2c6605049646
more tidying, especially to remove real_of_posnat
 paulson parents: 
10750diff
changeset | 217 | (*"neg" is used in rewrite rules for binary comparisons*) | 
| 
2c6605049646
more tidying, especially to remove real_of_posnat
 paulson parents: 
10750diff
changeset | 218 | Goal "hypreal_of_nat (number_of v :: nat) = \ | 
| 12018 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 paulson parents: 
11713diff
changeset | 219 | \ (if neg (number_of v) then 0 \ | 
| 10778 
2c6605049646
more tidying, especially to remove real_of_posnat
 paulson parents: 
10750diff
changeset | 220 | \ else (number_of v :: hypreal))"; | 
| 
2c6605049646
more tidying, especially to remove real_of_posnat
 paulson parents: 
10750diff
changeset | 221 | by (simp_tac (simpset() addsimps [hypreal_of_nat_def]) 1); | 
| 
2c6605049646
more tidying, especially to remove real_of_posnat
 paulson parents: 
10750diff
changeset | 222 | qed "hypreal_of_nat_number_of"; | 
| 
2c6605049646
more tidying, especially to remove real_of_posnat
 paulson parents: 
10750diff
changeset | 223 | Addsimps [hypreal_of_nat_number_of]; | 
| 
2c6605049646
more tidying, especially to remove real_of_posnat
 paulson parents: 
10750diff
changeset | 224 | |
| 12018 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 paulson parents: 
11713diff
changeset | 225 | Goal "hypreal_of_nat 0 = 0"; | 
| 10778 
2c6605049646
more tidying, especially to remove real_of_posnat
 paulson parents: 
10750diff
changeset | 226 | by (simp_tac (simpset() delsimps [numeral_0_eq_0] | 
| 12018 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 paulson parents: 
11713diff
changeset | 227 | addsimps [numeral_0_eq_0 RS sym]) 1); | 
| 10778 
2c6605049646
more tidying, especially to remove real_of_posnat
 paulson parents: 
10750diff
changeset | 228 | qed "hypreal_of_nat_zero"; | 
| 
2c6605049646
more tidying, especially to remove real_of_posnat
 paulson parents: 
10750diff
changeset | 229 | Addsimps [hypreal_of_nat_zero]; | 
| 12018 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 paulson parents: 
11713diff
changeset | 230 | |
| 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 paulson parents: 
11713diff
changeset | 231 | Goal "hypreal_of_nat 1 = 1"; | 
| 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 paulson parents: 
11713diff
changeset | 232 | by (simp_tac (simpset() addsimps [hypreal_of_nat_Suc]) 1); | 
| 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 paulson parents: 
11713diff
changeset | 233 | qed "hypreal_of_nat_one"; | 
| 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 paulson parents: 
11713diff
changeset | 234 | Addsimps [hypreal_of_nat_one]; |