Numerals and simprocs for types real and hypreal. The abstract
constants 0, 1 and binary numerals work harmoniously.
--- a/src/HOL/Hyperreal/HRealAbs.ML Thu Nov 01 21:12:13 2001 +0100
+++ b/src/HOL/Hyperreal/HRealAbs.ML Fri Nov 02 17:55:24 2001 +0100
@@ -32,56 +32,55 @@
(adapted version of previously proved theorems about abs)
------------------------------------------------------------*)
-Goal "abs (Numeral0::hypreal) = Numeral0";
+Goal "abs (0::hypreal) = 0";
by (simp_tac (simpset() addsimps [hrabs_def]) 1);
qed "hrabs_zero";
Addsimps [hrabs_zero];
-Goal "(Numeral0::hypreal)<=x ==> abs x = x";
+Goal "abs (1::hypreal) = 1";
+by (simp_tac (simpset() addsimps [hrabs_def]) 1);
+qed "hrabs_one";
+Addsimps [hrabs_one];
+
+Goal "(0::hypreal)<=x ==> abs x = x";
by (asm_simp_tac (simpset() addsimps [hrabs_def]) 1);
qed "hrabs_eqI1";
-Goal "(Numeral0::hypreal)<x ==> abs x = x";
+Goal "(0::hypreal)<x ==> abs x = x";
by (asm_simp_tac (simpset() addsimps [order_less_imp_le, hrabs_eqI1]) 1);
qed "hrabs_eqI2";
-Goal "x<(Numeral0::hypreal) ==> abs x = -x";
+Goal "x<(0::hypreal) ==> abs x = -x";
by (asm_simp_tac (simpset() addsimps [hypreal_le_def, hrabs_def]) 1);
qed "hrabs_minus_eqI2";
-Goal "x<=(Numeral0::hypreal) ==> abs x = -x";
-by (auto_tac (claset() addDs [order_antisym], simpset() addsimps [hrabs_def]));
-qed "hrabs_minus_eqI1";
+Goal "x<=(0::hypreal) ==> abs x = -x";
+by (auto_tac (claset() addDs [order_antisym], simpset() addsimps [hrabs_def]));qed "hrabs_minus_eqI1";
-Goal "(Numeral0::hypreal)<= abs x";
-by (auto_tac (claset() addDs [hypreal_minus_zero_less_iff RS iffD2,
- hypreal_less_asym],
- simpset() addsimps [hypreal_le_def, hrabs_def]));
+Goal "(0::hypreal)<= abs x";
+by (simp_tac (simpset() addsimps [hrabs_def]) 1);
qed "hrabs_ge_zero";
Goal "abs(abs x) = abs (x::hypreal)";
-by (auto_tac (claset() addDs [hypreal_minus_zero_less_iff RS iffD2,
- hypreal_less_asym],
- simpset() addsimps [hypreal_le_def, hrabs_def]));
+by (simp_tac (simpset() addsimps [hrabs_def]) 1);
qed "hrabs_idempotent";
Addsimps [hrabs_idempotent];
-Goalw [hrabs_def] "(abs x = (Numeral0::hypreal)) = (x=Numeral0)";
+Goalw [hrabs_def] "(abs x = (0::hypreal)) = (x=0)";
by (Simp_tac 1);
qed "hrabs_zero_iff";
AddIffs [hrabs_zero_iff];
Goalw [hrabs_def] "(x::hypreal) <= abs x";
-by (auto_tac (claset() addDs [not_hypreal_leE, order_less_imp_le],
- simpset() addsimps [hypreal_le_zero_iff RS sym]));
+by (Simp_tac 1);
qed "hrabs_ge_self";
Goalw [hrabs_def] "-(x::hypreal) <= abs x";
-by (simp_tac (simpset() addsimps [hypreal_ge_zero_iff]) 1);
+by (Simp_tac 1);
qed "hrabs_ge_minus_self";
-(* very short proof by "transfer" *)
-Goal "abs(x*(y::hypreal)) = (abs x)*(abs y)";
+(* proof by "transfer" *)
+Goal "abs(x*(y::hypreal)) = (abs x)*(abs y)";
by (res_inst_tac [("z","x")] eq_Abs_hypreal 1);
by (res_inst_tac [("z","y")] eq_Abs_hypreal 1);
by (auto_tac (claset(),
@@ -90,48 +89,31 @@
Addsimps [hrabs_mult];
Goal "abs(inverse(x)) = inverse(abs(x::hypreal))";
-by (hypreal_div_undefined_case_tac "x=Numeral0" 1);
-by (simp_tac (simpset() addsimps [HYPREAL_DIVIDE_ZERO]) 1);
-by (res_inst_tac [("z","x")] eq_Abs_hypreal 1);
-by (auto_tac (claset(),
- simpset() addsimps [hypreal_hrabs,
- hypreal_inverse,hypreal_zero_def]));
-by (ultra_tac (claset(), simpset() addsimps [abs_inverse]) 1);
+by (simp_tac (simpset() addsimps [hypreal_minus_inverse, hrabs_def]) 1);
qed "hrabs_inverse";
-Goal "abs(x+(y::hypreal)) <= abs x + abs y";
-by (res_inst_tac [("z","x")] eq_Abs_hypreal 1);
-by (res_inst_tac [("z","y")] eq_Abs_hypreal 1);
-by (auto_tac (claset(),
- simpset() addsimps [hypreal_hrabs, hypreal_add,hypreal_le,
- abs_triangle_ineq]));
+Goalw [hrabs_def] "abs(x+(y::hypreal)) <= abs x + abs y";
+by (Simp_tac 1);
qed "hrabs_triangle_ineq";
Goal "abs((w::hypreal) + x + y) <= abs(w) + abs(x) + abs(y)";
-by (auto_tac (claset() addSIs [hrabs_triangle_ineq RS order_trans,
- hypreal_add_left_le_mono1],
- simpset() addsimps [hypreal_add_assoc]));
+by (simp_tac (simpset() addsimps [hrabs_triangle_ineq RS order_trans]) 1);
qed "hrabs_triangle_ineq_three";
-Goalw [hrabs_def] "abs(-x)=abs((x::hypreal))";
-by (auto_tac (claset() addSDs [not_hypreal_leE, hypreal_less_asym]
- addIs [hypreal_le_anti_sym],
- simpset() addsimps [hypreal_ge_zero_iff]));
+Goalw [hrabs_def] "abs(-x) = abs((x::hypreal))";
+by (Simp_tac 1);
qed "hrabs_minus_cancel";
Addsimps [hrabs_minus_cancel];
-val prem1::prem2::rest = goal thy
- "[| abs x < r; abs y < s |] ==> abs(x+y) < r + (s::hypreal)";
-by (rtac order_le_less_trans 1);
-by (rtac hrabs_triangle_ineq 1);
-by (rtac ([prem1,prem2] MRS hypreal_add_less_mono) 1);
+Goalw [hrabs_def] "[| abs x < r; abs y < s |] ==> abs(x+y) < r + (s::hypreal)";
+by (asm_full_simp_tac (simpset() addsplits [split_if_asm]) 1);
qed "hrabs_add_less";
Goal "[| abs x<r; abs y<s |] ==> abs x * abs y < r * (s::hypreal)";
-by (subgoal_tac "Numeral0 < r" 1);
+by (subgoal_tac "0 < r" 1);
by (asm_full_simp_tac (simpset() addsimps [hrabs_def]
addsplits [split_if_asm]) 2);
-by (case_tac "y = Numeral0" 1);
+by (case_tac "y = 0" 1);
by (asm_full_simp_tac (simpset() addsimps [hypreal_0_less_mult_iff]) 1);
by (rtac hypreal_mult_less_mono 1);
by (auto_tac (claset(),
@@ -139,40 +121,27 @@
addsplits [split_if_asm]));
qed "hrabs_mult_less";
-Goal "((Numeral0::hypreal) < abs x) = (x ~= 0)";
+Goal "((0::hypreal) < abs x) = (x ~= 0)";
by (simp_tac (simpset() addsimps [hrabs_def]) 1);
by (arith_tac 1);
qed "hypreal_0_less_abs_iff";
Addsimps [hypreal_0_less_abs_iff];
-Goal "abs x < r ==> (Numeral0::hypreal) < r";
+Goal "abs x < r ==> (0::hypreal) < r";
by (blast_tac (claset() addSIs [order_le_less_trans, hrabs_ge_zero]) 1);
qed "hrabs_less_gt_zero";
Goal "abs x = (x::hypreal) | abs x = -x";
-by (cut_inst_tac [("x","Numeral0"),("y","x")] hypreal_linear 1);
-by (fast_tac (claset() addIs [hrabs_eqI2,hrabs_minus_eqI2,
- hrabs_zero]) 1);
+by (simp_tac (simpset() addsimps [hrabs_def]) 1);
qed "hrabs_disj";
Goal "abs x = (y::hypreal) ==> x = y | -x = y";
-by (dtac sym 1);
-by (hyp_subst_tac 1);
-by (res_inst_tac [("x1","x")] (hrabs_disj RS disjE) 1);
-by (REPEAT(Asm_simp_tac 1));
+by (asm_full_simp_tac (simpset() addsimps [hrabs_def]
+ addsplits [split_if_asm]) 1);
qed "hrabs_eq_disj";
-Goal "(abs x < (r::hypreal)) = (-r < x & x < r)";
-by (Step_tac 1);
-by (rtac (hypreal_less_swap_iff RS iffD2) 1);
-by (asm_simp_tac (simpset() addsimps [(hrabs_ge_minus_self
- RS order_le_less_trans)]) 1);
-by (asm_simp_tac (simpset() addsimps [(hrabs_ge_self
- RS order_le_less_trans)]) 1);
-by (EVERY1 [dtac (hypreal_less_swap_iff RS iffD1), rotate_tac 1,
- dtac (hypreal_minus_minus RS subst),
- cut_inst_tac [("x","x")] hrabs_disj, dtac disjE ]);
-by (assume_tac 3 THEN Auto_tac);
+Goalw [hrabs_def] "(abs x < (r::hypreal)) = (-r < x & x < r)";
+by Auto_tac;
qed "hrabs_interval_iff";
Goal "(abs x < (r::hypreal)) = (- x < r & x < r)";
@@ -247,14 +216,19 @@
(*"neg" is used in rewrite rules for binary comparisons*)
Goal "hypreal_of_nat (number_of v :: nat) = \
-\ (if neg (number_of v) then Numeral0 \
+\ (if neg (number_of v) then 0 \
\ else (number_of v :: hypreal))";
by (simp_tac (simpset() addsimps [hypreal_of_nat_def]) 1);
qed "hypreal_of_nat_number_of";
Addsimps [hypreal_of_nat_number_of];
-Goal "hypreal_of_nat 0 = Numeral0";
+Goal "hypreal_of_nat 0 = 0";
by (simp_tac (simpset() delsimps [numeral_0_eq_0]
- addsimps [numeral_0_eq_0 RS sym]) 1);
+ addsimps [numeral_0_eq_0 RS sym]) 1);
qed "hypreal_of_nat_zero";
Addsimps [hypreal_of_nat_zero];
+
+Goal "hypreal_of_nat 1 = 1";
+by (simp_tac (simpset() addsimps [hypreal_of_nat_Suc]) 1);
+qed "hypreal_of_nat_one";
+Addsimps [hypreal_of_nat_one];
--- a/src/HOL/Hyperreal/HSeries.ML Thu Nov 01 21:12:13 2001 +0100
+++ b/src/HOL/Hyperreal/HSeries.ML Fri Nov 02 17:55:24 2001 +0100
@@ -35,7 +35,7 @@
(* Theorem corresponding to base case in def of sumr *)
Goalw [hypnat_zero_def]
- "sumhr (m,0,f) = Numeral0";
+ "sumhr (m,0,f) = 0";
by (res_inst_tac [("z","m")] eq_Abs_hypnat 1);
by (auto_tac (claset(),
simpset() addsimps [sumhr, symmetric hypreal_zero_def]));
@@ -44,31 +44,26 @@
(* Theorem corresponding to recursive case in def of sumr *)
Goalw [hypnat_one_def]
- "sumhr(m,n+(1::hypnat),f) = (if n + (1::hypnat) <= m then Numeral0 \
+ "sumhr(m,n+(1::hypnat),f) = (if n + (1::hypnat) <= m then 0 \
\ else sumhr(m,n,f) + (*fNat* f) n)";
-by (simp_tac (HOL_ss addsimps
- [zero_eq_numeral_0 RS sym, hypreal_zero_def]) 1);
by (res_inst_tac [("z","m")] eq_Abs_hypnat 1);
by (res_inst_tac [("z","n")] eq_Abs_hypnat 1);
by (auto_tac (claset(),
- simpset() addsimps [sumhr, hypnat_add,hypnat_le,starfunNat,hypreal_add]));
+ simpset() addsimps [sumhr, hypnat_add,hypnat_le,starfunNat,hypreal_add,
+ hypreal_zero_def]));
by (ALLGOALS(Ultra_tac));
qed "sumhr_if";
-Goalw [hypnat_one_def] "sumhr (n + (1::hypnat), n, f) = Numeral0";
-by (simp_tac (HOL_ss addsimps
- [zero_eq_numeral_0 RS sym, hypreal_zero_def]) 1);
+Goalw [hypnat_one_def] "sumhr (n + (1::hypnat), n, f) = 0";
by (res_inst_tac [("z","n")] eq_Abs_hypnat 1);
by (auto_tac (claset(),
- simpset() addsimps [sumhr, hypnat_add]));
+ simpset() addsimps [sumhr, hypnat_add, hypreal_zero_def]));
qed "sumhr_Suc_zero";
Addsimps [sumhr_Suc_zero];
-Goal "sumhr (n,n,f) = Numeral0";
-by (simp_tac (HOL_ss addsimps
- [zero_eq_numeral_0 RS sym, hypreal_zero_def]) 1);
+Goal "sumhr (n,n,f) = 0";
by (res_inst_tac [("z","n")] eq_Abs_hypnat 1);
-by (auto_tac (claset(), simpset() addsimps [sumhr]));
+by (auto_tac (claset(), simpset() addsimps [sumhr, hypreal_zero_def]));
qed "sumhr_eq_bounds";
Addsimps [sumhr_eq_bounds];
@@ -80,13 +75,11 @@
qed "sumhr_Suc";
Addsimps [sumhr_Suc];
-Goal "sumhr(m+k,k,f) = Numeral0";
-by (simp_tac (HOL_ss addsimps
- [zero_eq_numeral_0 RS sym, hypreal_zero_def]) 1);
+Goal "sumhr(m+k,k,f) = 0";
by (res_inst_tac [("z","m")] eq_Abs_hypnat 1);
by (res_inst_tac [("z","k")] eq_Abs_hypnat 1);
by (auto_tac (claset(),
- simpset() addsimps [sumhr, hypnat_add]));
+ simpset() addsimps [sumhr, hypnat_add, hypreal_zero_def]));
qed "sumhr_add_lbound_zero";
Addsimps [sumhr_add_lbound_zero];
@@ -156,13 +149,11 @@
hypreal_minus,sumr_add RS sym]) 1);
qed "sumhr_add_mult_const";
-Goal "n < m ==> sumhr (m,n,f) = Numeral0";
-by (simp_tac (HOL_ss addsimps
- [zero_eq_numeral_0 RS sym, hypreal_zero_def]) 1);
+Goal "n < m ==> sumhr (m,n,f) = 0";
by (res_inst_tac [("z","m")] eq_Abs_hypnat 1);
by (res_inst_tac [("z","n")] eq_Abs_hypnat 1);
by (auto_tac (claset() addEs [FreeUltrafilterNat_subset],
- simpset() addsimps [sumhr,hypnat_less]));
+ simpset() addsimps [sumhr,hypnat_less, hypreal_zero_def]));
qed "sumhr_less_bounds_zero";
Addsimps [sumhr_less_bounds_zero];
@@ -185,25 +176,23 @@
by summing to some infinite hypernatural (such as whn)
-----------------------------------------------------------------*)
Goalw [hypnat_omega_def,hypnat_zero_def]
- "sumhr(0,whn,%i. Numeral1) = hypreal_of_hypnat whn";
+ "sumhr(0,whn,%i. 1) = hypreal_of_hypnat whn";
by (auto_tac (claset(),
simpset() addsimps [sumhr, hypreal_of_hypnat]));
qed "sumhr_hypreal_of_hypnat_omega";
Goalw [hypnat_omega_def,hypnat_zero_def,omega_def]
- "sumhr(0, whn, %i. Numeral1) = omega - Numeral1";
+ "sumhr(0, whn, %i. 1) = omega - 1";
by (simp_tac (HOL_ss addsimps
- [one_eq_numeral_1 RS sym, hypreal_one_def]) 1);
+ [hypreal_numeral_1_eq_1, hypreal_one_def]) 1);
by (auto_tac (claset(),
simpset() addsimps [sumhr, hypreal_diff, real_of_nat_Suc]));
qed "sumhr_hypreal_omega_minus_one";
Goalw [hypnat_zero_def, hypnat_omega_def]
- "sumhr(0, whn + whn, %i. (-Numeral1) ^ (i+1)) = Numeral0";
-by (simp_tac (HOL_ss addsimps
- [zero_eq_numeral_0 RS sym, hypreal_zero_def]) 1);
-by (simp_tac (simpset() addsimps [sumhr,hypnat_add,double_lemma]
- delsimps [realpow_Suc]) 1);
+ "sumhr(0, whn + whn, %i. (-1) ^ (i+1)) = 0";
+by (simp_tac (simpset() delsimps [realpow_Suc]
+ addsimps [sumhr,hypnat_add,double_lemma, hypreal_zero_def]) 1);
qed "sumhr_minus_one_realpow_zero";
Addsimps [sumhr_minus_one_realpow_zero];
@@ -223,7 +212,7 @@
qed "starfunNat_sumr";
Goal "sumhr (0, M, f) @= sumhr (0, N, f) \
-\ ==> abs (sumhr (M, N, f)) @= Numeral0";
+\ ==> abs (sumhr (M, N, f)) @= 0";
by (cut_inst_tac [("x","M"),("y","N")] hypnat_linear 1);
by (auto_tac (claset(), simpset() addsimps [approx_refl]));
by (dtac (approx_sym RS (approx_minus_iff RS iffD1)) 1);
@@ -265,12 +254,12 @@
sums_unique]) 1);
qed "NSsums_unique";
-Goal "ALL m. n <= Suc m --> f(m) = Numeral0 ==> f NSsums (sumr 0 n f)";
+Goal "ALL m. n <= Suc m --> f(m) = 0 ==> f NSsums (sumr 0 n f)";
by (asm_simp_tac (simpset() addsimps [sums_NSsums_iff RS sym, series_zero]) 1);
qed "NSseries_zero";
Goal "NSsummable f = \
-\ (ALL M: HNatInfinite. ALL N: HNatInfinite. abs (sumhr(M,N,f)) @= Numeral0)";
+\ (ALL M: HNatInfinite. ALL N: HNatInfinite. abs (sumhr(M,N,f)) @= 0)";
by (auto_tac (claset(),
simpset() addsimps [summable_NSsummable_iff RS sym,
summable_convergent_sumr_iff, convergent_NSconvergent_iff,
@@ -287,17 +276,16 @@
(*-------------------------------------------------------------------
Terms of a convergent series tend to zero
-------------------------------------------------------------------*)
-Goalw [NSLIMSEQ_def] "NSsummable f ==> f ----NS> Numeral0";
+Goalw [NSLIMSEQ_def] "NSsummable f ==> f ----NS> 0";
by (auto_tac (claset(), simpset() addsimps [NSsummable_NSCauchy]));
by (dtac bspec 1 THEN Auto_tac);
by (dres_inst_tac [("x","N + (1::hypnat)")] bspec 1);
-by (auto_tac (claset() addIs [HNatInfinite_add_one,
- approx_hrabs_zero_cancel],
- simpset() addsimps [rename_numerals hypreal_of_real_zero]));
+by (auto_tac (claset() addIs [HNatInfinite_add_one, approx_hrabs_zero_cancel],
+ simpset()));
qed "NSsummable_NSLIMSEQ_zero";
(* Easy to prove stsandard case now *)
-Goal "summable f ==> f ----> Numeral0";
+Goal "summable f ==> f ----> 0";
by (auto_tac (claset(),
simpset() addsimps [summable_NSsummable_iff,
LIMSEQ_NSLIMSEQ_iff, NSsummable_NSLIMSEQ_zero]));
--- a/src/HOL/Hyperreal/HyperArith0.ML Thu Nov 01 21:12:13 2001 +0100
+++ b/src/HOL/Hyperreal/HyperArith0.ML Fri Nov 02 17:55:24 2001 +0100
@@ -8,7 +8,12 @@
Also, common factor cancellation
*)
-Goal "((x * y = Numeral0) = (x = Numeral0 | y = (Numeral0::hypreal)))";
+Goal "x - - y = x + (y::hypreal)";
+by (Simp_tac 1);
+qed "hypreal_diff_minus_eq";
+Addsimps [hypreal_diff_minus_eq];
+
+Goal "((x * y = 0) = (x = 0 | y = (0::hypreal)))";
by Auto_tac;
by (cut_inst_tac [("x","x"),("y","y")] hypreal_mult_zero_disj 1);
by Auto_tac;
@@ -17,85 +22,85 @@
(** Division and inverse **)
-Goal "Numeral0/x = (Numeral0::hypreal)";
+Goal "0/x = (0::hypreal)";
by (simp_tac (simpset() addsimps [hypreal_divide_def]) 1);
qed "hypreal_0_divide";
Addsimps [hypreal_0_divide];
-Goal "((Numeral0::hypreal) < inverse x) = (Numeral0 < x)";
-by (case_tac "x=Numeral0" 1);
-by (asm_simp_tac (HOL_ss addsimps [rename_numerals HYPREAL_INVERSE_ZERO]) 1);
+Goal "((0::hypreal) < inverse x) = (0 < x)";
+by (case_tac "x=0" 1);
+by (asm_simp_tac (HOL_ss addsimps [HYPREAL_INVERSE_ZERO]) 1);
by (auto_tac (claset() addDs [hypreal_inverse_less_0],
simpset() addsimps [linorder_neq_iff,
hypreal_inverse_gt_0]));
qed "hypreal_0_less_inverse_iff";
Addsimps [hypreal_0_less_inverse_iff];
-Goal "(inverse x < (Numeral0::hypreal)) = (x < Numeral0)";
-by (case_tac "x=Numeral0" 1);
-by (asm_simp_tac (HOL_ss addsimps [rename_numerals HYPREAL_INVERSE_ZERO]) 1);
+Goal "(inverse x < (0::hypreal)) = (x < 0)";
+by (case_tac "x=0" 1);
+by (asm_simp_tac (HOL_ss addsimps [HYPREAL_INVERSE_ZERO]) 1);
by (auto_tac (claset() addDs [hypreal_inverse_less_0],
simpset() addsimps [linorder_neq_iff,
hypreal_inverse_gt_0]));
qed "hypreal_inverse_less_0_iff";
Addsimps [hypreal_inverse_less_0_iff];
-Goal "((Numeral0::hypreal) <= inverse x) = (Numeral0 <= x)";
+Goal "((0::hypreal) <= inverse x) = (0 <= x)";
by (simp_tac (simpset() addsimps [linorder_not_less RS sym]) 1);
qed "hypreal_0_le_inverse_iff";
Addsimps [hypreal_0_le_inverse_iff];
-Goal "(inverse x <= (Numeral0::hypreal)) = (x <= Numeral0)";
+Goal "(inverse x <= (0::hypreal)) = (x <= 0)";
by (simp_tac (simpset() addsimps [linorder_not_less RS sym]) 1);
qed "hypreal_inverse_le_0_iff";
Addsimps [hypreal_inverse_le_0_iff];
-Goalw [hypreal_divide_def] "x/(Numeral0::hypreal) = Numeral0";
-by (stac (rename_numerals HYPREAL_INVERSE_ZERO) 1);
+Goalw [hypreal_divide_def] "x/(0::hypreal) = 0";
+by (stac (HYPREAL_INVERSE_ZERO) 1);
by (Simp_tac 1);
qed "HYPREAL_DIVIDE_ZERO";
-Goal "inverse (x::hypreal) = Numeral1/x";
+Goal "inverse (x::hypreal) = 1/x";
by (simp_tac (simpset() addsimps [hypreal_divide_def]) 1);
qed "hypreal_inverse_eq_divide";
-Goal "((Numeral0::hypreal) < x/y) = (Numeral0 < x & Numeral0 < y | x < Numeral0 & y < Numeral0)";
+Goal "((0::hypreal) < x/y) = (0 < x & 0 < y | x < 0 & y < 0)";
by (simp_tac (simpset() addsimps [hypreal_divide_def, hypreal_0_less_mult_iff]) 1);
qed "hypreal_0_less_divide_iff";
Addsimps [inst "x" "number_of ?w" hypreal_0_less_divide_iff];
-Goal "(x/y < (Numeral0::hypreal)) = (Numeral0 < x & y < Numeral0 | x < Numeral0 & Numeral0 < y)";
+Goal "(x/y < (0::hypreal)) = (0 < x & y < 0 | x < 0 & 0 < y)";
by (simp_tac (simpset() addsimps [hypreal_divide_def, hypreal_mult_less_0_iff]) 1);
qed "hypreal_divide_less_0_iff";
Addsimps [inst "x" "number_of ?w" hypreal_divide_less_0_iff];
-Goal "((Numeral0::hypreal) <= x/y) = ((x <= Numeral0 | Numeral0 <= y) & (Numeral0 <= x | y <= Numeral0))";
+Goal "((0::hypreal) <= x/y) = ((x <= 0 | 0 <= y) & (0 <= x | y <= 0))";
by (simp_tac (simpset() addsimps [hypreal_divide_def, hypreal_0_le_mult_iff]) 1);
by Auto_tac;
qed "hypreal_0_le_divide_iff";
Addsimps [inst "x" "number_of ?w" hypreal_0_le_divide_iff];
-Goal "(x/y <= (Numeral0::hypreal)) = ((x <= Numeral0 | y <= Numeral0) & (Numeral0 <= x | Numeral0 <= y))";
+Goal "(x/y <= (0::hypreal)) = ((x <= 0 | y <= 0) & (0 <= x | 0 <= y))";
by (simp_tac (simpset() addsimps [hypreal_divide_def,
hypreal_mult_le_0_iff]) 1);
by Auto_tac;
qed "hypreal_divide_le_0_iff";
Addsimps [inst "x" "number_of ?w" hypreal_divide_le_0_iff];
-Goal "(inverse(x::hypreal) = Numeral0) = (x = Numeral0)";
+Goal "(inverse(x::hypreal) = 0) = (x = 0)";
by (auto_tac (claset(),
- simpset() addsimps [rename_numerals HYPREAL_INVERSE_ZERO]));
+ simpset() addsimps [HYPREAL_INVERSE_ZERO]));
by (rtac ccontr 1);
-by (blast_tac (claset() addDs [rename_numerals hypreal_inverse_not_zero]) 1);
+by (blast_tac (claset() addDs [hypreal_inverse_not_zero]) 1);
qed "hypreal_inverse_zero_iff";
Addsimps [hypreal_inverse_zero_iff];
-Goal "(x/y = Numeral0) = (x=Numeral0 | y=(Numeral0::hypreal))";
+Goal "(x/y = 0) = (x=0 | y=(0::hypreal))";
by (auto_tac (claset(), simpset() addsimps [hypreal_divide_def]));
qed "hypreal_divide_eq_0_iff";
Addsimps [hypreal_divide_eq_0_iff];
-Goal "h ~= (Numeral0::hypreal) ==> h/h = Numeral1";
+Goal "h ~= (0::hypreal) ==> h/h = 1";
by (asm_simp_tac
(simpset() addsimps [hypreal_divide_def, hypreal_mult_inverse_left]) 1);
qed "hypreal_divide_self_eq";
@@ -140,7 +145,7 @@
by (ALLGOALS (asm_full_simp_tac (simpset() addsimps [hypreal_mult_commute])));
qed "hypreal_mult_le_mono2_neg";
-Goal "(m*k < n*k) = (((Numeral0::hypreal) < k & m<n) | (k < Numeral0 & n<m))";
+Goal "(m*k < n*k) = (((0::hypreal) < k & m<n) | (k < 0 & n<m))";
by (case_tac "k = (0::hypreal)" 1);
by (auto_tac (claset(),
simpset() addsimps [linorder_neq_iff,
@@ -155,32 +160,32 @@
hypreal_mult_le_mono1_neg]));
qed "hypreal_mult_less_cancel2";
-Goal "(m*k <= n*k) = (((Numeral0::hypreal) < k --> m<=n) & (k < Numeral0 --> n<=m))";
+Goal "(m*k <= n*k) = (((0::hypreal) < k --> m<=n) & (k < 0 --> n<=m))";
by (simp_tac (simpset() addsimps [linorder_not_less RS sym,
hypreal_mult_less_cancel2]) 1);
qed "hypreal_mult_le_cancel2";
-Goal "(k*m < k*n) = (((Numeral0::hypreal) < k & m<n) | (k < Numeral0 & n<m))";
+Goal "(k*m < k*n) = (((0::hypreal) < k & m<n) | (k < 0 & n<m))";
by (simp_tac (simpset() addsimps [inst "z" "k" hypreal_mult_commute,
hypreal_mult_less_cancel2]) 1);
qed "hypreal_mult_less_cancel1";
-Goal "!!k::hypreal. (k*m <= k*n) = ((Numeral0 < k --> m<=n) & (k < Numeral0 --> n<=m))";
+Goal "!!k::hypreal. (k*m <= k*n) = ((0 < k --> m<=n) & (k < 0 --> n<=m))";
by (simp_tac (simpset() addsimps [linorder_not_less RS sym,
hypreal_mult_less_cancel1]) 1);
qed "hypreal_mult_le_cancel1";
-Goal "!!k::hypreal. (k*m = k*n) = (k = Numeral0 | m=n)";
+Goal "!!k::hypreal. (k*m = k*n) = (k = 0 | m=n)";
by (case_tac "k=0" 1);
by (auto_tac (claset(), simpset() addsimps [hypreal_mult_left_cancel]));
qed "hypreal_mult_eq_cancel1";
-Goal "!!k::hypreal. (m*k = n*k) = (k = Numeral0 | m=n)";
+Goal "!!k::hypreal. (m*k = n*k) = (k = 0 | m=n)";
by (case_tac "k=0" 1);
by (auto_tac (claset(), simpset() addsimps [hypreal_mult_right_cancel]));
qed "hypreal_mult_eq_cancel2";
-Goal "!!k::hypreal. k~=Numeral0 ==> (k*m) / (k*n) = (m/n)";
+Goal "!!k::hypreal. k~=0 ==> (k*m) / (k*n) = (m/n)";
by (asm_simp_tac
(simpset() addsimps [hypreal_divide_def, hypreal_inverse_distrib]) 1);
by (subgoal_tac "k * m * (inverse k * inverse n) = \
@@ -190,7 +195,7 @@
qed "hypreal_mult_div_cancel1";
(*For ExtractCommonTerm*)
-Goal "(k*m) / (k*n) = (if k = (Numeral0::hypreal) then Numeral0 else m/n)";
+Goal "(k*m) / (k*n) = (if k = (0::hypreal) then 0 else m/n)";
by (simp_tac (simpset() addsimps [hypreal_mult_div_cancel1]) 1);
qed "hypreal_mult_div_cancel_disj";
@@ -206,14 +211,11 @@
struct
val mk_coeff = mk_coeff
val dest_coeff = dest_coeff 1
- val trans_tac = trans_tac
- val norm_tac = ALLGOALS (simp_tac (HOL_ss addsimps mult_plus_1s))
+ val trans_tac = Real_Numeral_Simprocs.trans_tac
+ val norm_tac =
+ ALLGOALS (simp_tac (HOL_ss addsimps hypreal_minus_from_mult_simps @ mult_1s))
THEN ALLGOALS (simp_tac (HOL_ss addsimps bin_simps@hypreal_mult_minus_simps))
- THEN ALLGOALS
- (simp_tac
- (HOL_ss addsimps [eq_hypreal_number_of, mult_hypreal_number_of,
- hypreal_mult_number_of_left]@
- hypreal_minus_from_mult_simps @ hypreal_mult_ac))
+ THEN ALLGOALS (simp_tac (HOL_ss addsimps hypreal_mult_ac))
val numeral_simp_tac =
ALLGOALS (simp_tac (HOL_ss addsimps rel_hypreal_number_of@bin_simps))
val simplify_meta_eq = simplify_meta_eq
@@ -221,7 +223,8 @@
structure DivCancelNumeralFactor = CancelNumeralFactorFun
(open CancelNumeralFactorCommon
- val prove_conv = prove_conv "hyprealdiv_cancel_numeral_factor"
+ val prove_conv = Real_Numeral_Simprocs.prove_conv
+ "hyprealdiv_cancel_numeral_factor"
val mk_bal = HOLogic.mk_binop "HOL.divide"
val dest_bal = HOLogic.dest_bin "HOL.divide" hyprealT
val cancel = hypreal_mult_div_cancel1 RS trans
@@ -230,7 +233,8 @@
structure EqCancelNumeralFactor = CancelNumeralFactorFun
(open CancelNumeralFactorCommon
- val prove_conv = prove_conv "hyprealeq_cancel_numeral_factor"
+ val prove_conv = Real_Numeral_Simprocs.prove_conv
+ "hyprealeq_cancel_numeral_factor"
val mk_bal = HOLogic.mk_eq
val dest_bal = HOLogic.dest_bin "op =" hyprealT
val cancel = hypreal_mult_eq_cancel1 RS trans
@@ -239,7 +243,8 @@
structure LessCancelNumeralFactor = CancelNumeralFactorFun
(open CancelNumeralFactorCommon
- val prove_conv = prove_conv "hyprealless_cancel_numeral_factor"
+ val prove_conv = Real_Numeral_Simprocs.prove_conv
+ "hyprealless_cancel_numeral_factor"
val mk_bal = HOLogic.mk_binrel "op <"
val dest_bal = HOLogic.dest_bin "op <" hyprealT
val cancel = hypreal_mult_less_cancel1 RS trans
@@ -248,7 +253,8 @@
structure LeCancelNumeralFactor = CancelNumeralFactorFun
(open CancelNumeralFactorCommon
- val prove_conv = prove_conv "hyprealle_cancel_numeral_factor"
+ val prove_conv = Real_Numeral_Simprocs.prove_conv
+ "hyprealle_cancel_numeral_factor"
val mk_bal = HOLogic.mk_binrel "op <="
val dest_bal = HOLogic.dest_bin "op <=" hyprealT
val cancel = hypreal_mult_le_cancel1 RS trans
@@ -288,7 +294,7 @@
set trace_simp;
fun test s = (Goal s; by (Simp_tac 1));
-test "Numeral0 <= (y::hypreal) * -2";
+test "0 <= (y::hypreal) * -2";
test "9*x = 12 * (y::hypreal)";
test "(9*x) / (12 * (y::hypreal)) = z";
test "9*x < 12 * (y::hypreal)";
@@ -332,13 +338,14 @@
val mk_coeff = mk_coeff
val dest_coeff = dest_coeff
val find_first = find_first []
- val trans_tac = trans_tac
+ val trans_tac = Real_Numeral_Simprocs.trans_tac
val norm_tac = ALLGOALS (simp_tac (HOL_ss addsimps mult_1s@hypreal_mult_ac))
end;
structure EqCancelFactor = ExtractCommonTermFun
(open CancelFactorCommon
- val prove_conv = prove_conv "hypreal_eq_cancel_factor"
+ val prove_conv = Real_Numeral_Simprocs.prove_conv
+ "hypreal_eq_cancel_factor"
val mk_bal = HOLogic.mk_eq
val dest_bal = HOLogic.dest_bin "op =" hyprealT
val simplify_meta_eq = cancel_simplify_meta_eq hypreal_mult_eq_cancel1
@@ -347,7 +354,8 @@
structure DivideCancelFactor = ExtractCommonTermFun
(open CancelFactorCommon
- val prove_conv = prove_conv "hypreal_divide_cancel_factor"
+ val prove_conv = Real_Numeral_Simprocs.prove_conv
+ "hypreal_divide_cancel_factor"
val mk_bal = HOLogic.mk_binop "HOL.divide"
val dest_bal = HOLogic.dest_bin "HOL.divide" hyprealT
val simplify_meta_eq = cancel_simplify_meta_eq hypreal_mult_div_cancel_disj
@@ -391,7 +399,7 @@
(*** Simplification of inequalities involving literal divisors ***)
-Goal "Numeral0<z ==> ((x::hypreal) <= y/z) = (x*z <= y)";
+Goal "0<z ==> ((x::hypreal) <= y/z) = (x*z <= y)";
by (subgoal_tac "(x*z <= y) = (x*z <= (y/z)*z)" 1);
by (asm_simp_tac (simpset() addsimps [hypreal_divide_def, hypreal_mult_assoc]) 2);
by (etac ssubst 1);
@@ -400,7 +408,7 @@
qed "pos_hypreal_le_divide_eq";
Addsimps [inst "z" "number_of ?w" pos_hypreal_le_divide_eq];
-Goal "z<Numeral0 ==> ((x::hypreal) <= y/z) = (y <= x*z)";
+Goal "z<0 ==> ((x::hypreal) <= y/z) = (y <= x*z)";
by (subgoal_tac "(y <= x*z) = ((y/z)*z <= x*z)" 1);
by (asm_simp_tac (simpset() addsimps [hypreal_divide_def, hypreal_mult_assoc]) 2);
by (etac ssubst 1);
@@ -409,7 +417,7 @@
qed "neg_hypreal_le_divide_eq";
Addsimps [inst "z" "number_of ?w" neg_hypreal_le_divide_eq];
-Goal "Numeral0<z ==> (y/z <= (x::hypreal)) = (y <= x*z)";
+Goal "0<z ==> (y/z <= (x::hypreal)) = (y <= x*z)";
by (subgoal_tac "(y <= x*z) = ((y/z)*z <= x*z)" 1);
by (asm_simp_tac (simpset() addsimps [hypreal_divide_def, hypreal_mult_assoc]) 2);
by (etac ssubst 1);
@@ -418,7 +426,7 @@
qed "pos_hypreal_divide_le_eq";
Addsimps [inst "z" "number_of ?w" pos_hypreal_divide_le_eq];
-Goal "z<Numeral0 ==> (y/z <= (x::hypreal)) = (x*z <= y)";
+Goal "z<0 ==> (y/z <= (x::hypreal)) = (x*z <= y)";
by (subgoal_tac "(x*z <= y) = (x*z <= (y/z)*z)" 1);
by (asm_simp_tac (simpset() addsimps [hypreal_divide_def, hypreal_mult_assoc]) 2);
by (etac ssubst 1);
@@ -427,7 +435,7 @@
qed "neg_hypreal_divide_le_eq";
Addsimps [inst "z" "number_of ?w" neg_hypreal_divide_le_eq];
-Goal "Numeral0<z ==> ((x::hypreal) < y/z) = (x*z < y)";
+Goal "0<z ==> ((x::hypreal) < y/z) = (x*z < y)";
by (subgoal_tac "(x*z < y) = (x*z < (y/z)*z)" 1);
by (asm_simp_tac (simpset() addsimps [hypreal_divide_def, hypreal_mult_assoc]) 2);
by (etac ssubst 1);
@@ -436,7 +444,7 @@
qed "pos_hypreal_less_divide_eq";
Addsimps [inst "z" "number_of ?w" pos_hypreal_less_divide_eq];
-Goal "z<Numeral0 ==> ((x::hypreal) < y/z) = (y < x*z)";
+Goal "z<0 ==> ((x::hypreal) < y/z) = (y < x*z)";
by (subgoal_tac "(y < x*z) = ((y/z)*z < x*z)" 1);
by (asm_simp_tac (simpset() addsimps [hypreal_divide_def, hypreal_mult_assoc]) 2);
by (etac ssubst 1);
@@ -445,7 +453,7 @@
qed "neg_hypreal_less_divide_eq";
Addsimps [inst "z" "number_of ?w" neg_hypreal_less_divide_eq];
-Goal "Numeral0<z ==> (y/z < (x::hypreal)) = (y < x*z)";
+Goal "0<z ==> (y/z < (x::hypreal)) = (y < x*z)";
by (subgoal_tac "(y < x*z) = ((y/z)*z < x*z)" 1);
by (asm_simp_tac (simpset() addsimps [hypreal_divide_def, hypreal_mult_assoc]) 2);
by (etac ssubst 1);
@@ -454,7 +462,7 @@
qed "pos_hypreal_divide_less_eq";
Addsimps [inst "z" "number_of ?w" pos_hypreal_divide_less_eq];
-Goal "z<Numeral0 ==> (y/z < (x::hypreal)) = (x*z < y)";
+Goal "z<0 ==> (y/z < (x::hypreal)) = (x*z < y)";
by (subgoal_tac "(x*z < y) = (x*z < (y/z)*z)" 1);
by (asm_simp_tac (simpset() addsimps [hypreal_divide_def, hypreal_mult_assoc]) 2);
by (etac ssubst 1);
@@ -463,7 +471,7 @@
qed "neg_hypreal_divide_less_eq";
Addsimps [inst "z" "number_of ?w" neg_hypreal_divide_less_eq];
-Goal "z~=Numeral0 ==> ((x::hypreal) = y/z) = (x*z = y)";
+Goal "z~=0 ==> ((x::hypreal) = y/z) = (x*z = y)";
by (subgoal_tac "(x*z = y) = (x*z = (y/z)*z)" 1);
by (asm_simp_tac (simpset() addsimps [hypreal_divide_def, hypreal_mult_assoc]) 2);
by (etac ssubst 1);
@@ -472,7 +480,7 @@
qed "hypreal_eq_divide_eq";
Addsimps [inst "z" "number_of ?w" hypreal_eq_divide_eq];
-Goal "z~=Numeral0 ==> (y/z = (x::hypreal)) = (y = x*z)";
+Goal "z~=0 ==> (y/z = (x::hypreal)) = (y = x*z)";
by (subgoal_tac "(y = x*z) = ((y/z)*z = x*z)" 1);
by (asm_simp_tac (simpset() addsimps [hypreal_divide_def, hypreal_mult_assoc]) 2);
by (etac ssubst 1);
@@ -481,37 +489,37 @@
qed "hypreal_divide_eq_eq";
Addsimps [inst "z" "number_of ?w" hypreal_divide_eq_eq];
-Goal "(m/k = n/k) = (k = Numeral0 | m = (n::hypreal))";
-by (case_tac "k=Numeral0" 1);
+Goal "(m/k = n/k) = (k = 0 | m = (n::hypreal))";
+by (case_tac "k=0" 1);
by (asm_simp_tac (simpset() addsimps [HYPREAL_DIVIDE_ZERO]) 1);
by (asm_simp_tac (simpset() addsimps [hypreal_divide_eq_eq, hypreal_eq_divide_eq,
hypreal_mult_eq_cancel2]) 1);
qed "hypreal_divide_eq_cancel2";
-Goal "(k/m = k/n) = (k = Numeral0 | m = (n::hypreal))";
-by (case_tac "m=Numeral0 | n = Numeral0" 1);
+Goal "(k/m = k/n) = (k = 0 | m = (n::hypreal))";
+by (case_tac "m=0 | n = 0" 1);
by (auto_tac (claset(),
simpset() addsimps [HYPREAL_DIVIDE_ZERO, hypreal_divide_eq_eq,
hypreal_eq_divide_eq, hypreal_mult_eq_cancel1]));
qed "hypreal_divide_eq_cancel1";
-Goal "[| Numeral0 < r; Numeral0 < x|] ==> (inverse x < inverse (r::hypreal)) = (r < x)";
+Goal "[| 0 < r; 0 < x|] ==> (inverse x < inverse (r::hypreal)) = (r < x)";
by (auto_tac (claset() addIs [hypreal_inverse_less_swap], simpset()));
by (res_inst_tac [("t","r")] (hypreal_inverse_inverse RS subst) 1);
by (res_inst_tac [("t","x")] (hypreal_inverse_inverse RS subst) 1);
by (auto_tac (claset() addIs [hypreal_inverse_less_swap],
simpset() delsimps [hypreal_inverse_inverse]
- addsimps [hypreal_inverse_gt_zero]));
+ addsimps [hypreal_inverse_gt_0]));
qed "hypreal_inverse_less_iff";
-Goal "[| Numeral0 < r; Numeral0 < x|] ==> (inverse x <= inverse r) = (r <= (x::hypreal))";
+Goal "[| 0 < r; 0 < x|] ==> (inverse x <= inverse r) = (r <= (x::hypreal))";
by (asm_simp_tac (simpset() addsimps [linorder_not_less RS sym,
hypreal_inverse_less_iff]) 1);
qed "hypreal_inverse_le_iff";
(** Division by 1, -1 **)
-Goal "(x::hypreal)/Numeral1 = x";
+Goal "(x::hypreal)/1 = x";
by (simp_tac (simpset() addsimps [hypreal_divide_def]) 1);
qed "hypreal_divide_1";
Addsimps [hypreal_divide_1];
@@ -521,12 +529,12 @@
qed "hypreal_divide_minus1";
Addsimps [hypreal_divide_minus1];
-Goal "-1/(x::hypreal) = - (Numeral1/x)";
+Goal "-1/(x::hypreal) = - (1/x)";
by (simp_tac (simpset() addsimps [hypreal_divide_def, hypreal_minus_inverse]) 1);
qed "hypreal_minus1_divide";
Addsimps [hypreal_minus1_divide];
-Goal "[| (Numeral0::hypreal) < d1; Numeral0 < d2 |] ==> EX e. Numeral0 < e & e < d1 & e < d2";
+Goal "[| (0::hypreal) < d1; 0 < d2 |] ==> EX e. 0 < e & e < d1 & e < d2";
by (res_inst_tac [("x","(min d1 d2)/2")] exI 1);
by (asm_simp_tac (simpset() addsimps [min_def]) 1);
qed "hypreal_lbound_gt_zero";
@@ -560,7 +568,7 @@
by Auto_tac;
qed "hypreal_minus_equation";
-Goal "(x + - a = (Numeral0::hypreal)) = (x=a)";
+Goal "(x + - a = (0::hypreal)) = (x=a)";
by (arith_tac 1);
qed "hypreal_add_minus_iff";
Addsimps [hypreal_add_minus_iff];
@@ -587,45 +595,49 @@
Addsimps (map (inst "y" "number_of ?v")
[hypreal_minus_less, hypreal_minus_le, hypreal_minus_equation]);
+Addsimps (map (simplify (simpset()) o inst "x" "1")
+ [hypreal_less_minus, hypreal_le_minus, hypreal_equation_minus]);
+Addsimps (map (simplify (simpset()) o inst "y" "1")
+ [hypreal_minus_less, hypreal_minus_le, hypreal_minus_equation]);
-(*** Simprules combining x+y and Numeral0 ***)
+(*** Simprules combining x+y and 0 ***)
-Goal "(x+y = (Numeral0::hypreal)) = (y = -x)";
+Goal "(x+y = (0::hypreal)) = (y = -x)";
by Auto_tac;
qed "hypreal_add_eq_0_iff";
AddIffs [hypreal_add_eq_0_iff];
-Goal "(x+y < (Numeral0::hypreal)) = (y < -x)";
+Goal "(x+y < (0::hypreal)) = (y < -x)";
by Auto_tac;
qed "hypreal_add_less_0_iff";
AddIffs [hypreal_add_less_0_iff];
-Goal "((Numeral0::hypreal) < x+y) = (-x < y)";
+Goal "((0::hypreal) < x+y) = (-x < y)";
by Auto_tac;
qed "hypreal_0_less_add_iff";
AddIffs [hypreal_0_less_add_iff];
-Goal "(x+y <= (Numeral0::hypreal)) = (y <= -x)";
+Goal "(x+y <= (0::hypreal)) = (y <= -x)";
by Auto_tac;
qed "hypreal_add_le_0_iff";
AddIffs [hypreal_add_le_0_iff];
-Goal "((Numeral0::hypreal) <= x+y) = (-x <= y)";
+Goal "((0::hypreal) <= x+y) = (-x <= y)";
by Auto_tac;
qed "hypreal_0_le_add_iff";
AddIffs [hypreal_0_le_add_iff];
-(** Simprules combining x-y and Numeral0; see also hypreal_less_iff_diff_less_0 etc
+(** Simprules combining x-y and 0; see also hypreal_less_iff_diff_less_0 etc
in HyperBin
**)
-Goal "((Numeral0::hypreal) < x-y) = (y < x)";
+Goal "((0::hypreal) < x-y) = (y < x)";
by Auto_tac;
qed "hypreal_0_less_diff_iff";
AddIffs [hypreal_0_less_diff_iff];
-Goal "((Numeral0::hypreal) <= x-y) = (y <= x)";
+Goal "((0::hypreal) <= x-y) = (y <= x)";
by Auto_tac;
qed "hypreal_0_le_diff_iff";
AddIffs [hypreal_0_le_diff_iff];
@@ -657,7 +669,7 @@
qed "hypreal_dense";
-(*Replaces "inverse #nn" by Numeral1/#nn *)
+(*Replaces "inverse #nn" by 1/#nn *)
Addsimps [inst "x" "number_of ?w" hypreal_inverse_eq_divide];
--- a/src/HOL/Hyperreal/HyperBin.ML Thu Nov 01 21:12:13 2001 +0100
+++ b/src/HOL/Hyperreal/HyperBin.ML Fri Nov 02 17:55:24 2001 +0100
@@ -13,13 +13,13 @@
qed "hypreal_number_of";
Addsimps [hypreal_number_of];
-Goalw [hypreal_number_of_def] "(0::hypreal) = Numeral0";
-by (simp_tac (simpset() addsimps [hypreal_of_real_zero RS sym]) 1);
-qed "zero_eq_numeral_0";
+Goalw [hypreal_number_of_def] "Numeral0 = (0::hypreal)";
+by (Simp_tac 1);
+qed "hypreal_numeral_0_eq_0";
-Goalw [hypreal_number_of_def] "(1::hypreal) = Numeral1";
-by (simp_tac (simpset() addsimps [hypreal_of_real_one RS sym]) 1);
-qed "one_eq_numeral_1";
+Goalw [hypreal_number_of_def] "Numeral1 = (1::hypreal)";
+by (Simp_tac 1);
+qed "hypreal_numeral_1_eq_1";
(** Addition **)
@@ -57,15 +57,15 @@
qed "mult_hypreal_number_of";
Addsimps [mult_hypreal_number_of];
-Goal "(2::hypreal) = Numeral1 + Numeral1";
-by (Simp_tac 1);
+Goal "(2::hypreal) = 1 + 1";
+by (simp_tac (simpset() addsimps [hypreal_numeral_1_eq_1 RS sym]) 1);
val lemma = result();
(*For specialist use: NOT as default simprules*)
Goal "2 * z = (z+z::hypreal)";
by (simp_tac (simpset ()
addsimps [lemma, hypreal_add_mult_distrib,
- one_eq_numeral_1 RS sym]) 1);
+ hypreal_numeral_1_eq_1]) 1);
qed "hypreal_mult_2";
Goal "z * 2 = (z+z::hypreal)";
@@ -107,46 +107,20 @@
(*** New versions of existing theorems involving 0, 1 ***)
-Goal "- Numeral1 = (-1::hypreal)";
-by (Simp_tac 1);
-qed "minus_numeral_one";
+Goal "- 1 = (-1::hypreal)";
+by (simp_tac (simpset() addsimps [hypreal_numeral_1_eq_1 RS sym]) 1);
+qed "hypreal_minus_1_eq_m1";
-(*Maps 0 to Numeral0 and (1::hypreal) to Numeral1 and -(Numeral1) to -1*)
+(*Maps 0 to Numeral0 and 1 to Numeral1 and -(Numeral1) to -1*)
val hypreal_numeral_ss =
- real_numeral_ss addsimps [zero_eq_numeral_0, one_eq_numeral_1,
- minus_numeral_one];
+ real_numeral_ss addsimps [hypreal_numeral_0_eq_0 RS sym,
+ hypreal_numeral_1_eq_1 RS sym,
+ hypreal_minus_1_eq_m1];
fun rename_numerals th =
asm_full_simplify hypreal_numeral_ss (Thm.transfer (the_context ()) th);
-(*Now insert some identities previously stated for 0 and 1*)
-
-(** HyperDef **)
-
-Addsimps (map rename_numerals
- [hypreal_minus_zero, hypreal_minus_zero_iff,
- hypreal_add_zero_left, hypreal_add_zero_right,
- hypreal_diff_zero, hypreal_diff_zero_right,
- hypreal_mult_0_right, hypreal_mult_0,
- hypreal_mult_1_right, hypreal_mult_1,
- hypreal_inverse_1, hypreal_minus_zero_less_iff]);
-
-bind_thm ("hypreal_0_less_mult_iff",
- rename_numerals hypreal_zero_less_mult_iff);
-bind_thm ("hypreal_0_le_mult_iff",
- rename_numerals hypreal_zero_le_mult_iff);
-bind_thm ("hypreal_mult_less_0_iff",
- rename_numerals hypreal_mult_less_zero_iff);
-bind_thm ("hypreal_mult_le_0_iff",
- rename_numerals hypreal_mult_le_zero_iff);
-
-bind_thm ("hypreal_inverse_less_0", rename_numerals hypreal_inverse_less_zero);
-bind_thm ("hypreal_inverse_gt_0", rename_numerals hypreal_inverse_gt_zero);
-
-Addsimps [zero_eq_numeral_0,one_eq_numeral_1];
-
-
(** Simplification of arithmetic when nested to the right **)
Goal "number_of v + (number_of w + z) = (number_of(bin_add v w) + z::hypreal)";
@@ -176,15 +150,15 @@
(** Combining of literal coefficients in sums of products **)
-Goal "(x < y) = (x-y < (Numeral0::hypreal))";
+Goal "(x < y) = (x-y < (0::hypreal))";
by (simp_tac (simpset() addsimps [hypreal_diff_less_eq]) 1);
qed "hypreal_less_iff_diff_less_0";
-Goal "(x = y) = (x-y = (Numeral0::hypreal))";
+Goal "(x = y) = (x-y = (0::hypreal))";
by (simp_tac (simpset() addsimps [hypreal_diff_eq_eq]) 1);
qed "hypreal_eq_iff_diff_eq_0";
-Goal "(x <= y) = (x-y <= (Numeral0::hypreal))";
+Goal "(x <= y) = (x-y <= (0::hypreal))";
by (simp_tac (simpset() addsimps [hypreal_diff_le_eq]) 1);
qed "hypreal_le_iff_diff_le_0";
@@ -209,7 +183,7 @@
Goal "!!i::hypreal. (i*u + m = j*u + n) = ((i-j)*u + m = n)";
by (asm_simp_tac
(simpset() addsimps [hypreal_diff_def, hypreal_add_mult_distrib]@
- hypreal_add_ac@rel_iff_rel_0_rls) 1);
+ hypreal_add_ac@rel_iff_rel_0_rls) 1);
qed "hypreal_eq_add_iff1";
Goal "!!i::hypreal. (i*u + m = j*u + n) = (m = (j-i)*u + n)";
@@ -242,26 +216,31 @@
hypreal_add_ac@rel_iff_rel_0_rls) 1);
qed "hypreal_le_add_iff2";
+Goal "-1 * (z::hypreal) = -z";
+by (simp_tac (simpset() addsimps [hypreal_minus_1_eq_m1 RS sym,
+ hypreal_mult_minus_1]) 1);
+qed "hypreal_mult_minus1";
+Addsimps [hypreal_mult_minus1];
+
Goal "(z::hypreal) * -1 = -z";
-by (stac (minus_numeral_one RS sym) 1);
-by (stac (hypreal_minus_mult_eq2 RS sym) 1);
-by Auto_tac;
-qed "hypreal_mult_minus_1_right";
-Addsimps [hypreal_mult_minus_1_right];
+by (stac hypreal_mult_commute 1 THEN rtac hypreal_mult_minus1 1);
+qed "hypreal_mult_minus1_right";
+Addsimps [hypreal_mult_minus1_right];
-Goal "-1 * (z::hypreal) = -z";
-by (simp_tac (simpset() addsimps [hypreal_mult_commute]) 1);
-qed "hypreal_mult_minus_1";
-Addsimps [hypreal_mult_minus_1];
+Addsimps [hypreal_numeral_0_eq_0, hypreal_numeral_1_eq_1];
structure Hyperreal_Numeral_Simprocs =
struct
+(*Maps 0 to Numeral0 and 1 to Numeral1 so that arithmetic in simprocs
+ isn't complicated by the abstract 0 and 1.*)
+val numeral_syms = [hypreal_numeral_0_eq_0 RS sym,
+ hypreal_numeral_1_eq_1 RS sym];
+
(*Utilities*)
-
val hyprealT = Type("HyperDef.hypreal",[]);
fun mk_numeral n = HOLogic.number_of_const hyprealT $ HOLogic.mk_bin n;
@@ -271,11 +250,11 @@
val find_first_numeral = Real_Numeral_Simprocs.find_first_numeral;
val zero = mk_numeral 0;
-val mk_plus = HOLogic.mk_binop "op +";
+val mk_plus = Real_Numeral_Simprocs.mk_plus;
val uminus_const = Const ("uminus", hyprealT --> hyprealT);
-(*Thus mk_sum[t] yields t+Numeral0; longer sums don't have a trailing zero*)
+(*Thus mk_sum[t] yields t+0; longer sums don't have a trailing zero*)
fun mk_sum [] = zero
| mk_sum [t,u] = mk_plus (t, u)
| mk_sum (t :: ts) = mk_plus (t, mk_sum ts);
@@ -335,24 +314,23 @@
handle TERM _ => find_first_coeff (t::past) u terms;
-(*Simplify Numeral1*n and n*Numeral1 to n*)
+(*Simplify Numeral0+n, n+Numeral0, Numeral1*n, n*Numeral1*)
val add_0s = map rename_numerals
[hypreal_add_zero_left, hypreal_add_zero_right];
-val mult_plus_1s = map rename_numerals
- [hypreal_mult_1, hypreal_mult_1_right];
-val mult_minus_1s = map rename_numerals
- [hypreal_mult_minus_1, hypreal_mult_minus_1_right];
-val mult_1s = mult_plus_1s @ mult_minus_1s;
+val mult_1s = map rename_numerals [hypreal_mult_1, hypreal_mult_1_right] @
+ [hypreal_mult_minus1, hypreal_mult_minus1_right];
(*To perform binary arithmetic*)
val bin_simps =
- [add_hypreal_number_of, hypreal_add_number_of_left,
- minus_hypreal_number_of, diff_hypreal_number_of, mult_hypreal_number_of,
+ [hypreal_numeral_0_eq_0 RS sym, hypreal_numeral_1_eq_1 RS sym,
+ add_hypreal_number_of, hypreal_add_number_of_left,
+ minus_hypreal_number_of,
+ diff_hypreal_number_of, mult_hypreal_number_of,
hypreal_mult_number_of_left] @ bin_arith_simps @ bin_rel_simps;
(*To evaluate binary negations of coefficients*)
val hypreal_minus_simps = NCons_simps @
- [minus_hypreal_number_of,
+ [hypreal_minus_1_eq_m1, minus_hypreal_number_of,
bin_minus_1, bin_minus_0, bin_minus_Pls, bin_minus_Min,
bin_pred_1, bin_pred_0, bin_pred_Pls, bin_pred_Min];
@@ -374,24 +352,6 @@
val hypreal_mult_minus_simps =
[hypreal_mult_assoc, hypreal_minus_mult_eq1, hypreal_minus_mult_eq_1_to_2];
-(*Apply the given rewrite (if present) just once*)
-fun trans_tac None = all_tac
- | trans_tac (Some th) = ALLGOALS (rtac (th RS trans));
-
-fun prove_conv name tacs sg (hyps: thm list) (t,u) =
- if t aconv u then None
- else
- let val ct = cterm_of sg (HOLogic.mk_Trueprop (HOLogic.mk_eq (t, u)))
- in Some
- (prove_goalw_cterm [] ct (K tacs)
- handle ERROR => error
- ("The error(s) above occurred while trying to prove " ^
- string_of_cterm ct ^ "\nInternal failure of simproc " ^ name))
- end;
-
-(*version without the hyps argument*)
-fun prove_conv_nohyps name tacs sg = prove_conv name tacs sg [];
-
(*Final simplification: cancel + and * *)
val simplify_meta_eq =
Int_Numeral_Simprocs.simplify_meta_eq
@@ -409,7 +369,7 @@
val mk_coeff = mk_coeff
val dest_coeff = dest_coeff 1
val find_first_coeff = find_first_coeff []
- val trans_tac = trans_tac
+ val trans_tac = Real_Numeral_Simprocs.trans_tac
val norm_tac =
ALLGOALS (simp_tac (HOL_ss addsimps add_0s@mult_1s@diff_simps@
hypreal_minus_simps@hypreal_add_ac))
@@ -424,7 +384,7 @@
structure EqCancelNumerals = CancelNumeralsFun
(open CancelNumeralsCommon
- val prove_conv = prove_conv "hyprealeq_cancel_numerals"
+ val prove_conv = Real_Numeral_Simprocs.prove_conv "hyprealeq_cancel_numerals"
val mk_bal = HOLogic.mk_eq
val dest_bal = HOLogic.dest_bin "op =" hyprealT
val bal_add1 = hypreal_eq_add_iff1 RS trans
@@ -433,7 +393,7 @@
structure LessCancelNumerals = CancelNumeralsFun
(open CancelNumeralsCommon
- val prove_conv = prove_conv "hyprealless_cancel_numerals"
+ val prove_conv = Real_Numeral_Simprocs.prove_conv "hyprealless_cancel_numerals"
val mk_bal = HOLogic.mk_binrel "op <"
val dest_bal = HOLogic.dest_bin "op <" hyprealT
val bal_add1 = hypreal_less_add_iff1 RS trans
@@ -442,7 +402,7 @@
structure LeCancelNumerals = CancelNumeralsFun
(open CancelNumeralsCommon
- val prove_conv = prove_conv "hyprealle_cancel_numerals"
+ val prove_conv = Real_Numeral_Simprocs.prove_conv "hyprealle_cancel_numerals"
val mk_bal = HOLogic.mk_binrel "op <="
val dest_bal = HOLogic.dest_bin "op <=" hyprealT
val bal_add1 = hypreal_le_add_iff1 RS trans
@@ -476,8 +436,9 @@
val mk_coeff = mk_coeff
val dest_coeff = dest_coeff 1
val left_distrib = left_hypreal_add_mult_distrib RS trans
- val prove_conv = prove_conv_nohyps "hypreal_combine_numerals"
- val trans_tac = trans_tac
+ val prove_conv = Real_Numeral_Simprocs.prove_conv_nohyps
+ "hypreal_combine_numerals"
+ val trans_tac = Real_Numeral_Simprocs.trans_tac
val norm_tac =
ALLGOALS (simp_tac (HOL_ss addsimps add_0s@mult_1s@diff_simps@
hypreal_minus_simps@hypreal_add_ac))
@@ -516,8 +477,52 @@
[hypreal_mult_1, hypreal_mult_1_right]
(([th, cancel_th]) MRS trans);
+(*** Making constant folding work for 0 and 1 too ***)
+
+structure HyperrealAbstractNumeralsData =
+ struct
+ val dest_eq = HOLogic.dest_eq o HOLogic.dest_Trueprop o concl_of
+ val is_numeral = Bin_Simprocs.is_numeral
+ val numeral_0_eq_0 = hypreal_numeral_0_eq_0
+ val numeral_1_eq_1 = hypreal_numeral_1_eq_1
+ val prove_conv = Real_Numeral_Simprocs.prove_conv_nohyps
+ "hypreal_abstract_numerals"
+ fun norm_tac simps = ALLGOALS (simp_tac (HOL_ss addsimps simps))
+ val simplify_meta_eq = Bin_Simprocs.simplify_meta_eq
+ end
+
+structure HyperrealAbstractNumerals =
+ AbstractNumeralsFun (HyperrealAbstractNumeralsData)
+
+(*For addition, we already have rules for the operand 0.
+ Multiplication is omitted because there are already special rules for
+ both 0 and 1 as operands. Unary minus is trivial, just have - 1 = -1.
+ For the others, having three patterns is a compromise between just having
+ one (many spurious calls) and having nine (just too many!) *)
+val eval_numerals =
+ map prep_simproc
+ [("hypreal_add_eval_numerals",
+ prep_pats ["(m::hypreal) + 1", "(m::hypreal) + number_of v"],
+ HyperrealAbstractNumerals.proc add_hypreal_number_of),
+ ("hypreal_diff_eval_numerals",
+ prep_pats ["(m::hypreal) - 1", "(m::hypreal) - number_of v"],
+ HyperrealAbstractNumerals.proc diff_hypreal_number_of),
+ ("hypreal_eq_eval_numerals",
+ prep_pats ["(m::hypreal) = 0", "(m::hypreal) = 1",
+ "(m::hypreal) = number_of v"],
+ HyperrealAbstractNumerals.proc eq_hypreal_number_of),
+ ("hypreal_less_eval_numerals",
+ prep_pats ["(m::hypreal) < 0", "(m::hypreal) < 1",
+ "(m::hypreal) < number_of v"],
+ HyperrealAbstractNumerals.proc less_hypreal_number_of),
+ ("hypreal_le_eval_numerals",
+ prep_pats ["(m::hypreal) <= 0", "(m::hypreal) <= 1",
+ "(m::hypreal) <= number_of v"],
+ HyperrealAbstractNumerals.proc le_hypreal_number_of_eq_not_less)]
+
end;
+Addsimprocs Hyperreal_Numeral_Simprocs.eval_numerals;
Addsimprocs Hyperreal_Numeral_Simprocs.cancel_numerals;
Addsimprocs [Hyperreal_Numeral_Simprocs.combine_numerals];
@@ -581,8 +586,6 @@
Addsimprocs [Hyperreal_Times_Assoc.conv];
-Addsimps [rename_numerals hypreal_of_real_zero_iff];
-
(*Simplification of x-y < 0, etc.*)
AddIffs [hypreal_less_iff_diff_less_0 RS sym];
AddIffs [hypreal_eq_iff_diff_eq_0 RS sym];
@@ -603,6 +606,12 @@
qed "number_of_le_hypreal_of_real_iff";
Addsimps [number_of_le_hypreal_of_real_iff];
+Goal "(hypreal_of_real z = number_of w) = (z = number_of w)";
+by (stac (hypreal_of_real_eq_iff RS sym) 1);
+by (Simp_tac 1);
+qed "hypreal_of_real_eq_number_of_iff";
+Addsimps [hypreal_of_real_eq_number_of_iff];
+
Goal "(hypreal_of_real z < number_of w) = (z < number_of w)";
by (stac (hypreal_of_real_less_iff RS sym) 1);
by (Simp_tac 1);
@@ -615,6 +624,20 @@
qed "hypreal_of_real_le_number_of_iff";
Addsimps [hypreal_of_real_le_number_of_iff];
+(*As above, for the special case of zero*)
+Addsimps
+ (map (simplify (simpset()) o inst "w" "Pls")
+ [hypreal_of_real_eq_number_of_iff,
+ hypreal_of_real_le_number_of_iff, hypreal_of_real_less_number_of_iff,
+ number_of_le_hypreal_of_real_iff, number_of_less_hypreal_of_real_iff]);
+
+(*As above, for the special case of one*)
+Addsimps
+ (map (simplify (simpset()) o inst "w" "Pls BIT True")
+ [hypreal_of_real_eq_number_of_iff,
+ hypreal_of_real_le_number_of_iff, hypreal_of_real_less_number_of_iff,
+ number_of_le_hypreal_of_real_iff, number_of_less_hypreal_of_real_iff]);
+
(** <= monotonicity results: needed for arithmetic **)
Goal "[| i <= j; (0::hypreal) <= k |] ==> i*k <= j*k";
@@ -634,3 +657,4 @@
by (assume_tac 1);
qed "hypreal_mult_le_mono";
+Addsimps [hypreal_minus_1_eq_m1];
--- a/src/HOL/Hyperreal/HyperDef.ML Thu Nov 01 21:12:13 2001 +0100
+++ b/src/HOL/Hyperreal/HyperDef.ML Fri Nov 02 17:55:24 2001 +0100
@@ -311,11 +311,8 @@
Goal "(-x = 0) = (x = (0::hypreal))";
by (res_inst_tac [("z","x")] eq_Abs_hypreal 1);
-by (auto_tac (claset(),
- simpset() addsimps [hypreal_zero_def, hypreal_minus, eq_commute] @
- real_add_ac));
+by (auto_tac (claset(), simpset() addsimps [hypreal_zero_def, hypreal_minus]));
qed "hypreal_minus_zero_iff";
-
Addsimps [hypreal_minus_zero_iff];
@@ -430,59 +427,17 @@
by (simp_tac (simpset() addsimps [hypreal_add_commute]) 1);
qed "hypreal_minus_distrib1";
-Goal "(x + - (y::hypreal)) + (y + - z) = x + -z";
-by (res_inst_tac [("w1","y")] (hypreal_add_commute RS subst) 1);
-by (simp_tac (simpset() addsimps [hypreal_add_left_commute,
- hypreal_add_assoc]) 1);
-by (simp_tac (simpset() addsimps [hypreal_add_commute]) 1);
-qed "hypreal_add_minus_cancel1";
-
Goal "((x::hypreal) + y = x + z) = (y = z)";
by (Step_tac 1);
by (dres_inst_tac [("f","%t.-x + t")] arg_cong 1);
by (asm_full_simp_tac (simpset() addsimps [hypreal_add_assoc RS sym]) 1);
qed "hypreal_add_left_cancel";
-Goal "z + (x + (y + -z)) = x + (y::hypreal)";
-by (simp_tac (simpset() addsimps hypreal_add_ac) 1);
-qed "hypreal_add_minus_cancel2";
-Addsimps [hypreal_add_minus_cancel2];
-
-Goal "y + -(x + y) = -(x::hypreal)";
-by (Full_simp_tac 1);
-by (rtac (hypreal_add_left_commute RS subst) 1);
-by (Full_simp_tac 1);
-qed "hypreal_add_minus_cancel";
-Addsimps [hypreal_add_minus_cancel];
-
-Goal "y + -(y + x) = -(x::hypreal)";
-by (simp_tac (simpset() addsimps [hypreal_add_assoc RS sym]) 1);
-qed "hypreal_add_minus_cancelc";
-Addsimps [hypreal_add_minus_cancelc];
-
-Goal "(z + -x) + (y + -z) = (y + -(x::hypreal))";
-by (full_simp_tac
- (simpset() addsimps [hypreal_minus_add_distrib RS sym,
- hypreal_add_left_cancel] @ hypreal_add_ac
- delsimps [hypreal_minus_add_distrib]) 1);
-qed "hypreal_add_minus_cancel3";
-Addsimps [hypreal_add_minus_cancel3];
-
Goal "(y + (x::hypreal)= z + x) = (y = z)";
by (simp_tac (simpset() addsimps [hypreal_add_commute,
hypreal_add_left_cancel]) 1);
qed "hypreal_add_right_cancel";
-Goal "z + (y + -z) = (y::hypreal)";
-by (simp_tac (simpset() addsimps hypreal_add_ac) 1);
-qed "hypreal_add_minus_cancel4";
-Addsimps [hypreal_add_minus_cancel4];
-
-Goal "z + (w + (x + (-z + y))) = w + x + (y::hypreal)";
-by (simp_tac (simpset() addsimps hypreal_add_ac) 1);
-qed "hypreal_add_minus_cancel5";
-Addsimps [hypreal_add_minus_cancel5];
-
Goal "z + ((- z) + w) = (w::hypreal)";
by (simp_tac (simpset() addsimps [hypreal_add_assoc RS sym]) 1);
qed "hypreal_add_minus_cancelA";
@@ -535,22 +490,24 @@
by (res_inst_tac [("z","z")] eq_Abs_hypreal 1);
by (asm_full_simp_tac (simpset() addsimps [hypreal_mult]) 1);
qed "hypreal_mult_1";
+Addsimps [hypreal_mult_1];
Goal "z * (1::hypreal) = z";
by (simp_tac (simpset() addsimps [hypreal_mult_commute,
hypreal_mult_1]) 1);
qed "hypreal_mult_1_right";
+Addsimps [hypreal_mult_1_right];
Goalw [hypreal_zero_def] "0 * z = (0::hypreal)";
by (res_inst_tac [("z","z")] eq_Abs_hypreal 1);
-by (asm_full_simp_tac (simpset() addsimps [hypreal_mult,real_mult_0]) 1);
+by (asm_full_simp_tac (simpset() addsimps [hypreal_mult]) 1);
qed "hypreal_mult_0";
+Addsimps [hypreal_mult_0];
Goal "z * 0 = (0::hypreal)";
-by (simp_tac (simpset() addsimps [hypreal_mult_commute, hypreal_mult_0]) 1);
+by (simp_tac (simpset() addsimps [hypreal_mult_commute]) 1);
qed "hypreal_mult_0_right";
-
-Addsimps [hypreal_mult_0,hypreal_mult_0_right];
+Addsimps [hypreal_mult_0_right];
Goal "-(x * y) = -x * (y::hypreal)";
by (res_inst_tac [("z","x")] eq_Abs_hypreal 1);
@@ -570,7 +527,18 @@
(*Pull negations out*)
Addsimps [hypreal_minus_mult_eq2 RS sym, hypreal_minus_mult_eq1 RS sym];
-Goal "-x*y = (x::hypreal)*-y";
+Goal "(- (1::hypreal)) * z = -z";
+by (Simp_tac 1);
+qed "hypreal_mult_minus_1";
+Addsimps [hypreal_mult_minus_1];
+
+Goal "z * (- (1::hypreal)) = -z";
+by (stac hypreal_mult_commute 1);
+by (Simp_tac 1);
+qed "hypreal_mult_minus_1_right";
+Addsimps [hypreal_mult_minus_1_right];
+
+Goal "(-x) * y = (x::hypreal) * -y";
by Auto_tac;
qed "hypreal_minus_mult_commute";
@@ -599,10 +567,6 @@
by (simp_tac (simpset() addsimps [hypreal_mult_commute',hypreal_add_mult_distrib]) 1);
qed "hypreal_add_mult_distrib2";
-bind_thms ("hypreal_mult_simps", [hypreal_mult_1, hypreal_mult_1_right]);
-Addsimps hypreal_mult_simps;
-
-(* 07/00 *)
Goalw [hypreal_diff_def] "((z1::hypreal) - z2) * w = (z1 * w) - (z2 * w)";
by (simp_tac (simpset() addsimps [hypreal_add_mult_distrib]) 1);
@@ -622,13 +586,13 @@
(**** multiplicative inverse on hypreal ****)
Goalw [congruent_def]
- "congruent hyprel (%X. hyprel``{%n. if X n = Numeral0 then Numeral0 else inverse(X n)})";
+ "congruent hyprel (%X. hyprel``{%n. if X n = 0 then 0 else inverse(X n)})";
by (Auto_tac THEN Ultra_tac 1);
qed "hypreal_inverse_congruent";
Goalw [hypreal_inverse_def]
"inverse (Abs_hypreal(hyprel``{%n. X n})) = \
-\ Abs_hypreal(hyprel `` {%n. if X n = Numeral0 then Numeral0 else inverse(X n)})";
+\ Abs_hypreal(hyprel `` {%n. if X n = 0 then 0 else inverse(X n)})";
by (res_inst_tac [("f","Abs_hypreal")] arg_cong 1);
by (simp_tac (simpset() addsimps
[hyprel_in_hypreal RS Abs_hypreal_inverse,
@@ -807,41 +771,12 @@
by (Ultra_tac 1);
qed "hypreal_less";
-(*---------------------------------------------------------------------------------
- Hyperreals as a linearly ordered field
- ---------------------------------------------------------------------------------*)
-(*** sum order
-Goalw [hypreal_zero_def]
- "[| 0 < x; 0 < y |] ==> (0::hypreal) < x + y";
-by (res_inst_tac [("z","x")] eq_Abs_hypreal 1);
-by (res_inst_tac [("z","y")] eq_Abs_hypreal 1);
-by (auto_tac (claset(),simpset() addsimps
- [hypreal_less_def,hypreal_add]));
-by (auto_tac (claset() addSIs [exI],simpset() addsimps
- [hypreal_less_def,hypreal_add]));
-by (ultra_tac (claset() addIs [real_add_order],simpset()) 1);
-qed "hypreal_add_order";
-***)
+(*----------------------------------------------------------------------------
+ Trichotomy: the hyperreals are linearly ordered
+ ---------------------------------------------------------------------------*)
-(*** mult order
-Goalw [hypreal_zero_def]
- "[| 0 < x; 0 < y |] ==> (0::hypreal) < x * y";
-by (res_inst_tac [("z","x")] eq_Abs_hypreal 1);
-by (res_inst_tac [("z","y")] eq_Abs_hypreal 1);
-by (auto_tac (claset() addSIs [exI],simpset() addsimps
- [hypreal_less_def,hypreal_mult]));
-by (ultra_tac (claset() addIs [rename_numerals real_mult_order],
- simpset()) 1);
-qed "hypreal_mult_order";
-****)
-
-
-(*---------------------------------------------------------------------------------
- Trichotomy of the hyperreals
- --------------------------------------------------------------------------------*)
-
-Goalw [hyprel_def] "EX x. x: hyprel `` {%n. Numeral0}";
-by (res_inst_tac [("x","%n. Numeral0")] exI 1);
+Goalw [hyprel_def] "EX x. x: hyprel `` {%n. 0}";
+by (res_inst_tac [("x","%n. 0")] exI 1);
by (Step_tac 1);
by (auto_tac (claset() addSIs [FreeUltrafilterNat_Nat_set], simpset()));
qed "lemma_hyprel_0_mem";
@@ -924,12 +859,6 @@
simpset()));
qed "hypreal_not_eq_minus_iff";
-Goal "(x+y = (0::hypreal)) = (x = -y)";
-by (stac hypreal_eq_minus_iff 1);
-by Auto_tac;
-qed "hypreal_add_eq_0_iff";
-AddIffs [hypreal_add_eq_0_iff];
-
(*** linearity ***)
@@ -1048,12 +977,18 @@
simpset() addsimps [hypreal_zero_def, hypreal_less,hypreal_minus]));
by (ALLGOALS(Ultra_tac));
qed "hypreal_minus_zero_less_iff2";
+Addsimps [hypreal_minus_zero_less_iff2];
-Goalw [hypreal_le_def] "((0::hypreal) <= -r) = (r <= (0::hypreal))";
+Goalw [hypreal_le_def] "((0::hypreal) <= -r) = (r <= 0)";
by (simp_tac (simpset() addsimps [hypreal_minus_zero_less_iff2]) 1);
qed "hypreal_minus_zero_le_iff";
Addsimps [hypreal_minus_zero_le_iff];
+Goalw [hypreal_le_def] "(-r <= (0::hypreal)) = (0 <= r)";
+by (simp_tac (simpset() addsimps [hypreal_minus_zero_less_iff2]) 1);
+qed "hypreal_minus_zero_le_iff2";
+Addsimps [hypreal_minus_zero_le_iff2];
+
(*----------------------------------------------------------
hypreal_of_real preserves field and order properties
-----------------------------------------------------------*)
@@ -1098,32 +1033,30 @@
qed "hypreal_of_real_minus";
Addsimps [hypreal_of_real_minus];
-(*DON'T insert this or the next one as default simprules.
- They are used in both orientations and anyway aren't the ones we finally
- need, which would use binary literals.*)
-Goalw [hypreal_of_real_def,hypreal_one_def] "hypreal_of_real Numeral1 = (1::hypreal)";
-by (Step_tac 1);
+Goalw [hypreal_of_real_def,hypreal_one_def] "hypreal_of_real 1 = (1::hypreal)";
+by (Simp_tac 1);
qed "hypreal_of_real_one";
+Addsimps [hypreal_of_real_one];
-Goalw [hypreal_of_real_def,hypreal_zero_def] "hypreal_of_real Numeral0 = 0";
-by (Step_tac 1);
+Goalw [hypreal_of_real_def,hypreal_zero_def] "hypreal_of_real 0 = 0";
+by (Simp_tac 1);
qed "hypreal_of_real_zero";
+Addsimps [hypreal_of_real_zero];
-Goal "(hypreal_of_real r = 0) = (r = Numeral0)";
+Goal "(hypreal_of_real r = 0) = (r = 0)";
by (auto_tac (claset() addIs [FreeUltrafilterNat_P],
simpset() addsimps [hypreal_of_real_def,
hypreal_zero_def,FreeUltrafilterNat_Nat_set]));
qed "hypreal_of_real_zero_iff";
Goal "hypreal_of_real (inverse r) = inverse (hypreal_of_real r)";
-by (case_tac "r=Numeral0" 1);
+by (case_tac "r=0" 1);
by (asm_simp_tac (simpset() addsimps [REAL_DIVIDE_ZERO, INVERSE_ZERO,
- HYPREAL_INVERSE_ZERO, hypreal_of_real_zero]) 1);
+ HYPREAL_INVERSE_ZERO]) 1);
by (res_inst_tac [("c1","hypreal_of_real r")]
(hypreal_mult_left_cancel RS iffD1) 1);
by (stac (hypreal_of_real_mult RS sym) 2);
-by (auto_tac (claset(),
- simpset() addsimps [hypreal_of_real_one, hypreal_of_real_zero_iff]));
+by (auto_tac (claset(), simpset() addsimps [hypreal_of_real_zero_iff]));
qed "hypreal_of_real_inverse";
Addsimps [hypreal_of_real_inverse];
--- a/src/HOL/Hyperreal/HyperDef.thy Thu Nov 01 21:12:13 2001 +0100
+++ b/src/HOL/Hyperreal/HyperDef.thy Fri Nov 02 17:55:24 2001 +0100
@@ -30,10 +30,10 @@
defs
hypreal_zero_def
- "0 == Abs_hypreal(hyprel``{%n::nat. (Numeral0::real)})"
+ "0 == Abs_hypreal(hyprel``{%n::nat. (0::real)})"
hypreal_one_def
- "1 == Abs_hypreal(hyprel``{%n::nat. (Numeral1::real)})"
+ "1 == Abs_hypreal(hyprel``{%n::nat. (1::real)})"
hypreal_minus_def
"- P == Abs_hypreal(UN X: Rep_hypreal(P). hyprel``{%n::nat. - (X n)})"
@@ -43,7 +43,7 @@
hypreal_inverse_def
"inverse P == Abs_hypreal(UN X: Rep_hypreal(P).
- hyprel``{%n. if X n = Numeral0 then Numeral0 else inverse (X n)})"
+ hyprel``{%n. if X n = 0 then 0 else inverse (X n)})"
hypreal_divide_def
"P / Q::hypreal == P * inverse Q"
--- a/src/HOL/Hyperreal/HyperNat.ML Thu Nov 01 21:12:13 2001 +0100
+++ b/src/HOL/Hyperreal/HyperNat.ML Fri Nov 02 17:55:24 2001 +0100
@@ -1246,17 +1246,13 @@
Addsimps [hypnat_of_nat_eq_cancel];
Goalw [hypnat_zero_def]
- "hypreal_of_hypnat 0 = Numeral0";
-by (simp_tac (HOL_ss addsimps
- [zero_eq_numeral_0 RS sym, hypreal_zero_def]) 1);
-by (simp_tac (simpset() addsimps [hypreal_of_hypnat, real_of_nat_zero]) 1);
+ "hypreal_of_hypnat 0 = 0";
+by (simp_tac (simpset() addsimps [hypreal_zero_def, hypreal_of_hypnat]) 1);
qed "hypreal_of_hypnat_zero";
Goalw [hypnat_one_def]
- "hypreal_of_hypnat (1::hypnat) = Numeral1";
-by (simp_tac (HOL_ss addsimps
- [one_eq_numeral_1 RS sym, hypreal_one_def]) 1);
-by (simp_tac (simpset() addsimps [hypreal_of_hypnat, real_of_nat_one]) 1);
+ "hypreal_of_hypnat (1::hypnat) = 1";
+by (simp_tac (simpset() addsimps [hypreal_one_def, hypreal_of_hypnat]) 1);
qed "hypreal_of_hypnat_one";
Goal "hypreal_of_hypnat (m + n) = hypreal_of_hypnat m + hypreal_of_hypnat n";
@@ -1283,7 +1279,7 @@
qed "hypreal_of_hypnat_less_iff";
Addsimps [hypreal_of_hypnat_less_iff];
-Goal "(hypreal_of_hypnat N = Numeral0) = (N = 0)";
+Goal "(hypreal_of_hypnat N = 0) = (N = 0)";
by (simp_tac (simpset() addsimps [hypreal_of_hypnat_zero RS sym]) 1);
qed "hypreal_of_hypnat_eq_zero_iff";
Addsimps [hypreal_of_hypnat_eq_zero_iff];
--- a/src/HOL/Hyperreal/HyperOrd.ML Thu Nov 01 21:12:13 2001 +0100
+++ b/src/HOL/Hyperreal/HyperOrd.ML Fri Nov 02 17:55:24 2001 +0100
@@ -101,16 +101,7 @@
by (Full_simp_tac 1);
qed "hypreal_lt_zero_iff";
-Goalw [hypreal_le_def] "((0::hypreal) <= x) = (-x <= x)";
-by (auto_tac (claset(), simpset() addsimps [hypreal_lt_zero_iff RS sym]));
-qed "hypreal_ge_zero_iff";
-
-Goalw [hypreal_le_def] "(x <= (0::hypreal)) = (x <= -x)";
-by (auto_tac (claset(), simpset() addsimps [hypreal_gt_zero_iff RS sym]));
-qed "hypreal_le_zero_iff";
-
-Goalw [hypreal_zero_def]
- "[| 0 < x; 0 < y |] ==> (0::hypreal) < x + y";
+Goalw [hypreal_zero_def] "[| 0 < x; 0 < y |] ==> (0::hypreal) < x + y";
by (res_inst_tac [("z","x")] eq_Abs_hypreal 1);
by (res_inst_tac [("z","y")] eq_Abs_hypreal 1);
by (auto_tac (claset(),
@@ -126,14 +117,12 @@
simpset()));
qed "hypreal_add_order_le";
-Goalw [hypreal_zero_def]
- "[| 0 < x; 0 < y |] ==> (0::hypreal) < x * y";
+Goalw [hypreal_zero_def] "[| 0 < x; 0 < y |] ==> (0::hypreal) < x * y";
by (res_inst_tac [("z","x")] eq_Abs_hypreal 1);
by (res_inst_tac [("z","y")] eq_Abs_hypreal 1);
by (auto_tac (claset() addSIs [exI],
simpset() addsimps [hypreal_less_def,hypreal_mult]));
-by (ultra_tac (claset() addIs [rename_numerals real_mult_order],
- simpset()) 1);
+by (ultra_tac (claset() addIs [real_mult_order], simpset()) 1);
qed "hypreal_mult_order";
Goal "[| x < 0; y < 0 |] ==> (0::hypreal) < x * y";
@@ -150,8 +139,8 @@
qed "hypreal_mult_less_zero";
Goalw [hypreal_one_def,hypreal_zero_def,hypreal_less_def] "0 < (1::hypreal)";
-by (res_inst_tac [("x","%n. Numeral0")] exI 1);
-by (res_inst_tac [("x","%n. Numeral1")] exI 1);
+by (res_inst_tac [("x","%n. 0")] exI 1);
+by (res_inst_tac [("x","%n. 1")] exI 1);
by (auto_tac (claset(),
simpset() addsimps [real_zero_less_one, FreeUltrafilterNat_Nat_set]));
qed "hypreal_zero_less_one";
@@ -332,29 +321,15 @@
by (dtac hypreal_mult_less_zero1 1 THEN assume_tac 1);
by (auto_tac (claset() addIs [hypreal_zero_less_one RS hypreal_less_asym],
simpset() addsimps [hypreal_minus_zero_less_iff]));
-qed "hypreal_inverse_gt_zero";
+qed "hypreal_inverse_gt_0";
Goal "x < 0 ==> inverse (x::hypreal) < 0";
by (ftac hypreal_not_refl2 1);
by (dtac (hypreal_minus_zero_less_iff RS iffD2) 1);
by (rtac (hypreal_minus_zero_less_iff RS iffD1) 1);
by (stac (hypreal_minus_inverse RS sym) 1);
-by (auto_tac (claset() addIs [hypreal_inverse_gt_zero], simpset()));
-qed "hypreal_inverse_less_zero";
-
-Goal "[| 0 <= x; 0 <= y |] ==> (x+y = 0) = (x = 0 & y = (0::hypreal))";
-by (auto_tac (claset() addIs [order_antisym], simpset()));
-qed "hypreal_add_nonneg_eq_0_iff";
-
-Goal "(x*y = 0) = (x = 0 | y = (0::hypreal))";
-by Auto_tac;
-by (blast_tac (claset() addDs [hypreal_mult_zero_disj]) 1);
-qed "hypreal_mult_is_0";
-
-Goal "(x*x + y*y + z*z = 0) = (x = 0 & y = 0 & z = (0::hypreal))";
-by (simp_tac (HOL_ss addsimps [hypreal_le_square, hypreal_le_add_order,
- hypreal_add_nonneg_eq_0_iff, hypreal_mult_is_0]) 1);
-qed "hypreal_three_squares_add_zero_iff";
+by (auto_tac (claset() addIs [hypreal_inverse_gt_0], simpset()));
+qed "hypreal_inverse_less_0";
Goal "(x::hypreal)*x <= x*x + y*y";
by (res_inst_tac [("z","x")] eq_Abs_hypreal 1);
@@ -471,17 +446,17 @@
by (res_inst_tac [("t","r")] (hypreal_inverse_inverse RS subst) 1);
by (res_inst_tac [("t","x")] (hypreal_inverse_inverse RS subst) 1);
by (rtac hypreal_inverse_less_swap 1);
-by (auto_tac (claset(), simpset() addsimps [hypreal_inverse_gt_zero]));
+by (auto_tac (claset(), simpset() addsimps [hypreal_inverse_gt_0]));
qed "hypreal_inverse_less_iff";
Goal "[| 0 < z; x < y |] ==> x * inverse z < y * inverse (z::hypreal)";
by (blast_tac (claset() addSIs [hypreal_mult_less_mono1,
- hypreal_inverse_gt_zero]) 1);
+ hypreal_inverse_gt_0]) 1);
qed "hypreal_mult_inverse_less_mono1";
Goal "[| 0 < z; x < y |] ==> inverse z * x < inverse (z::hypreal) * y";
by (blast_tac (claset() addSIs [hypreal_mult_less_mono2,
- hypreal_inverse_gt_zero]) 1);
+ hypreal_inverse_gt_0]) 1);
qed "hypreal_mult_inverse_less_mono2";
Goal "[| (0::hypreal) < z; x*z < y*z |] ==> x < y";
@@ -526,25 +501,25 @@
by (ALLGOALS (dtac hypreal_mult_less_zero THEN' assume_tac));
by (auto_tac (claset() addDs [order_less_not_sym],
simpset() addsimps [hypreal_mult_commute]));
-qed "hypreal_zero_less_mult_iff";
+qed "hypreal_0_less_mult_iff";
Goal "((0::hypreal) <= x*y) = (0 <= x & 0 <= y | x <= 0 & y <= 0)";
by (auto_tac (claset() addDs [hypreal_mult_zero_disj],
simpset() addsimps [order_le_less, linorder_not_less,
- hypreal_zero_less_mult_iff]));
-qed "hypreal_zero_le_mult_iff";
+ hypreal_0_less_mult_iff]));
+qed "hypreal_0_le_mult_iff";
Goal "(x*y < (0::hypreal)) = (0 < x & y < 0 | x < 0 & 0 < y)";
by (auto_tac (claset(),
- simpset() addsimps [hypreal_zero_le_mult_iff,
+ simpset() addsimps [hypreal_0_le_mult_iff,
linorder_not_le RS sym]));
by (auto_tac (claset() addDs [order_less_not_sym],
simpset() addsimps [linorder_not_le]));
-qed "hypreal_mult_less_zero_iff";
+qed "hypreal_mult_less_0_iff";
Goal "(x*y <= (0::hypreal)) = (0 <= x & y <= 0 | x <= 0 & 0 <= y)";
by (auto_tac (claset() addDs [order_less_not_sym],
- simpset() addsimps [hypreal_zero_less_mult_iff,
+ simpset() addsimps [hypreal_0_less_mult_iff,
linorder_not_less RS sym]));
-qed "hypreal_mult_le_zero_iff";
+qed "hypreal_mult_le_0_iff";
--- a/src/HOL/Hyperreal/HyperPow.ML Thu Nov 01 21:12:13 2001 +0100
+++ b/src/HOL/Hyperreal/HyperPow.ML Fri Nov 02 17:55:24 2001 +0100
@@ -6,19 +6,19 @@
Exponentials on the hyperreals
*)
-Goal "(Numeral0::hypreal) ^ (Suc n) = 0";
-by (Auto_tac);
+Goal "(0::hypreal) ^ (Suc n) = 0";
+by Auto_tac;
qed "hrealpow_zero";
Addsimps [hrealpow_zero];
-Goal "r ~= (Numeral0::hypreal) --> r ^ n ~= 0";
+Goal "r ~= (0::hypreal) --> r ^ n ~= 0";
by (induct_tac "n" 1);
by Auto_tac;
qed_spec_mp "hrealpow_not_zero";
-Goal "r ~= (Numeral0::hypreal) --> inverse(r ^ n) = (inverse r) ^ n";
+Goal "r ~= (0::hypreal) --> inverse(r ^ n) = (inverse r) ^ n";
by (induct_tac "n" 1);
-by (Auto_tac);
+by Auto_tac;
by (forw_inst_tac [("n","n")] hrealpow_not_zero 1);
by (auto_tac (claset(), simpset() addsimps [hypreal_inverse_distrib]));
qed_spec_mp "hrealpow_inverse";
@@ -42,40 +42,40 @@
by (Simp_tac 1);
qed "hrealpow_two";
-Goal "(Numeral0::hypreal) <= r --> Numeral0 <= r ^ n";
+Goal "(0::hypreal) <= r --> 0 <= r ^ n";
by (induct_tac "n" 1);
by (auto_tac (claset(), simpset() addsimps [hypreal_0_le_mult_iff]));
qed_spec_mp "hrealpow_ge_zero";
-Goal "(Numeral0::hypreal) < r --> Numeral0 < r ^ n";
+Goal "(0::hypreal) < r --> 0 < r ^ n";
by (induct_tac "n" 1);
by (auto_tac (claset(), simpset() addsimps [hypreal_0_less_mult_iff]));
qed_spec_mp "hrealpow_gt_zero";
-Goal "x <= y & (Numeral0::hypreal) < x --> x ^ n <= y ^ n";
+Goal "x <= y & (0::hypreal) < x --> x ^ n <= y ^ n";
by (induct_tac "n" 1);
by (auto_tac (claset() addSIs [hypreal_mult_le_mono], simpset()));
by (asm_simp_tac (simpset() addsimps [hrealpow_ge_zero]) 1);
qed_spec_mp "hrealpow_le";
-Goal "x < y & (Numeral0::hypreal) < x & 0 < n --> x ^ n < y ^ n";
+Goal "x < y & (0::hypreal) < x & 0 < n --> x ^ n < y ^ n";
by (induct_tac "n" 1);
by (auto_tac (claset() addIs [hypreal_mult_less_mono,gr0I],
simpset() addsimps [hrealpow_gt_zero]));
qed "hrealpow_less";
-Goal "Numeral1 ^ n = (Numeral1::hypreal)";
+Goal "1 ^ n = (1::hypreal)";
by (induct_tac "n" 1);
-by (Auto_tac);
+by Auto_tac;
qed "hrealpow_eq_one";
Addsimps [hrealpow_eq_one];
-Goal "abs(-(Numeral1 ^ n)) = (Numeral1::hypreal)";
+Goal "abs(-(1 ^ n)) = (1::hypreal)";
by Auto_tac;
qed "hrabs_minus_hrealpow_one";
Addsimps [hrabs_minus_hrealpow_one];
-Goal "abs(-1 ^ n) = (Numeral1::hypreal)";
+Goal "abs(-1 ^ n) = (1::hypreal)";
by (induct_tac "n" 1);
by Auto_tac;
qed "hrabs_hrealpow_minus_one";
@@ -86,32 +86,43 @@
by (auto_tac (claset(), simpset() addsimps hypreal_mult_ac));
qed "hrealpow_mult";
-Goal "(Numeral0::hypreal) <= r ^Suc (Suc 0)";
+Goal "(0::hypreal) <= r ^ Suc (Suc 0)";
by (auto_tac (claset(), simpset() addsimps [hypreal_0_le_mult_iff]));
qed "hrealpow_two_le";
Addsimps [hrealpow_two_le];
-Goal "(Numeral0::hypreal) <= u ^ Suc (Suc 0) + v ^ Suc (Suc 0)";
-by (simp_tac (HOL_ss addsimps [hrealpow_two_le,
- rename_numerals hypreal_le_add_order]) 1);
+Goal "(0::hypreal) <= u ^ Suc (Suc 0) + v ^ Suc (Suc 0)";
+by (simp_tac (HOL_ss addsimps [hrealpow_two_le, hypreal_le_add_order]) 1);
qed "hrealpow_two_le_add_order";
Addsimps [hrealpow_two_le_add_order];
-Goal "(Numeral0::hypreal) <= u ^ Suc (Suc 0) + v ^ Suc (Suc 0) + w ^ Suc (Suc 0)";
-by (simp_tac (HOL_ss addsimps [hrealpow_two_le,
- rename_numerals hypreal_le_add_order]) 1);
+Goal "(0::hypreal) <= u ^ Suc (Suc 0) + v ^ Suc (Suc 0) + w ^ Suc (Suc 0)";
+by (simp_tac (HOL_ss addsimps [hrealpow_two_le, hypreal_le_add_order]) 1);
qed "hrealpow_two_le_add_order2";
Addsimps [hrealpow_two_le_add_order2];
-Goal "(x ^ Suc (Suc 0) + y ^ Suc (Suc 0) + z ^ Suc (Suc 0) = (Numeral0::hypreal)) = (x = Numeral0 & y = Numeral0 & z = Numeral0)";
+Goal "[| 0 <= x; 0 <= y |] ==> (x+y = 0) = (x = 0 & y = (0::hypreal))";
+by (auto_tac (claset() addIs [order_antisym], simpset()));
+qed "hypreal_add_nonneg_eq_0_iff";
+
+Goal "(x*y = 0) = (x = 0 | y = (0::hypreal))";
+by Auto_tac;
+qed "hypreal_mult_is_0";
+
+Goal "(x*x + y*y + z*z = 0) = (x = 0 & y = 0 & z = (0::hypreal))";
+by (simp_tac (HOL_ss addsimps [hypreal_le_square, hypreal_le_add_order,
+ hypreal_add_nonneg_eq_0_iff, hypreal_mult_is_0]) 1);
+qed "hypreal_three_squares_add_zero_iff";
+
+Goal "(x ^ Suc (Suc 0) + y ^ Suc (Suc 0) + z ^ Suc (Suc 0) = (0::hypreal)) = (x = 0 & y = 0 & z = 0)";
by (simp_tac (HOL_ss addsimps
- [rename_numerals hypreal_three_squares_add_zero_iff, hrealpow_two]) 1);
+ [hypreal_three_squares_add_zero_iff, hrealpow_two]) 1);
qed "hrealpow_three_squares_add_zero_iff";
Addsimps [hrealpow_three_squares_add_zero_iff];
Goal "abs(x ^ Suc (Suc 0)) = (x::hypreal) ^ Suc (Suc 0)";
by (auto_tac (claset(),
- simpset() addsimps [hrabs_def, hypreal_0_le_mult_iff]));
+ simpset() addsimps [hrabs_def, hypreal_0_le_mult_iff]));
qed "hrabs_hrealpow_two";
Addsimps [hrabs_hrealpow_two];
@@ -121,21 +132,21 @@
qed "hrealpow_two_hrabs";
Addsimps [hrealpow_two_hrabs];
-Goal "(Numeral1::hypreal) < r ==> Numeral1 < r ^ Suc (Suc 0)";
+Goal "(1::hypreal) < r ==> 1 < r ^ Suc (Suc 0)";
by (auto_tac (claset(), simpset() addsimps [hrealpow_two]));
-by (res_inst_tac [("y","Numeral1*Numeral1")] order_le_less_trans 1);
+by (res_inst_tac [("y","1*1")] order_le_less_trans 1);
by (rtac hypreal_mult_less_mono 2);
by Auto_tac;
qed "hrealpow_two_gt_one";
-Goal "(Numeral1::hypreal) <= r ==> Numeral1 <= r ^ Suc (Suc 0)";
+Goal "(1::hypreal) <= r ==> 1 <= r ^ Suc (Suc 0)";
by (etac (order_le_imp_less_or_eq RS disjE) 1);
by (etac (hrealpow_two_gt_one RS order_less_imp_le) 1);
by Auto_tac;
qed "hrealpow_two_ge_one";
-Goal "(Numeral1::hypreal) <= 2 ^ n";
-by (res_inst_tac [("y","Numeral1 ^ n")] order_trans 1);
+Goal "(1::hypreal) <= 2 ^ n";
+by (res_inst_tac [("y","1 ^ n")] order_trans 1);
by (rtac hrealpow_le 2);
by Auto_tac;
qed "two_hrealpow_ge_one";
@@ -149,9 +160,9 @@
qed "two_hrealpow_gt";
Addsimps [two_hrealpow_gt,two_hrealpow_ge_one];
-Goal "-1 ^ (2*n) = (Numeral1::hypreal)";
+Goal "-1 ^ (2*n) = (1::hypreal)";
by (induct_tac "n" 1);
-by (Auto_tac);
+by Auto_tac;
qed "hrealpow_minus_one";
Goal "n+n = (2*n::nat)";
@@ -159,24 +170,24 @@
qed "double_lemma";
(*ugh: need to get rid fo the n+n*)
-Goal "-1 ^ (n + n) = (Numeral1::hypreal)";
+Goal "-1 ^ (n + n) = (1::hypreal)";
by (auto_tac (claset(),
simpset() addsimps [double_lemma, hrealpow_minus_one]));
qed "hrealpow_minus_one2";
Addsimps [hrealpow_minus_one2];
Goal "(-(x::hypreal)) ^ Suc (Suc 0) = x ^ Suc (Suc 0)";
-by (Auto_tac);
+by Auto_tac;
qed "hrealpow_minus_two";
Addsimps [hrealpow_minus_two];
-Goal "(Numeral0::hypreal) < r & r < Numeral1 --> r ^ Suc n < r ^ n";
+Goal "(0::hypreal) < r & r < 1 --> r ^ Suc n < r ^ n";
by (induct_tac "n" 1);
by (auto_tac (claset(),
simpset() addsimps [hypreal_mult_less_mono2]));
qed_spec_mp "hrealpow_Suc_less";
-Goal "(Numeral0::hypreal) <= r & r < Numeral1 --> r ^ Suc n <= r ^ n";
+Goal "(0::hypreal) <= r & r < 1 --> r ^ Suc n <= r ^ n";
by (induct_tac "n" 1);
by (auto_tac (claset() addIs [order_less_imp_le]
addSDs [order_le_imp_less_or_eq,hrealpow_Suc_less],
@@ -186,9 +197,7 @@
Goal "Abs_hypreal(hyprel``{%n. X n}) ^ m = Abs_hypreal(hyprel``{%n. (X n) ^ m})";
by (induct_tac "m" 1);
by (auto_tac (claset(),
- simpset() delsimps [one_eq_numeral_1]
- addsimps [hypreal_one_def, hypreal_mult,
- one_eq_numeral_1 RS sym]));
+ simpset() addsimps [hypreal_one_def, hypreal_mult]));
qed "hrealpow";
Goal "(x + (y::hypreal)) ^ Suc (Suc 0) = \
@@ -204,8 +213,8 @@
property for the real rather than prove it directly
using induction: proof is much simpler this way!
---------------------------------------------------------------*)
-Goal "[|(Numeral0::hypreal) <= x; Numeral0 <= y;x ^ Suc n <= y ^ Suc n |] ==> x <= y";
-by (full_simp_tac (simpset() addsimps [rename_numerals hypreal_zero_def]) 1);
+Goal "[|(0::hypreal) <= x; 0 <= y;x ^ Suc n <= y ^ Suc n |] ==> x <= y";
+by (full_simp_tac (simpset() addsimps [hypreal_zero_def]) 1);
by (res_inst_tac [("z","x")] eq_Abs_hypreal 1);
by (res_inst_tac [("z","y")] eq_Abs_hypreal 1);
by (auto_tac (claset(),
@@ -241,15 +250,15 @@
by (Fuf_tac 1);
qed "hyperpow";
-Goalw [hypnat_one_def] "(Numeral0::hypreal) pow (n + (1::hypnat)) = Numeral0";
-by (simp_tac (simpset() addsimps [rename_numerals hypreal_zero_def]) 1);
+Goalw [hypnat_one_def] "(0::hypreal) pow (n + (1::hypnat)) = 0";
+by (simp_tac (simpset() addsimps [hypreal_zero_def]) 1);
by (res_inst_tac [("z","n")] eq_Abs_hypnat 1);
by (auto_tac (claset(), simpset() addsimps [hyperpow,hypnat_add]));
qed "hyperpow_zero";
Addsimps [hyperpow_zero];
-Goal "r ~= (Numeral0::hypreal) --> r pow n ~= Numeral0";
-by (simp_tac (simpset() addsimps [rename_numerals hypreal_zero_def]) 1);
+Goal "r ~= (0::hypreal) --> r pow n ~= 0";
+by (simp_tac (simpset() addsimps [hypreal_zero_def]) 1);
by (res_inst_tac [("z","n")] eq_Abs_hypnat 1);
by (res_inst_tac [("z","r")] eq_Abs_hypreal 1);
by (auto_tac (claset(), simpset() addsimps [hyperpow]));
@@ -258,8 +267,8 @@
simpset()) 1);
qed_spec_mp "hyperpow_not_zero";
-Goal "r ~= (Numeral0::hypreal) --> inverse(r pow n) = (inverse r) pow n";
-by (simp_tac (simpset() addsimps [rename_numerals hypreal_zero_def]) 1);
+Goal "r ~= (0::hypreal) --> inverse(r pow n) = (inverse r) pow n";
+by (simp_tac (simpset() addsimps [hypreal_zero_def]) 1);
by (res_inst_tac [("z","n")] eq_Abs_hypnat 1);
by (res_inst_tac [("z","r")] eq_Abs_hypreal 1);
by (auto_tac (claset() addSDs [FreeUltrafilterNat_Compl_mem],
@@ -298,24 +307,24 @@
simpset() addsimps [hyperpow,hypnat_add, hypreal_mult]));
qed "hyperpow_two";
-Goal "(Numeral0::hypreal) < r --> Numeral0 < r pow n";
-by (simp_tac (simpset() addsimps [rename_numerals hypreal_zero_def]) 1);
+Goal "(0::hypreal) < r --> 0 < r pow n";
+by (simp_tac (simpset() addsimps [hypreal_zero_def]) 1);
by (res_inst_tac [("z","n")] eq_Abs_hypnat 1);
by (res_inst_tac [("z","r")] eq_Abs_hypreal 1);
by (auto_tac (claset() addSEs [FreeUltrafilterNat_subset, realpow_gt_zero],
simpset() addsimps [hyperpow,hypreal_less, hypreal_le]));
qed_spec_mp "hyperpow_gt_zero";
-Goal "(Numeral0::hypreal) <= r --> Numeral0 <= r pow n";
-by (simp_tac (simpset() addsimps [rename_numerals hypreal_zero_def]) 1);
+Goal "(0::hypreal) <= r --> 0 <= r pow n";
+by (simp_tac (simpset() addsimps [hypreal_zero_def]) 1);
by (res_inst_tac [("z","n")] eq_Abs_hypnat 1);
by (res_inst_tac [("z","r")] eq_Abs_hypreal 1);
by (auto_tac (claset() addSEs [FreeUltrafilterNat_subset, realpow_ge_zero],
simpset() addsimps [hyperpow,hypreal_le]));
qed "hyperpow_ge_zero";
-Goal "(Numeral0::hypreal) < x & x <= y --> x pow n <= y pow n";
-by (full_simp_tac (simpset() addsimps [rename_numerals hypreal_zero_def]) 1);
+Goal "(0::hypreal) < x & x <= y --> x pow n <= y pow n";
+by (full_simp_tac (simpset() addsimps [hypreal_zero_def]) 1);
by (res_inst_tac [("z","n")] eq_Abs_hypnat 1);
by (res_inst_tac [("z","x")] eq_Abs_hypreal 1);
by (res_inst_tac [("z","y")] eq_Abs_hypreal 1);
@@ -326,27 +335,26 @@
by (auto_tac (claset() addIs [realpow_le], simpset()));
qed_spec_mp "hyperpow_le";
-Goal "Numeral1 pow n = (Numeral1::hypreal)";
-by (simp_tac (HOL_ss addsimps [one_eq_numeral_1 RS sym, hypreal_one_def]) 1);
+Goal "1 pow n = (1::hypreal)";
by (res_inst_tac [("z","n")] eq_Abs_hypnat 1);
-by (auto_tac (claset(), simpset() addsimps [hyperpow]));
+by (auto_tac (claset(), simpset() addsimps [hypreal_one_def, hyperpow]));
qed "hyperpow_eq_one";
Addsimps [hyperpow_eq_one];
-Goal "abs(-(Numeral1 pow n)) = (Numeral1::hypreal)";
-by (simp_tac (HOL_ss addsimps [one_eq_numeral_1 RS sym, hypreal_one_def]) 1);
+Goal "abs(-(1 pow n)) = (1::hypreal)";
by (res_inst_tac [("z","n")] eq_Abs_hypnat 1);
-by (auto_tac (claset(), simpset() addsimps [hyperpow,hypreal_hrabs]));
+by (auto_tac (claset(),
+ simpset() addsimps [hyperpow, hypreal_hrabs, hypreal_one_def]));
qed "hrabs_minus_hyperpow_one";
Addsimps [hrabs_minus_hyperpow_one];
-Goal "abs(-1 pow n) = (Numeral1::hypreal)";
+Goal "abs(-1 pow n) = (1::hypreal)";
by (subgoal_tac "abs((- (1::hypreal)) pow n) = (1::hypreal)" 1);
by (Asm_full_simp_tac 1);
-by (simp_tac (HOL_ss addsimps [hypreal_one_def]) 1);
by (res_inst_tac [("z","n")] eq_Abs_hypnat 1);
by (auto_tac (claset(),
- simpset() addsimps [hyperpow,hypreal_minus,hypreal_hrabs]));
+ simpset() addsimps [hypreal_one_def, hyperpow,hypreal_minus,
+ hypreal_hrabs]));
qed "hrabs_hyperpow_minus_one";
Addsimps [hrabs_hyperpow_minus_one];
@@ -358,7 +366,7 @@
simpset() addsimps [hyperpow, hypreal_mult,realpow_mult]));
qed "hyperpow_mult";
-Goal "(Numeral0::hypreal) <= r pow ((1::hypnat) + (1::hypnat))";
+Goal "(0::hypreal) <= r pow ((1::hypnat) + (1::hypnat))";
by (auto_tac (claset(),
simpset() addsimps [hyperpow_two, hypreal_0_le_mult_iff]));
qed "hyperpow_two_le";
@@ -375,21 +383,21 @@
Addsimps [hyperpow_two_hrabs];
(*? very similar to hrealpow_two_gt_one *)
-Goal "(Numeral1::hypreal) < r ==> Numeral1 < r pow ((1::hypnat) + (1::hypnat))";
+Goal "(1::hypreal) < r ==> 1 < r pow ((1::hypnat) + (1::hypnat))";
by (auto_tac (claset(), simpset() addsimps [hyperpow_two]));
-by (res_inst_tac [("y","Numeral1*Numeral1")] order_le_less_trans 1);
+by (res_inst_tac [("y","1*1")] order_le_less_trans 1);
by (rtac hypreal_mult_less_mono 2);
by Auto_tac;
qed "hyperpow_two_gt_one";
-Goal "(Numeral1::hypreal) <= r ==> Numeral1 <= r pow ((1::hypnat) + (1::hypnat))";
+Goal "(1::hypreal) <= r ==> 1 <= r pow ((1::hypnat) + (1::hypnat))";
by (auto_tac (claset() addSDs [order_le_imp_less_or_eq]
addIs [hyperpow_two_gt_one,order_less_imp_le],
simpset()));
qed "hyperpow_two_ge_one";
-Goal "(Numeral1::hypreal) <= 2 pow n";
-by (res_inst_tac [("y","Numeral1 pow n")] order_trans 1);
+Goal "(1::hypreal) <= 2 pow n";
+by (res_inst_tac [("y","1 pow n")] order_trans 1);
by (rtac hyperpow_le 2);
by Auto_tac;
qed "two_hyperpow_ge_one";
@@ -397,7 +405,7 @@
Addsimps [simplify (simpset()) realpow_minus_one];
-Goal "-1 pow (((1::hypnat) + (1::hypnat))*n) = (Numeral1::hypreal)";
+Goal "-1 pow (((1::hypnat) + (1::hypnat))*n) = (1::hypreal)";
by (subgoal_tac "(-((1::hypreal))) pow (((1::hypnat) + (1::hypnat))*n) = (1::hypreal)" 1);
by (Asm_full_simp_tac 1);
by (simp_tac (HOL_ss addsimps [hypreal_one_def]) 1);
@@ -409,52 +417,44 @@
Addsimps [hyperpow_minus_one2];
Goalw [hypnat_one_def]
- "(Numeral0::hypreal) < r & r < Numeral1 --> r pow (n + (1::hypnat)) < r pow n";
-by (full_simp_tac
- (HOL_ss addsimps [zero_eq_numeral_0 RS sym, hypreal_zero_def,
- one_eq_numeral_1 RS sym, hypreal_one_def]) 1);
+ "(0::hypreal) < r & r < 1 --> r pow (n + (1::hypnat)) < r pow n";
by (res_inst_tac [("z","n")] eq_Abs_hypnat 1);
by (res_inst_tac [("z","r")] eq_Abs_hypreal 1);
by (auto_tac (claset() addSDs [conjI RS realpow_Suc_less]
addEs [FreeUltrafilterNat_Int RS FreeUltrafilterNat_subset],
- simpset() addsimps [hyperpow,hypreal_less,hypnat_add]));
+ simpset() addsimps [hypreal_zero_def, hypreal_one_def,
+ hyperpow, hypreal_less, hypnat_add]));
qed_spec_mp "hyperpow_Suc_less";
Goalw [hypnat_one_def]
- "Numeral0 <= r & r < (Numeral1::hypreal) --> r pow (n + (1::hypnat)) <= r pow n";
-by (full_simp_tac
- (HOL_ss addsimps [zero_eq_numeral_0 RS sym, hypreal_zero_def,
- one_eq_numeral_1 RS sym, hypreal_one_def]) 1);
+ "0 <= r & r < (1::hypreal) --> r pow (n + (1::hypnat)) <= r pow n";
by (res_inst_tac [("z","n")] eq_Abs_hypnat 1);
by (res_inst_tac [("z","r")] eq_Abs_hypreal 1);
-by (auto_tac (claset() addSDs [conjI RS realpow_Suc_le] addEs
- [FreeUltrafilterNat_Int RS FreeUltrafilterNat_subset ],
- simpset() addsimps [hyperpow,hypreal_le,hypnat_add,
- hypreal_less]));
+by (auto_tac (claset() addSDs [conjI RS realpow_Suc_le]
+ addEs [FreeUltrafilterNat_Int RS FreeUltrafilterNat_subset],
+ simpset() addsimps [hypreal_zero_def, hypreal_one_def, hyperpow,
+ hypreal_le,hypnat_add, hypreal_less]));
qed_spec_mp "hyperpow_Suc_le";
Goalw [hypnat_one_def]
- "(Numeral0::hypreal) <= r & r < Numeral1 & n < N --> r pow N <= r pow n";
-by (full_simp_tac
- (HOL_ss addsimps [zero_eq_numeral_0 RS sym, hypreal_zero_def,
- one_eq_numeral_1 RS sym, hypreal_one_def]) 1);
+ "(0::hypreal) <= r & r < 1 & n < N --> r pow N <= r pow n";
by (res_inst_tac [("z","n")] eq_Abs_hypnat 1);
by (res_inst_tac [("z","N")] eq_Abs_hypnat 1);
by (res_inst_tac [("z","r")] eq_Abs_hypreal 1);
by (auto_tac (claset(),
- simpset() addsimps [hyperpow, hypreal_le,hypreal_less,hypnat_less]));
+ simpset() addsimps [hyperpow, hypreal_le,hypreal_less,
+ hypnat_less, hypreal_zero_def, hypreal_one_def]));
by (etac (FreeUltrafilterNat_Int RS FreeUltrafilterNat_subset) 1);
by (etac FreeUltrafilterNat_Int 1);
-by (auto_tac (claset() addSDs [conjI RS realpow_less_le],
- simpset()));
+by (auto_tac (claset() addSDs [conjI RS realpow_less_le], simpset()));
qed_spec_mp "hyperpow_less_le";
-Goal "[| (Numeral0::hypreal) <= r; r < Numeral1 |] \
+Goal "[| (0::hypreal) <= r; r < 1 |] \
\ ==> ALL N n. n < N --> r pow N <= r pow n";
by (blast_tac (claset() addSIs [hyperpow_less_le]) 1);
qed "hyperpow_less_le2";
-Goal "[| Numeral0 <= r; r < (Numeral1::hypreal); N : HNatInfinite |] \
+Goal "[| 0 <= r; r < (1::hypreal); N : HNatInfinite |] \
\ ==> ALL n: Nats. r pow N <= r pow n";
by (auto_tac (claset() addSIs [hyperpow_less_le],
simpset() addsimps [HNatInfinite_iff]));
@@ -471,25 +471,25 @@
qed "hyperpow_SReal";
Addsimps [hyperpow_SReal];
-Goal "N : HNatInfinite ==> (Numeral0::hypreal) pow N = 0";
+Goal "N : HNatInfinite ==> (0::hypreal) pow N = 0";
by (dtac HNatInfinite_is_Suc 1);
-by (Auto_tac);
+by Auto_tac;
qed "hyperpow_zero_HNatInfinite";
Addsimps [hyperpow_zero_HNatInfinite];
-Goal "[| (Numeral0::hypreal) <= r; r < Numeral1; n <= N |] ==> r pow N <= r pow n";
+Goal "[| (0::hypreal) <= r; r < 1; n <= N |] ==> r pow N <= r pow n";
by (dres_inst_tac [("y","N")] hypnat_le_imp_less_or_eq 1);
by (auto_tac (claset() addIs [hyperpow_less_le], simpset()));
qed "hyperpow_le_le";
-Goal "[| (Numeral0::hypreal) < r; r < Numeral1 |] ==> r pow (n + (1::hypnat)) <= r";
+Goal "[| (0::hypreal) < r; r < 1 |] ==> r pow (n + (1::hypnat)) <= r";
by (dres_inst_tac [("n","(1::hypnat)")] (order_less_imp_le RS hyperpow_le_le) 1);
-by (Auto_tac);
+by Auto_tac;
qed "hyperpow_Suc_le_self";
-Goal "[| (Numeral0::hypreal) <= r; r < Numeral1 |] ==> r pow (n + (1::hypnat)) <= r";
+Goal "[| (0::hypreal) <= r; r < 1 |] ==> r pow (n + (1::hypnat)) <= r";
by (dres_inst_tac [("n","(1::hypnat)")] hyperpow_le_le 1);
-by (Auto_tac);
+by Auto_tac;
qed "hyperpow_Suc_le_self2";
Goalw [Infinitesimal_def]
--- a/src/HOL/Hyperreal/Lim.ML Thu Nov 01 21:12:13 2001 +0100
+++ b/src/HOL/Hyperreal/Lim.ML Fri Nov 02 17:55:24 2001 +0100
@@ -65,16 +65,16 @@
(*----------------------------------------------
LIM_zero
----------------------------------------------*)
-Goal "f -- a --> l ==> (%x. f(x) + -l) -- a --> Numeral0";
-by (res_inst_tac [("z1","l")] (rename_numerals (real_add_minus RS subst)) 1);
+Goal "f -- a --> l ==> (%x. f(x) + -l) -- a --> 0";
+by (res_inst_tac [("z1","l")] ((real_add_minus RS subst)) 1);
by (rtac LIM_add_minus 1 THEN Auto_tac);
qed "LIM_zero";
(*--------------------------
Limit not zero
--------------------------*)
-Goalw [LIM_def] "k \\<noteq> Numeral0 ==> ~ ((%x. k) -- x --> Numeral0)";
-by (res_inst_tac [("R1.0","k"),("R2.0","Numeral0")] real_linear_less2 1);
+Goalw [LIM_def] "k \\<noteq> 0 ==> ~ ((%x. k) -- x --> 0)";
+by (res_inst_tac [("R1.0","k"),("R2.0","0")] real_linear_less2 1);
by (auto_tac (claset(), simpset() addsimps [real_abs_def]));
by (res_inst_tac [("x","-k")] exI 1);
by (res_inst_tac [("x","k")] exI 2);
@@ -85,7 +85,7 @@
by Auto_tac;
qed "LIM_not_zero";
-(* [| k \\<noteq> Numeral0; (%x. k) -- x --> Numeral0 |] ==> R *)
+(* [| k \\<noteq> 0; (%x. k) -- x --> 0 |] ==> R *)
bind_thm("LIM_not_zeroE", LIM_not_zero RS notE);
Goal "(%x. k) -- x --> L ==> k = L";
@@ -108,9 +108,9 @@
LIM_mult_zero
-------------*)
Goalw [LIM_def]
- "[| f -- x --> Numeral0; g -- x --> Numeral0 |] ==> (%x. f(x)*g(x)) -- x --> Numeral0";
+ "[| f -- x --> 0; g -- x --> 0 |] ==> (%x. f(x)*g(x)) -- x --> 0";
by Safe_tac;
-by (dres_inst_tac [("x","Numeral1")] spec 1);
+by (dres_inst_tac [("x","1")] spec 1);
by (dres_inst_tac [("x","r")] spec 1);
by (cut_facts_tac [real_zero_less_one] 1);
by (asm_full_simp_tac (simpset() addsimps
@@ -146,7 +146,7 @@
by (auto_tac (claset(), simpset() addsimps [real_add_minus_iff]));
qed "LIM_equal";
-Goal "[| (%x. f(x) + -g(x)) -- a --> Numeral0; g -- a --> l |] \
+Goal "[| (%x. f(x) + -g(x)) -- a --> 0; g -- a --> l |] \
\ ==> f -- a --> l";
by (dtac LIM_add 1 THEN assume_tac 1);
by (auto_tac (claset(), simpset() addsimps [real_add_assoc]));
@@ -181,17 +181,17 @@
Limit: NS definition ==> standard definition
---------------------------------------------------------------------*)
-Goal "\\<forall>s. Numeral0 < s --> (\\<exists>xa. xa \\<noteq> x & \
+Goal "\\<forall>s. 0 < s --> (\\<exists>xa. xa \\<noteq> x & \
\ abs (xa + - x) < s & r \\<le> abs (f xa + -L)) \
\ ==> \\<forall>n::nat. \\<exists>xa. xa \\<noteq> x & \
\ abs(xa + -x) < inverse(real(Suc n)) & r \\<le> abs(f xa + -L)";
by (Clarify_tac 1);
by (cut_inst_tac [("n1","n")]
- (real_of_nat_Suc_gt_zero RS rename_numerals real_inverse_gt_zero) 1);
+ (real_of_nat_Suc_gt_zero RS real_inverse_gt_0) 1);
by Auto_tac;
val lemma_LIM = result();
-Goal "\\<forall>s. Numeral0 < s --> (\\<exists>xa. xa \\<noteq> x & \
+Goal "\\<forall>s. 0 < s --> (\\<exists>xa. xa \\<noteq> x & \
\ abs (xa + - x) < s & r \\<le> abs (f xa + -L)) \
\ ==> \\<exists>X. \\<forall>n::nat. X n \\<noteq> x & \
\ abs(X n + -x) < inverse(real(Suc n)) & r \\<le> abs(f (X n) + -L)";
@@ -320,7 +320,7 @@
NSLIM_inverse
-----------------------------*)
Goalw [NSLIM_def]
- "[| f -- a --NS> L; L \\<noteq> Numeral0 |] \
+ "[| f -- a --NS> L; L \\<noteq> 0 |] \
\ ==> (%x. inverse(f(x))) -- a --NS> (inverse L)";
by (Clarify_tac 1);
by (dtac spec 1);
@@ -329,28 +329,28 @@
qed "NSLIM_inverse";
Goal "[| f -- a --> L; \
-\ L \\<noteq> Numeral0 |] ==> (%x. inverse(f(x))) -- a --> (inverse L)";
+\ L \\<noteq> 0 |] ==> (%x. inverse(f(x))) -- a --> (inverse L)";
by (asm_full_simp_tac (simpset() addsimps [LIM_NSLIM_iff, NSLIM_inverse]) 1);
qed "LIM_inverse";
(*------------------------------
NSLIM_zero
------------------------------*)
-Goal "f -- a --NS> l ==> (%x. f(x) + -l) -- a --NS> Numeral0";
-by (res_inst_tac [("z1","l")] (rename_numerals (real_add_minus RS subst)) 1);
+Goal "f -- a --NS> l ==> (%x. f(x) + -l) -- a --NS> 0";
+by (res_inst_tac [("z1","l")] ((real_add_minus RS subst)) 1);
by (rtac NSLIM_add_minus 1 THEN Auto_tac);
qed "NSLIM_zero";
-Goal "f -- a --> l ==> (%x. f(x) + -l) -- a --> Numeral0";
+Goal "f -- a --> l ==> (%x. f(x) + -l) -- a --> 0";
by (asm_full_simp_tac (simpset() addsimps [LIM_NSLIM_iff, NSLIM_zero]) 1);
qed "LIM_zero2";
-Goal "(%x. f(x) - l) -- x --NS> Numeral0 ==> f -- x --NS> l";
+Goal "(%x. f(x) - l) -- x --NS> 0 ==> f -- x --NS> l";
by (dres_inst_tac [("g","%x. l"),("m","l")] NSLIM_add 1);
by (auto_tac (claset(),simpset() addsimps [real_diff_def, real_add_assoc]));
qed "NSLIM_zero_cancel";
-Goal "(%x. f(x) - l) -- x --> Numeral0 ==> f -- x --> l";
+Goal "(%x. f(x) - l) -- x --> 0 ==> f -- x --> l";
by (dres_inst_tac [("g","%x. l"),("m","l")] LIM_add 1);
by (auto_tac (claset(),simpset() addsimps [real_diff_def, real_add_assoc]));
qed "LIM_zero_cancel";
@@ -359,17 +359,17 @@
(*--------------------------
NSLIM_not_zero
--------------------------*)
-Goalw [NSLIM_def] "k \\<noteq> Numeral0 ==> ~ ((%x. k) -- x --NS> Numeral0)";
+Goalw [NSLIM_def] "k \\<noteq> 0 ==> ~ ((%x. k) -- x --NS> 0)";
by Auto_tac;
by (res_inst_tac [("x","hypreal_of_real x + epsilon")] exI 1);
by (auto_tac (claset() addIs [Infinitesimal_add_approx_self RS approx_sym],
- simpset() addsimps [rename_numerals hypreal_epsilon_not_zero]));
+ simpset() addsimps [hypreal_epsilon_not_zero]));
qed "NSLIM_not_zero";
-(* [| k \\<noteq> Numeral0; (%x. k) -- x --NS> Numeral0 |] ==> R *)
+(* [| k \\<noteq> 0; (%x. k) -- x --NS> 0 |] ==> R *)
bind_thm("NSLIM_not_zeroE", NSLIM_not_zero RS notE);
-Goal "k \\<noteq> Numeral0 ==> ~ ((%x. k) -- x --> Numeral0)";
+Goal "k \\<noteq> 0 ==> ~ ((%x. k) -- x --> 0)";
by (asm_full_simp_tac (simpset() addsimps [LIM_NSLIM_iff, NSLIM_not_zero]) 1);
qed "LIM_not_zero2";
@@ -405,16 +405,16 @@
(*--------------------
NSLIM_mult_zero
--------------------*)
-Goal "[| f -- x --NS> Numeral0; g -- x --NS> Numeral0 |] \
-\ ==> (%x. f(x)*g(x)) -- x --NS> Numeral0";
+Goal "[| f -- x --NS> 0; g -- x --NS> 0 |] \
+\ ==> (%x. f(x)*g(x)) -- x --NS> 0";
by (dtac NSLIM_mult 1 THEN Auto_tac);
qed "NSLIM_mult_zero";
(* we can use the corresponding thm LIM_mult2 *)
(* for standard definition of limit *)
-Goal "[| f -- x --> Numeral0; g -- x --> Numeral0 |] \
-\ ==> (%x. f(x)*g(x)) -- x --> Numeral0";
+Goal "[| f -- x --> 0; g -- x --> 0 |] \
+\ ==> (%x. f(x)*g(x)) -- x --> 0";
by (dtac LIM_mult2 1 THEN Auto_tac);
qed "LIM_mult_zero2";
@@ -499,8 +499,8 @@
--------------------------------------------------------------------------*)
(* Prove equivalence between NS limits - *)
(* seems easier than using standard def *)
-Goalw [NSLIM_def] "(f -- a --NS> L) = ((%h. f(a + h)) -- Numeral0 --NS> L)";
-by (auto_tac (claset(),simpset() addsimps [hypreal_of_real_zero]));
+Goalw [NSLIM_def] "(f -- a --NS> L) = ((%h. f(a + h)) -- 0 --NS> L)";
+by Auto_tac;
by (dres_inst_tac [("x","hypreal_of_real a + x")] spec 1);
by (dres_inst_tac [("x","-hypreal_of_real a + x")] spec 2);
by Safe_tac;
@@ -516,15 +516,15 @@
hypreal_add, real_add_assoc, approx_refl, hypreal_zero_def]));
qed "NSLIM_h_iff";
-Goal "(f -- a --NS> f a) = ((%h. f(a + h)) -- Numeral0 --NS> f a)";
+Goal "(f -- a --NS> f a) = ((%h. f(a + h)) -- 0 --NS> f a)";
by (rtac NSLIM_h_iff 1);
qed "NSLIM_isCont_iff";
-Goal "(f -- a --> f a) = ((%h. f(a + h)) -- Numeral0 --> f(a))";
+Goal "(f -- a --> f a) = ((%h. f(a + h)) -- 0 --> f(a))";
by (simp_tac (simpset() addsimps [LIM_NSLIM_iff, NSLIM_isCont_iff]) 1);
qed "LIM_isCont_iff";
-Goalw [isCont_def] "(isCont f x) = ((%h. f(x + h)) -- Numeral0 --> f(x))";
+Goalw [isCont_def] "(isCont f x) = ((%h. f(x + h)) -- 0 --> f(x))";
by (simp_tac (simpset() addsimps [LIM_isCont_iff]) 1);
qed "isCont_iff";
@@ -574,11 +574,11 @@
qed "isCont_minus";
Goalw [isCont_def]
- "[| isCont f x; f x \\<noteq> Numeral0 |] ==> isCont (%x. inverse (f x)) x";
+ "[| isCont f x; f x \\<noteq> 0 |] ==> isCont (%x. inverse (f x)) x";
by (blast_tac (claset() addIs [LIM_inverse]) 1);
qed "isCont_inverse";
-Goal "[| isNSCont f x; f x \\<noteq> Numeral0 |] ==> isNSCont (%x. inverse (f x)) x";
+Goal "[| isNSCont f x; f x \\<noteq> 0 |] ==> isNSCont (%x. inverse (f x)) x";
by (auto_tac (claset() addIs [isCont_inverse],simpset() addsimps
[isNSCont_isCont_iff]));
qed "isNSCont_inverse";
@@ -667,8 +667,10 @@
qed "isNSUContD";
Goalw [isUCont_def,isCont_def,LIM_def]
- "isUCont f ==> \\<exists>x. isCont f x";
-by (Force_tac 1);
+ "isUCont f ==> isCont f x";
+by (Clarify_tac 1);
+by (dtac spec 1);
+by (Blast_tac 1);
qed "isUCont_isCont";
Goalw [isNSUCont_def,isUCont_def,approx_def]
@@ -690,17 +692,17 @@
by (Ultra_tac 1);
qed "isUCont_isNSUCont";
-Goal "\\<forall>s. Numeral0 < s --> (\\<exists>z y. abs (z + - y) < s & r \\<le> abs (f z + -f y)) \
+Goal "\\<forall>s. 0 < s --> (\\<exists>z y. abs (z + - y) < s & r \\<le> abs (f z + -f y)) \
\ ==> \\<forall>n::nat. \\<exists>z y. \
\ abs(z + -y) < inverse(real(Suc n)) & \
\ r \\<le> abs(f z + -f y)";
by (Clarify_tac 1);
by (cut_inst_tac [("n1","n")]
- (real_of_nat_Suc_gt_zero RS rename_numerals real_inverse_gt_zero) 1);
+ (real_of_nat_Suc_gt_zero RS real_inverse_gt_0) 1);
by Auto_tac;
val lemma_LIMu = result();
-Goal "\\<forall>s. Numeral0 < s --> (\\<exists>z y. abs (z + - y) < s & r \\<le> abs (f z + -f y)) \
+Goal "\\<forall>s. 0 < s --> (\\<exists>z y. abs (z + - y) < s & r \\<le> abs (f z + -f y)) \
\ ==> \\<exists>X Y. \\<forall>n::nat. \
\ abs(X n + -(Y n)) < inverse(real(Suc n)) & \
\ r \\<le> abs(f (X n) + -f (Y n))";
@@ -745,23 +747,23 @@
Derivatives
------------------------------------------------------------------*)
Goalw [deriv_def]
- "(DERIV f x :> D) = ((%h. (f(x + h) + - f(x))/h) -- Numeral0 --> D)";
+ "(DERIV f x :> D) = ((%h. (f(x + h) + - f(x))/h) -- 0 --> D)";
by (Blast_tac 1);
qed "DERIV_iff";
Goalw [deriv_def]
- "(DERIV f x :> D) = ((%h. (f(x + h) + - f(x))/h) -- Numeral0 --NS> D)";
+ "(DERIV f x :> D) = ((%h. (f(x + h) + - f(x))/h) -- 0 --NS> D)";
by (simp_tac (simpset() addsimps [LIM_NSLIM_iff]) 1);
qed "DERIV_NS_iff";
Goalw [deriv_def]
"DERIV f x :> D \
-\ ==> (%h. (f(x + h) + - f(x))/h) -- Numeral0 --> D";
+\ ==> (%h. (f(x + h) + - f(x))/h) -- 0 --> D";
by (Blast_tac 1);
qed "DERIVD";
Goalw [deriv_def] "DERIV f x :> D ==> \
-\ (%h. (f(x + h) + - f(x))/h) -- Numeral0 --NS> D";
+\ (%h. (f(x + h) + - f(x))/h) -- 0 --NS> D";
by (asm_full_simp_tac (simpset() addsimps [LIM_NSLIM_iff]) 1);
qed "NS_DERIVD";
@@ -809,7 +811,7 @@
-------------------------------------------------------*)
Goalw [LIM_def]
- "((%h. (f(a + h) + - f(a))/h) -- Numeral0 --> D) = \
+ "((%h. (f(a + h) + - f(a))/h) -- 0 --> D) = \
\ ((%x. (f(x) + -f(a)) / (x + -a)) -- a --> D)";
by Safe_tac;
by (ALLGOALS(dtac spec));
@@ -836,8 +838,8 @@
(*--- first equivalence ---*)
Goalw [nsderiv_def,NSLIM_def]
- "(NSDERIV f x :> D) = ((%h. (f(x + h) + - f(x))/h) -- Numeral0 --NS> D)";
-by (auto_tac (claset(), simpset() addsimps [hypreal_of_real_zero]));
+ "(NSDERIV f x :> D) = ((%h. (f(x + h) + - f(x))/h) -- 0 --NS> D)";
+by Auto_tac;
by (dres_inst_tac [("x","xa")] bspec 1);
by (rtac ccontr 3);
by (dres_inst_tac [("x","h")] spec 3);
@@ -868,8 +870,7 @@
\ (*f* (%z. f z - f x)) u \\<approx> hypreal_of_real D * (u - hypreal_of_real x))";
by (auto_tac (claset(), simpset() addsimps [NSDERIV_iff2]));
by (case_tac "u = hypreal_of_real x" 1);
-by (auto_tac (claset(),
- simpset() addsimps [hypreal_diff_def, hypreal_of_real_zero]));
+by (auto_tac (claset(), simpset() addsimps [hypreal_diff_def]));
by (dres_inst_tac [("x","u")] spec 1);
by Auto_tac;
by (dres_inst_tac [("c","u - hypreal_of_real x"),("b","hypreal_of_real D")]
@@ -956,12 +957,12 @@
------------------------*)
(* use simple constant nslimit theorem *)
-Goal "(NSDERIV (%x. k) x :> Numeral0)";
+Goal "(NSDERIV (%x. k) x :> 0)";
by (simp_tac (simpset() addsimps [NSDERIV_NSLIM_iff]) 1);
qed "NSDERIV_const";
Addsimps [NSDERIV_const];
-Goal "(DERIV (%x. k) x :> Numeral0)";
+Goal "(DERIV (%x. k) x :> 0)";
by (simp_tac (simpset() addsimps [NSDERIV_DERIV_iff RS sym]) 1);
qed "DERIV_const";
Addsimps [DERIV_const];
@@ -1000,7 +1001,7 @@
Goal "[| (x + y) / z = hypreal_of_real D + yb; z \\<noteq> 0; \
\ z \\<in> Infinitesimal; yb \\<in> Infinitesimal |] \
-\ ==> x + y \\<approx> Numeral0";
+\ ==> x + y \\<approx> 0";
by (forw_inst_tac [("c1","z")] (hypreal_mult_right_cancel RS iffD2) 1
THEN assume_tac 1);
by (thin_tac "(x + y) / z = hypreal_of_real D + yb" 1);
@@ -1015,8 +1016,7 @@
by (asm_full_simp_tac (simpset() addsimps [NSDERIV_NSLIM_iff, NSLIM_def]) 1);
by (REPEAT (Step_tac 1));
by (auto_tac (claset(),
- simpset() addsimps [starfun_lambda_cancel, hypreal_of_real_zero,
- lemma_nsderiv1]));
+ simpset() addsimps [starfun_lambda_cancel, lemma_nsderiv1]));
by (simp_tac (simpset() addsimps [hypreal_add_divide_distrib]) 1);
by (REPEAT(dtac (bex_Infinitesimal_iff2 RS iffD2) 1));
by (auto_tac (claset(),
@@ -1127,13 +1127,13 @@
qed "incrementI2";
(* The Increment theorem -- Keisler p. 65 *)
-Goal "[| NSDERIV f x :> D; h \\<in> Infinitesimal; h \\<noteq> Numeral0 |] \
+Goal "[| NSDERIV f x :> D; h \\<in> Infinitesimal; h \\<noteq> 0 |] \
\ ==> \\<exists>e \\<in> Infinitesimal. increment f x h = hypreal_of_real(D)*h + e*h";
by (forw_inst_tac [("h","h")] incrementI2 1 THEN rewtac nsderiv_def);
by (dtac bspec 1 THEN Auto_tac);
by (dtac (bex_Infinitesimal_iff2 RS iffD2) 1 THEN Step_tac 1);
by (forw_inst_tac [("b1","hypreal_of_real(D) + y")]
- (rename_numerals (hypreal_mult_right_cancel RS iffD2)) 1);
+ ((hypreal_mult_right_cancel RS iffD2)) 1);
by (thin_tac "((*f* f) (hypreal_of_real(x) + h) + \
\ - hypreal_of_real (f x)) / h = hypreal_of_real(D) + y" 2);
by (assume_tac 1);
@@ -1143,15 +1143,15 @@
simpset() addsimps [hypreal_add_mult_distrib]));
qed "increment_thm";
-Goal "[| NSDERIV f x :> D; h \\<approx> Numeral0; h \\<noteq> Numeral0 |] \
+Goal "[| NSDERIV f x :> D; h \\<approx> 0; h \\<noteq> 0 |] \
\ ==> \\<exists>e \\<in> Infinitesimal. increment f x h = \
\ hypreal_of_real(D)*h + e*h";
by (blast_tac (claset() addSDs [mem_infmal_iff RS iffD2]
addSIs [increment_thm]) 1);
qed "increment_thm2";
-Goal "[| NSDERIV f x :> D; h \\<approx> Numeral0; h \\<noteq> Numeral0 |] \
-\ ==> increment f x h \\<approx> Numeral0";
+Goal "[| NSDERIV f x :> D; h \\<approx> 0; h \\<noteq> 0 |] \
+\ ==> increment f x h \\<approx> 0";
by (dtac increment_thm2 1 THEN auto_tac (claset() addSIs
[Infinitesimal_HFinite_mult2,HFinite_add],simpset() addsimps
[hypreal_add_mult_distrib RS sym,mem_infmal_iff RS sym]));
@@ -1172,16 +1172,16 @@
"[| NSDERIV g x :> D; \
\ (*f* g) (hypreal_of_real(x) + xa) = hypreal_of_real(g x);\
\ xa \\<in> Infinitesimal;\
-\ xa \\<noteq> Numeral0 \
-\ |] ==> D = Numeral0";
+\ xa \\<noteq> 0 \
+\ |] ==> D = 0";
by (dtac bspec 1);
by Auto_tac;
qed "NSDERIV_zero";
(* can be proved differently using NSLIM_isCont_iff *)
Goalw [nsderiv_def]
- "[| NSDERIV f x :> D; h \\<in> Infinitesimal; h \\<noteq> Numeral0 |] \
-\ ==> (*f* f) (hypreal_of_real(x) + h) + -hypreal_of_real(f x) \\<approx> Numeral0";
+ "[| NSDERIV f x :> D; h \\<in> Infinitesimal; h \\<noteq> 0 |] \
+\ ==> (*f* f) (hypreal_of_real(x) + h) + -hypreal_of_real(f x) \\<approx> 0";
by (asm_full_simp_tac (simpset() addsimps
[mem_infmal_iff RS sym]) 1);
by (rtac Infinitesimal_ratio 1);
@@ -1214,12 +1214,12 @@
----------------- \\<approx> Db
h
--------------------------------------------------------------*)
-Goal "[| NSDERIV g x :> Db; xa \\<in> Infinitesimal; xa \\<noteq> Numeral0 |] \
+Goal "[| NSDERIV g x :> Db; xa \\<in> Infinitesimal; xa \\<noteq> 0 |] \
\ ==> ((*f* g) (hypreal_of_real(x) + xa) + - hypreal_of_real(g x)) / xa \
\ \\<approx> hypreal_of_real(Db)";
by (auto_tac (claset(),
simpset() addsimps [NSDERIV_NSLIM_iff, NSLIM_def,
- hypreal_of_real_zero, mem_infmal_iff, starfun_lambda_cancel]));
+ mem_infmal_iff, starfun_lambda_cancel]));
qed "NSDERIVD2";
Goal "(z::hypreal) \\<noteq> 0 ==> x*y = (x*inverse(z))*(z*y)";
@@ -1232,14 +1232,13 @@
Goal "[| NSDERIV f (g x) :> Da; NSDERIV g x :> Db |] \
\ ==> NSDERIV (f o g) x :> Da * Db";
by (asm_simp_tac (simpset() addsimps [NSDERIV_NSLIM_iff,
- NSLIM_def,hypreal_of_real_zero,mem_infmal_iff RS sym]) 1 THEN Step_tac 1);
+ NSLIM_def,mem_infmal_iff RS sym]) 1 THEN Step_tac 1);
by (forw_inst_tac [("f","g")] NSDERIV_approx 1);
by (auto_tac (claset(),
simpset() addsimps [starfun_lambda_cancel2, starfun_o RS sym]));
by (case_tac "(*f* g) (hypreal_of_real(x) + xa) = hypreal_of_real (g x)" 1);
by (dres_inst_tac [("g","g")] NSDERIV_zero 1);
-by (auto_tac (claset(),
- simpset() addsimps [hypreal_divide_def, hypreal_of_real_zero]));
+by (auto_tac (claset(), simpset() addsimps [hypreal_divide_def]));
by (res_inst_tac [("z1","(*f* g) (hypreal_of_real(x) + xa) + -hypreal_of_real (g x)"),
("y1","inverse xa")] (lemma_chain RS ssubst) 1);
by (etac (hypreal_not_eq_minus_iff RS iffD1) 1);
@@ -1266,16 +1265,14 @@
(*------------------------------------------------------------------
Differentiation of natural number powers
------------------------------------------------------------------*)
-Goal "NSDERIV (%x. x) x :> Numeral1";
+Goal "NSDERIV (%x. x) x :> 1";
by (auto_tac (claset(),
- simpset() addsimps [NSDERIV_NSLIM_iff,
- NSLIM_def ,starfun_Id, hypreal_of_real_zero,
- hypreal_of_real_one]));
+ simpset() addsimps [NSDERIV_NSLIM_iff, NSLIM_def ,starfun_Id]));
qed "NSDERIV_Id";
Addsimps [NSDERIV_Id];
(*derivative of the identity function*)
-Goal "DERIV (%x. x) x :> Numeral1";
+Goal "DERIV (%x. x) x :> 1";
by (simp_tac (simpset() addsimps [NSDERIV_DERIV_iff RS sym]) 1);
qed "DERIV_Id";
Addsimps [DERIV_Id];
@@ -1314,9 +1311,9 @@
Power of -1
---------------------------------------------------------------*)
-(*Can't get rid of x \\<noteq> Numeral0 because it isn't continuous at zero*)
+(*Can't get rid of x \\<noteq> 0 because it isn't continuous at zero*)
Goalw [nsderiv_def]
- "x \\<noteq> Numeral0 ==> NSDERIV (%x. inverse(x)) x :> (- (inverse x ^ Suc (Suc 0)))";
+ "x \\<noteq> 0 ==> NSDERIV (%x. inverse(x)) x :> (- (inverse x ^ Suc (Suc 0)))";
by (rtac ballI 1 THEN Asm_full_simp_tac 1 THEN Step_tac 1);
by (forward_tac [Infinitesimal_add_not_zero] 1);
by (asm_full_simp_tac (simpset() addsimps [hypreal_add_commute]) 2);
@@ -1345,7 +1342,7 @@
qed "NSDERIV_inverse";
-Goal "x \\<noteq> Numeral0 ==> DERIV (%x. inverse(x)) x :> (-(inverse x ^ Suc (Suc 0)))";
+Goal "x \\<noteq> 0 ==> DERIV (%x. inverse(x)) x :> (-(inverse x ^ Suc (Suc 0)))";
by (asm_simp_tac (simpset() addsimps [NSDERIV_inverse,
NSDERIV_DERIV_iff RS sym] delsimps [realpow_Suc]) 1);
qed "DERIV_inverse";
@@ -1353,7 +1350,7 @@
(*--------------------------------------------------------------
Derivative of inverse
-------------------------------------------------------------*)
-Goal "[| DERIV f x :> d; f(x) \\<noteq> Numeral0 |] \
+Goal "[| DERIV f x :> d; f(x) \\<noteq> 0 |] \
\ ==> DERIV (%x. inverse(f x)) x :> (- (d * inverse(f(x) ^ Suc (Suc 0))))";
by (rtac (real_mult_commute RS subst) 1);
by (asm_simp_tac (simpset() addsimps [real_minus_mult_eq1,
@@ -1363,7 +1360,7 @@
by (blast_tac (claset() addSIs [DERIV_chain,DERIV_inverse]) 1);
qed "DERIV_inverse_fun";
-Goal "[| NSDERIV f x :> d; f(x) \\<noteq> Numeral0 |] \
+Goal "[| NSDERIV f x :> d; f(x) \\<noteq> 0 |] \
\ ==> NSDERIV (%x. inverse(f x)) x :> (- (d * inverse(f(x) ^ Suc (Suc 0))))";
by (asm_full_simp_tac (simpset() addsimps [NSDERIV_DERIV_iff,
DERIV_inverse_fun] delsimps [realpow_Suc]) 1);
@@ -1372,7 +1369,7 @@
(*--------------------------------------------------------------
Derivative of quotient
-------------------------------------------------------------*)
-Goal "[| DERIV f x :> d; DERIV g x :> e; g(x) \\<noteq> Numeral0 |] \
+Goal "[| DERIV f x :> d; DERIV g x :> e; g(x) \\<noteq> 0 |] \
\ ==> DERIV (%y. f(y) / (g y)) x :> (d*g(x) + -(e*f(x))) / (g(x) ^ Suc (Suc 0))";
by (dres_inst_tac [("f","g")] DERIV_inverse_fun 1);
by (dtac DERIV_mult 2);
@@ -1384,7 +1381,7 @@
real_minus_mult_eq2 RS sym]) 1);
qed "DERIV_quotient";
-Goal "[| NSDERIV f x :> d; DERIV g x :> e; g(x) \\<noteq> Numeral0 |] \
+Goal "[| NSDERIV f x :> d; DERIV g x :> e; g(x) \\<noteq> 0 |] \
\ ==> NSDERIV (%y. f(y) / (g y)) x :> (d*g(x) \
\ + -(e*f(x))) / (g(x) ^ Suc (Suc 0))";
by (asm_full_simp_tac (simpset() addsimps [NSDERIV_DERIV_iff,
@@ -1401,7 +1398,7 @@
by (res_inst_tac
[("x","%z. if z = x then l else (f(z) - f(x)) / (z - x)")] exI 1);
by (auto_tac (claset(),simpset() addsimps [real_mult_assoc,
- ARITH_PROVE "z \\<noteq> x ==> z - x \\<noteq> (Numeral0::real)"]));
+ ARITH_PROVE "z \\<noteq> x ==> z - x \\<noteq> (0::real)"]));
by (auto_tac (claset(),simpset() addsimps [isCont_iff,DERIV_iff]));
by (ALLGOALS(rtac (LIM_equal RS iffD1)));
by (auto_tac (claset(),simpset() addsimps [real_diff_def,real_mult_assoc]));
@@ -1466,7 +1463,7 @@
by (dtac real_less_sum_gt_zero 1);
by (dres_inst_tac [("x","f n + - lim f")] spec 1);
by Safe_tac;
-by (dres_inst_tac [("P","%na. no\\<le>na --> ?Q na"),("x","no + n")] spec 2);
+by (dres_inst_tac [("P","%na. no\\<le>na --> ?Q na"),("x","no + n")] spec 1);
by Auto_tac;
by (subgoal_tac "lim f \\<le> f(no + n)" 1);
by (induct_tac "no" 2);
@@ -1511,7 +1508,7 @@
Goal "[| \\<forall>n. f(n) \\<le> f(Suc n); \
\ \\<forall>n. g(Suc n) \\<le> g(n); \
\ \\<forall>n. f(n) \\<le> g(n); \
-\ (%n. f(n) - g(n)) ----> Numeral0 |] \
+\ (%n. f(n) - g(n)) ----> 0 |] \
\ ==> \\<exists>l. ((\\<forall>n. f(n) \\<le> l) & f ----> l) & \
\ ((\\<forall>n. l \\<le> g(n)) & g ----> l)";
by (dtac lemma_nest 1 THEN Auto_tac);
@@ -1546,7 +1543,7 @@
Goal "((x::real) = y / (2 * z)) = (2 * x = y/z)";
by Auto_tac;
-by (dres_inst_tac [("f","%u. (Numeral1/2)*u")] arg_cong 1);
+by (dres_inst_tac [("f","%u. (1/2)*u")] arg_cong 1);
by Auto_tac;
qed "eq_divide_2_times_iff";
@@ -1589,7 +1586,7 @@
Goal "[| \\<forall>a b c. P(a,b) & P(b,c) & a \\<le> b & b \\<le> c --> P(a,c); \
-\ \\<forall>x. \\<exists>d::real. Numeral0 < d & \
+\ \\<forall>x. \\<exists>d::real. 0 < d & \
\ (\\<forall>a b. a \\<le> x & x \\<le> b & (b - a) < d --> P(a,b)); \
\ a \\<le> b |] \
\ ==> P(a,b)";
@@ -1604,8 +1601,8 @@
by (rename_tac "l" 1);
by (dres_inst_tac [("x","l")] spec 1 THEN Clarify_tac 1);
by (rewtac LIMSEQ_def);
-by (dres_inst_tac [("P", "%r. Numeral0<r --> ?Q r"), ("x","d/2")] spec 1);
-by (dres_inst_tac [("P", "%r. Numeral0<r --> ?Q r"), ("x","d/2")] spec 1);
+by (dres_inst_tac [("P", "%r. 0<r --> ?Q r"), ("x","d/2")] spec 1);
+by (dres_inst_tac [("P", "%r. 0<r --> ?Q r"), ("x","d/2")] spec 1);
by (dtac real_less_half_sum 1);
by Safe_tac;
(*linear arithmetic bug if we just use Asm_simp_tac*)
@@ -1626,7 +1623,7 @@
Goal "((\\<forall>a b c. (a \\<le> b & b \\<le> c & P(a,b) & P(b,c)) --> P(a,c)) & \
-\ (\\<forall>x. \\<exists>d::real. Numeral0 < d & \
+\ (\\<forall>x. \\<exists>d::real. 0 < d & \
\ (\\<forall>a b. a \\<le> x & x \\<le> b & (b - a) < d --> P(a,b)))) \
\ --> (\\<forall>a b. a \\<le> b --> P(a,b))";
by (Clarify_tac 1);
@@ -1654,14 +1651,14 @@
by (rtac ccontr 1);
by (subgoal_tac "a \\<le> x & x \\<le> b" 1);
by (Asm_full_simp_tac 2);
-by (dres_inst_tac [("P", "%d. Numeral0<d --> ?P d"),("x","Numeral1")] spec 2);
+by (dres_inst_tac [("P", "%d. 0<d --> ?P d"),("x","1")] spec 2);
by (Step_tac 2);
by (Asm_full_simp_tac 2);
by (Asm_full_simp_tac 2);
by (REPEAT(blast_tac (claset() addIs [order_trans]) 2));
by (REPEAT(dres_inst_tac [("x","x")] spec 1));
by (Asm_full_simp_tac 1);
-by (dres_inst_tac [("P", "%r. ?P r --> (\\<exists>s. Numeral0<s & ?Q r s)"),
+by (dres_inst_tac [("P", "%r. ?P r --> (\\<exists>s. 0<s & ?Q r s)"),
("x","abs(y - f x)")] spec 1);
by Safe_tac;
by (asm_full_simp_tac (simpset() addsimps []) 1);
@@ -1738,15 +1735,15 @@
by (cut_inst_tac [("x","xb"),("y","xa")] linorder_linear 1);
by (Force_tac 1);
by (case_tac "a \\<le> x & x \\<le> b" 1);
-by (res_inst_tac [("x","Numeral1")] exI 2);
+by (res_inst_tac [("x","1")] exI 2);
by (Force_tac 2);
by (asm_full_simp_tac (simpset() addsimps [LIM_def,isCont_iff]) 1);
by (dres_inst_tac [("x","x")] spec 1 THEN Auto_tac);
by (thin_tac "\\<forall>M. \\<exists>x. a \\<le> x & x \\<le> b & ~ f x \\<le> M" 1);
-by (dres_inst_tac [("x","Numeral1")] spec 1);
+by (dres_inst_tac [("x","1")] spec 1);
by Auto_tac;
by (res_inst_tac [("x","s")] exI 1 THEN Clarify_tac 1);
-by (res_inst_tac [("x","abs(f x) + Numeral1")] exI 1 THEN Clarify_tac 1);
+by (res_inst_tac [("x","abs(f x) + 1")] exI 1 THEN Clarify_tac 1);
by (dres_inst_tac [("x","xa - x")] spec 1 THEN Safe_tac);
by (arith_tac 1);
by (arith_tac 1);
@@ -1803,23 +1800,23 @@
"\\<exists>k. \\<forall>x. a \\<le> x & x \\<le> b --> (%x. inverse(M - (f x))) x \\<le> k" 1);
by (rtac isCont_bounded 2);
by Safe_tac;
-by (subgoal_tac "\\<forall>x. a \\<le> x & x \\<le> b --> Numeral0 < inverse(M - f(x))" 1);
+by (subgoal_tac "\\<forall>x. a \\<le> x & x \\<le> b --> 0 < inverse(M - f(x))" 1);
by (Asm_full_simp_tac 1);
by Safe_tac;
by (asm_full_simp_tac (simpset() addsimps [real_less_diff_eq]) 2);
by (subgoal_tac
- "\\<forall>x. a \\<le> x & x \\<le> b --> (%x. inverse(M - (f x))) x < (k + Numeral1)" 1);
+ "\\<forall>x. a \\<le> x & x \\<le> b --> (%x. inverse(M - (f x))) x < (k + 1)" 1);
by Safe_tac;
by (res_inst_tac [("y","k")] order_le_less_trans 2);
by (asm_full_simp_tac (simpset() addsimps [real_zero_less_one]) 3);
by (Asm_full_simp_tac 2);
by (subgoal_tac "\\<forall>x. a \\<le> x & x \\<le> b --> \
-\ inverse(k + Numeral1) < inverse((%x. inverse(M - (f x))) x)" 1);
+\ inverse(k + 1) < inverse((%x. inverse(M - (f x))) x)" 1);
by Safe_tac;
by (rtac real_inverse_less_swap 2);
by (ALLGOALS Asm_full_simp_tac);
by (dres_inst_tac [("P", "%N. N<M --> ?Q N"),
- ("x","M - inverse(k + Numeral1)")] spec 1);
+ ("x","M - inverse(k + 1)")] spec 1);
by (Step_tac 1 THEN dtac real_leI 1);
by (dtac (real_le_diff_eq RS iffD1) 1);
by (REPEAT(dres_inst_tac [("x","a")] spec 1));
@@ -1879,11 +1876,11 @@
(*----------------------------------------------------------------------------*)
Goalw [deriv_def,LIM_def]
- "[| DERIV f x :> l; Numeral0 < l |] ==> \
-\ \\<exists>d. Numeral0 < d & (\\<forall>h. Numeral0 < h & h < d --> f(x) < f(x + h))";
+ "[| DERIV f x :> l; 0 < l |] \
+\ ==> \\<exists>d. 0 < d & (\\<forall>h. 0 < h & h < d --> f(x) < f(x + h))";
by (dtac spec 1 THEN Auto_tac);
by (res_inst_tac [("x","s")] exI 1 THEN Auto_tac);
-by (subgoal_tac "Numeral0 < l*h" 1);
+by (subgoal_tac "0 < l*h" 1);
by (asm_full_simp_tac (simpset() addsimps [real_0_less_mult_iff]) 2);
by (dres_inst_tac [("x","h")] spec 1);
by (asm_full_simp_tac
@@ -1893,11 +1890,11 @@
qed "DERIV_left_inc";
Goalw [deriv_def,LIM_def]
- "[| DERIV f x :> l; l < Numeral0 |] ==> \
-\ \\<exists>d. Numeral0 < d & (\\<forall>h. Numeral0 < h & h < d --> f(x) < f(x - h))";
+ "[| DERIV f x :> l; l < 0 |] ==> \
+\ \\<exists>d. 0 < d & (\\<forall>h. 0 < h & h < d --> f(x) < f(x - h))";
by (dres_inst_tac [("x","-l")] spec 1 THEN Auto_tac);
by (res_inst_tac [("x","s")] exI 1 THEN Auto_tac);
-by (subgoal_tac "l*h < Numeral0" 1);
+by (subgoal_tac "l*h < 0" 1);
by (asm_full_simp_tac (simpset() addsimps [real_mult_less_0_iff]) 2);
by (dres_inst_tac [("x","-h")] spec 1);
by (asm_full_simp_tac
@@ -1905,7 +1902,7 @@
pos_real_less_divide_eq,
symmetric real_diff_def]
addsplits [split_if_asm]) 1);
-by (subgoal_tac "Numeral0 < (f (x - h) - f x)/h" 1);
+by (subgoal_tac "0 < (f (x - h) - f x)/h" 1);
by (arith_tac 2);
by (asm_full_simp_tac
(simpset() addsimps [pos_real_less_divide_eq]) 1);
@@ -1913,9 +1910,9 @@
Goal "[| DERIV f x :> l; \
-\ \\<exists>d. Numeral0 < d & (\\<forall>y. abs(x - y) < d --> f(y) \\<le> f(x)) |] \
-\ ==> l = Numeral0";
-by (res_inst_tac [("R1.0","l"),("R2.0","Numeral0")] real_linear_less2 1);
+\ \\<exists>d. 0 < d & (\\<forall>y. abs(x - y) < d --> f(y) \\<le> f(x)) |] \
+\ ==> l = 0";
+by (res_inst_tac [("R1.0","l"),("R2.0","0")] real_linear_less2 1);
by Safe_tac;
by (dtac DERIV_left_dec 1);
by (dtac DERIV_left_inc 3);
@@ -1933,8 +1930,8 @@
(*----------------------------------------------------------------------------*)
Goal "[| DERIV f x :> l; \
-\ \\<exists>d::real. Numeral0 < d & (\\<forall>y. abs(x - y) < d --> f(x) \\<le> f(y)) |] \
-\ ==> l = Numeral0";
+\ \\<exists>d::real. 0 < d & (\\<forall>y. abs(x - y) < d --> f(x) \\<le> f(y)) |] \
+\ ==> l = 0";
by (dtac (DERIV_minus RS DERIV_local_max) 1);
by Auto_tac;
qed "DERIV_local_min";
@@ -1944,8 +1941,8 @@
(*----------------------------------------------------------------------------*)
Goal "[| DERIV f x :> l; \
-\ \\<exists>d. Numeral0 < d & (\\<forall>y. abs(x - y) < d --> f(x) = f(y)) |] \
-\ ==> l = Numeral0";
+\ \\<exists>d. 0 < d & (\\<forall>y. abs(x - y) < d --> f(x) = f(y)) |] \
+\ ==> l = 0";
by (auto_tac (claset() addSDs [DERIV_local_max],simpset()));
qed "DERIV_local_const";
@@ -1954,7 +1951,7 @@
(*----------------------------------------------------------------------------*)
Goal "[| a < x; x < b |] ==> \
-\ \\<exists>d::real. Numeral0 < d & (\\<forall>y. abs(x - y) < d --> a < y & y < b)";
+\ \\<exists>d::real. 0 < d & (\\<forall>y. abs(x - y) < d --> a < y & y < b)";
by (simp_tac (simpset() addsimps [abs_interval_iff]) 1);
by (cut_inst_tac [("x","x - a"),("y","b - x")] linorder_linear 1);
by Safe_tac;
@@ -1965,7 +1962,7 @@
qed "lemma_interval_lt";
Goal "[| a < x; x < b |] ==> \
-\ \\<exists>d::real. Numeral0 < d & (\\<forall>y. abs(x - y) < d --> a \\<le> y & y \\<le> b)";
+\ \\<exists>d::real. 0 < d & (\\<forall>y. abs(x - y) < d --> a \\<le> y & y \\<le> b)";
by (dtac lemma_interval_lt 1);
by Auto_tac;
by (auto_tac (claset() addSIs [exI] ,simpset()));
@@ -1975,13 +1972,13 @@
Rolle's Theorem
If f is defined and continuous on the finite closed interval [a,b]
and differentiable a least on the open interval (a,b), and f(a) = f(b),
- then x0 \\<in> (a,b) such that f'(x0) = Numeral0
+ then x0 \\<in> (a,b) such that f'(x0) = 0
----------------------------------------------------------------------*)
Goal "[| a < b; f(a) = f(b); \
\ \\<forall>x. a \\<le> x & x \\<le> b --> isCont f x; \
\ \\<forall>x. a < x & x < b --> f differentiable x \
-\ |] ==> \\<exists>z. a < z & z < b & DERIV f z :> Numeral0";
+\ |] ==> \\<exists>z. a < z & z < b & DERIV f z :> 0";
by (ftac (order_less_imp_le RS isCont_eq_Ub) 1);
by (EVERY1[assume_tac,Step_tac]);
by (ftac (order_less_imp_le RS isCont_eq_Lb) 1);
@@ -1992,7 +1989,7 @@
by (EVERY1[assume_tac,etac exE]);
by (res_inst_tac [("x","x")] exI 1 THEN Asm_full_simp_tac 1);
by (subgoal_tac "(\\<exists>l. DERIV f x :> l) & \
-\ (\\<exists>d. Numeral0 < d & (\\<forall>y. abs(x - y) < d --> f(y) \\<le> f(x)))" 1);
+\ (\\<exists>d. 0 < d & (\\<forall>y. abs(x - y) < d --> f(y) \\<le> f(x)))" 1);
by (Clarify_tac 1 THEN rtac conjI 2);
by (blast_tac (claset() addIs [differentiableD]) 2);
by (Blast_tac 2);
@@ -2004,7 +2001,7 @@
by (EVERY1[assume_tac,etac exE]);
by (res_inst_tac [("x","xa")] exI 1 THEN Asm_full_simp_tac 1);
by (subgoal_tac "(\\<exists>l. DERIV f xa :> l) & \
-\ (\\<exists>d. Numeral0 < d & (\\<forall>y. abs(xa - y) < d --> f(xa) \\<le> f(y)))" 1);
+\ (\\<exists>d. 0 < d & (\\<forall>y. abs(xa - y) < d --> f(xa) \\<le> f(y)))" 1);
by (Clarify_tac 1 THEN rtac conjI 2);
by (blast_tac (claset() addIs [differentiableD]) 2);
by (Blast_tac 2);
@@ -2030,7 +2027,7 @@
by (forw_inst_tac [("a","a"),("x","r")] lemma_interval 1);
by (EVERY1[assume_tac, etac exE]);
by (subgoal_tac "(\\<exists>l. DERIV f r :> l) & \
-\ (\\<exists>d. Numeral0 < d & (\\<forall>y. abs(r - y) < d --> f(r) = f(y)))" 1);
+\ (\\<exists>d. 0 < d & (\\<forall>y. abs(r - y) < d --> f(r) = f(y)))" 1);
by (Clarify_tac 1 THEN rtac conjI 2);
by (blast_tac (claset() addIs [differentiableD]) 2);
by (EVERY1[ftac DERIV_local_const, Blast_tac, Blast_tac]);
@@ -2098,7 +2095,7 @@
Goal "[| a < b; \
\ \\<forall>x. a \\<le> x & x \\<le> b --> isCont f x; \
-\ \\<forall>x. a < x & x < b --> DERIV f x :> Numeral0 |] \
+\ \\<forall>x. a < x & x < b --> DERIV f x :> 0 |] \
\ ==> (f b = f a)";
by (dtac MVT 1 THEN assume_tac 1);
by (blast_tac (claset() addIs [differentiableI]) 1);
@@ -2108,7 +2105,7 @@
Goal "[| a < b; \
\ \\<forall>x. a \\<le> x & x \\<le> b --> isCont f x; \
-\ \\<forall>x. a < x & x < b --> DERIV f x :> Numeral0 |] \
+\ \\<forall>x. a < x & x < b --> DERIV f x :> 0 |] \
\ ==> \\<forall>x. a \\<le> x & x \\<le> b --> f x = f a";
by Safe_tac;
by (dres_inst_tac [("x","a")] order_le_imp_less_or_eq 1);
@@ -2119,13 +2116,13 @@
Goal "[| a < b; \
\ \\<forall>x. a \\<le> x & x \\<le> b --> isCont f x; \
-\ \\<forall>x. a < x & x < b --> DERIV f x :> Numeral0; \
+\ \\<forall>x. a < x & x < b --> DERIV f x :> 0; \
\ a \\<le> x; x \\<le> b |] \
\ ==> f x = f a";
by (blast_tac (claset() addDs [DERIV_isconst1]) 1);
qed "DERIV_isconst2";
-Goal "\\<forall>x. DERIV f x :> Numeral0 ==> f(x) = f(y)";
+Goal "\\<forall>x. DERIV f x :> 0 ==> f(x) = f(y)";
by (res_inst_tac [("R1.0","x"),("R2.0","y")] real_linear_less2 1);
by (rtac sym 1);
by (auto_tac (claset() addIs [DERIV_isCont,DERIV_isconst_end],simpset()));
@@ -2182,7 +2179,7 @@
(* maximum at an end point, not in the middle. *)
(* ------------------------------------------------------------------------ *)
-Goal "[|Numeral0 < d; \\<forall>z. abs(z - x) \\<le> d --> g(f z) = z; \
+Goal "[|0 < d; \\<forall>z. abs(z - x) \\<le> d --> g(f z) = z; \
\ \\<forall>z. abs(z - x) \\<le> d --> isCont f z |] \
\ ==> ~(\\<forall>z. abs(z - x) \\<le> d --> f(z) \\<le> f(x))";
by (rtac notI 1);
@@ -2221,7 +2218,7 @@
(* Similar version for lower bound *)
(* ------------------------------------------------------------------------ *)
-Goal "[|Numeral0 < d; \\<forall>z. abs(z - x) \\<le> d --> g(f z) = z; \
+Goal "[|0 < d; \\<forall>z. abs(z - x) \\<le> d --> g(f z) = z; \
\ \\<forall>z. abs(z - x) \\<le> d --> isCont f z |] \
\ ==> ~(\\<forall>z. abs(z - x) \\<le> d --> f(x) \\<le> f(z))";
by (auto_tac (claset() addSDs [(asm_full_simplify (simpset())
@@ -2234,14 +2231,12 @@
(* Also from John's theory *)
(* ------------------------------------------------------------------------ *)
-Addsimps [zero_eq_numeral_0,one_eq_numeral_1];
-
-val lemma_le = ARITH_PROVE "Numeral0 \\<le> (d::real) ==> -d \\<le> d";
+val lemma_le = ARITH_PROVE "0 \\<le> (d::real) ==> -d \\<le> d";
(* FIXME: awful proof - needs improvement *)
-Goal "[| Numeral0 < d; \\<forall>z. abs(z - x) \\<le> d --> g(f z) = z; \
+Goal "[| 0 < d; \\<forall>z. abs(z - x) \\<le> d --> g(f z) = z; \
\ \\<forall>z. abs(z - x) \\<le> d --> isCont f z |] \
-\ ==> \\<exists>e. Numeral0 < e & \
+\ ==> \\<exists>e. 0 < e & \
\ (\\<forall>y. \
\ abs(y - f(x)) \\<le> e --> \
\ (\\<exists>z. abs(z - x) \\<le> d & (f z = y)))";
@@ -2255,10 +2250,10 @@
by (Asm_full_simp_tac 2);
by (subgoal_tac "L < f x & f x < M" 1);
by Safe_tac;
-by (dres_inst_tac [("x","L")] (ARITH_PROVE "x < y ==> Numeral0 < y - (x::real)") 1);
-by (dres_inst_tac [("x","f x")] (ARITH_PROVE "x < y ==> Numeral0 < y - (x::real)") 1);
+by (dres_inst_tac [("x","L")] (ARITH_PROVE "x < y ==> 0 < y - (x::real)") 1);
+by (dres_inst_tac [("x","f x")] (ARITH_PROVE "x < y ==> 0 < y - (x::real)") 1);
by (dres_inst_tac [("d1.0","f x - L"),("d2.0","M - f x")]
- (rename_numerals real_lbound_gt_zero) 1);
+ (real_lbound_gt_zero) 1);
by Safe_tac;
by (res_inst_tac [("x","e")] exI 1);
by Safe_tac;
@@ -2284,12 +2279,12 @@
(* Continuity of inverse function *)
(* ------------------------------------------------------------------------ *)
-Goal "[| Numeral0 < d; \\<forall>z. abs(z - x) \\<le> d --> g(f(z)) = z; \
+Goal "[| 0 < d; \\<forall>z. abs(z - x) \\<le> d --> g(f(z)) = z; \
\ \\<forall>z. abs(z - x) \\<le> d --> isCont f z |] \
\ ==> isCont g (f x)";
by (simp_tac (simpset() addsimps [isCont_iff,LIM_def]) 1);
by Safe_tac;
-by (dres_inst_tac [("d1.0","r")] (rename_numerals real_lbound_gt_zero) 1);
+by (dres_inst_tac [("d1.0","r")] (real_lbound_gt_zero) 1);
by (assume_tac 1 THEN Step_tac 1);
by (subgoal_tac "\\<forall>z. abs(z - x) \\<le> e --> (g(f z) = z)" 1);
by (Force_tac 2);
--- a/src/HOL/Hyperreal/Lim.thy Thu Nov 01 21:12:13 2001 +0100
+++ b/src/HOL/Hyperreal/Lim.thy Fri Nov 02 17:55:24 2001 +0100
@@ -15,8 +15,8 @@
LIM :: [real=>real,real,real] => bool
("((_)/ -- (_)/ --> (_))" [60, 0, 60] 60)
"f -- a --> L ==
- ALL r. Numeral0 < r -->
- (EX s. Numeral0 < s & (ALL x. (x ~= a & (abs(x + -a) < s)
+ ALL r. 0 < r -->
+ (EX s. 0 < s & (ALL x. (x ~= a & (abs(x + -a) < s)
--> abs(f x + -L) < r)))"
NSLIM :: [real=>real,real,real] => bool
@@ -36,7 +36,7 @@
(* differentiation: D is derivative of function f at x *)
deriv:: [real=>real,real,real] => bool
("(DERIV (_)/ (_)/ :> (_))" [60, 0, 60] 60)
- "DERIV f x :> D == ((%h. (f(x + h) + -f(x))/h) -- Numeral0 --> D)"
+ "DERIV f x :> D == ((%h. (f(x + h) + -f(x))/h) -- 0 --> D)"
nsderiv :: [real=>real,real,real] => bool
("(NSDERIV (_)/ (_)/ :> (_))" [60, 0, 60] 60)
@@ -55,8 +55,8 @@
inc = (*f* f)(hypreal_of_real x + h) + -hypreal_of_real (f x))"
isUCont :: (real=>real) => bool
- "isUCont f == (ALL r. Numeral0 < r -->
- (EX s. Numeral0 < s & (ALL x y. abs(x + -y) < s
+ "isUCont f == (ALL r. 0 < r -->
+ (EX s. 0 < s & (ALL x y. abs(x + -y) < s
--> abs(f x + -f y) < r)))"
isNSUCont :: (real=>real) => bool
--- a/src/HOL/Hyperreal/NSA.ML Thu Nov 01 21:12:13 2001 +0100
+++ b/src/HOL/Hyperreal/NSA.ML Fri Nov 02 17:55:24 2001 +0100
@@ -69,6 +69,20 @@
qed "SReal_number_of";
Addsimps [SReal_number_of];
+(** As always with numerals, 0 and 1 are special cases **)
+
+Goal "(0::hypreal) : Reals";
+by (stac (hypreal_numeral_0_eq_0 RS sym) 1);
+by (rtac SReal_number_of 1);
+qed "Reals_0";
+Addsimps [Reals_0];
+
+Goal "(1::hypreal) : Reals";
+by (stac (hypreal_numeral_1_eq_1 RS sym) 1);
+by (rtac SReal_number_of 1);
+qed "Reals_1";
+Addsimps [Reals_1];
+
Goalw [hypreal_divide_def] "r : Reals ==> r/(number_of w::hypreal) : Reals";
by (blast_tac (claset() addSIs [SReal_number_of, SReal_mult,
SReal_inverse]) 1);
@@ -211,9 +225,9 @@
Goalw [SReal_def,HFinite_def] "Reals <= HFinite";
by Auto_tac;
-by (res_inst_tac [("x","Numeral1 + abs(hypreal_of_real r)")] exI 1);
+by (res_inst_tac [("x","1 + abs(hypreal_of_real r)")] exI 1);
by (auto_tac (claset(), simpset() addsimps [hypreal_of_real_hrabs]));
-by (res_inst_tac [("x","Numeral1 + abs r")] exI 1);
+by (res_inst_tac [("x","1 + abs r")] exI 1);
by (Simp_tac 1);
qed "SReal_subset_HFinite";
@@ -238,8 +252,22 @@
qed "HFinite_number_of";
Addsimps [HFinite_number_of];
-Goal "[|x : HFinite; y <= x; Numeral0 <= y |] ==> y: HFinite";
-by (case_tac "x <= Numeral0" 1);
+(** As always with numerals, 0 and 1 are special cases **)
+
+Goal "0 : HFinite";
+by (stac (hypreal_numeral_0_eq_0 RS sym) 1);
+by (rtac HFinite_number_of 1);
+qed "HFinite_0";
+Addsimps [HFinite_0];
+
+Goal "1 : HFinite";
+by (stac (hypreal_numeral_1_eq_1 RS sym) 1);
+by (rtac HFinite_number_of 1);
+qed "HFinite_1";
+Addsimps [HFinite_1];
+
+Goal "[|x : HFinite; y <= x; 0 <= y |] ==> y: HFinite";
+by (case_tac "x <= 0" 1);
by (dres_inst_tac [("y","x")] order_trans 1);
by (dtac hypreal_le_anti_sym 2);
by (auto_tac (claset() addSDs [not_hypreal_leE], simpset()));
@@ -251,11 +279,11 @@
Set of infinitesimals is a subring of the hyperreals
------------------------------------------------------------------*)
Goalw [Infinitesimal_def]
- "x : Infinitesimal ==> ALL r: Reals. Numeral0 < r --> abs x < r";
+ "x : Infinitesimal ==> ALL r: Reals. 0 < r --> abs x < r";
by Auto_tac;
qed "InfinitesimalD";
-Goalw [Infinitesimal_def] "Numeral0 : Infinitesimal";
+Goalw [Infinitesimal_def] "0 : Infinitesimal";
by (simp_tac (simpset() addsimps [hrabs_zero]) 1);
qed "Infinitesimal_zero";
AddIffs [Infinitesimal_zero];
@@ -264,7 +292,7 @@
by Auto_tac;
qed "hypreal_sum_of_halves";
-Goal "Numeral0 < r ==> Numeral0 < r/(2::hypreal)";
+Goal "0 < r ==> 0 < r/(2::hypreal)";
by Auto_tac;
qed "hypreal_half_gt_zero";
@@ -290,10 +318,10 @@
Goalw [Infinitesimal_def]
"[| x : Infinitesimal; y : Infinitesimal |] ==> (x * y) : Infinitesimal";
by Auto_tac;
-by (case_tac "y=Numeral0" 1);
-by (cut_inst_tac [("u","abs x"),("v","Numeral1"),("x","abs y"),("y","r")]
+by (case_tac "y=0" 1);
+by (cut_inst_tac [("u","abs x"),("v","1"),("x","abs y"),("y","r")]
hypreal_mult_less_mono 2);
-by Auto_tac;
+by Auto_tac;
qed "Infinitesimal_mult";
Goal "[| x : Infinitesimal; y : HFinite |] ==> (x * y) : Infinitesimal";
@@ -332,27 +360,27 @@
Goalw [HInfinite_def] "[|x: HInfinite;y: HInfinite|] ==> (x*y): HInfinite";
by Auto_tac;
-by (eres_inst_tac [("x","Numeral1")] ballE 1);
+by (eres_inst_tac [("x","1")] ballE 1);
by (eres_inst_tac [("x","r")] ballE 1);
by (case_tac "y=0" 1);
-by (cut_inst_tac [("x","Numeral1"),("y","abs x"),
+by (cut_inst_tac [("x","1"),("y","abs x"),
("u","r"),("v","abs y")] hypreal_mult_less_mono 2);
by (auto_tac (claset(), simpset() addsimps hypreal_mult_ac));
qed "HInfinite_mult";
Goalw [HInfinite_def]
- "[|x: HInfinite; Numeral0 <= y; Numeral0 <= x|] ==> (x + y): HInfinite";
+ "[|x: HInfinite; 0 <= y; 0 <= x|] ==> (x + y): HInfinite";
by (auto_tac (claset() addSIs [hypreal_add_zero_less_le_mono],
simpset() addsimps [hrabs_eqI1, hypreal_add_commute,
hypreal_le_add_order]));
qed "HInfinite_add_ge_zero";
-Goal "[|x: HInfinite; Numeral0 <= y; Numeral0 <= x|] ==> (y + x): HInfinite";
+Goal "[|x: HInfinite; 0 <= y; 0 <= x|] ==> (y + x): HInfinite";
by (auto_tac (claset() addSIs [HInfinite_add_ge_zero],
simpset() addsimps [hypreal_add_commute]));
qed "HInfinite_add_ge_zero2";
-Goal "[|x: HInfinite; Numeral0 < y; Numeral0 < x|] ==> (x + y): HInfinite";
+Goal "[|x: HInfinite; 0 < y; 0 < x|] ==> (x + y): HInfinite";
by (blast_tac (claset() addIs [HInfinite_add_ge_zero,
order_less_imp_le]) 1);
qed "HInfinite_add_gt_zero";
@@ -361,14 +389,14 @@
by Auto_tac;
qed "HInfinite_minus_iff";
-Goal "[|x: HInfinite; y <= Numeral0; x <= Numeral0|] ==> (x + y): HInfinite";
+Goal "[|x: HInfinite; y <= 0; x <= 0|] ==> (x + y): HInfinite";
by (dtac (HInfinite_minus_iff RS iffD2) 1);
by (rtac (HInfinite_minus_iff RS iffD1) 1);
by (auto_tac (claset() addIs [HInfinite_add_ge_zero],
simpset() addsimps [hypreal_minus_zero_le_iff]));
qed "HInfinite_add_le_zero";
-Goal "[|x: HInfinite; y < Numeral0; x < Numeral0|] ==> (x + y): HInfinite";
+Goal "[|x: HInfinite; y < 0; x < 0|] ==> (x + y): HInfinite";
by (blast_tac (claset() addIs [HInfinite_add_le_zero,
order_less_imp_le]) 1);
qed "HInfinite_add_lt_zero";
@@ -378,11 +406,11 @@
by (auto_tac (claset() addIs [HFinite_mult,HFinite_add], simpset()));
qed "HFinite_sum_squares";
-Goal "x ~: Infinitesimal ==> x ~= Numeral0";
+Goal "x ~: Infinitesimal ==> x ~= 0";
by Auto_tac;
qed "not_Infinitesimal_not_zero";
-Goal "x: HFinite - Infinitesimal ==> x ~= Numeral0";
+Goal "x: HFinite - Infinitesimal ==> x ~= 0";
by Auto_tac;
qed "not_Infinitesimal_not_zero2";
@@ -441,7 +469,7 @@
by (fast_tac (claset() addDs [not_Infinitesimal_mult]) 1);
qed "Infinitesimal_mult_disj";
-Goal "x: HFinite-Infinitesimal ==> x ~= Numeral0";
+Goal "x: HFinite-Infinitesimal ==> x ~= 0";
by (Blast_tac 1);
qed "HFinite_Infinitesimal_not_zero";
@@ -455,7 +483,7 @@
Goalw [Infinitesimal_def,HFinite_def]
"Infinitesimal <= HFinite";
by Auto_tac;
-by (res_inst_tac [("x","Numeral1")] bexI 1);
+by (res_inst_tac [("x","1")] bexI 1);
by Auto_tac;
qed "Infinitesimal_subset_HFinite";
@@ -474,15 +502,15 @@
----------------------------------------------------------------------*)
Goalw [Infinitesimal_def,approx_def]
- "(x:Infinitesimal) = (x @= Numeral0)";
+ "(x:Infinitesimal) = (x @= 0)";
by (Simp_tac 1);
qed "mem_infmal_iff";
-Goalw [approx_def]" (x @= y) = (x + -y @= Numeral0)";
+Goalw [approx_def]" (x @= y) = (x + -y @= 0)";
by (Simp_tac 1);
qed "approx_minus_iff";
-Goalw [approx_def]" (x @= y) = (-y + x @= Numeral0)";
+Goalw [approx_def]" (x @= y) = (-y + x @= 0)";
by (simp_tac (simpset() addsimps [hypreal_add_commute]) 1);
qed "approx_minus_iff2";
@@ -510,10 +538,48 @@
by (blast_tac (claset() addIs [approx_sym, approx_trans]) 1);
qed "approx_trans3";
-Goal "(number_of w @= x) = (x @= number_of w)";
+Goal "(number_of w @= x) = (x @= number_of w)";
by (blast_tac (claset() addIs [approx_sym]) 1);
qed "number_of_approx_reorient";
-Addsimps [number_of_approx_reorient];
+
+Goal "(0 @= x) = (x @= 0)";
+by (blast_tac (claset() addIs [approx_sym]) 1);
+qed "zero_approx_reorient";
+
+Goal "(1 @= x) = (x @= 1)";
+by (blast_tac (claset() addIs [approx_sym]) 1);
+qed "one_approx_reorient";
+
+
+(*** re-orientation, following HOL/Integ/Bin.ML
+ We re-orient x @=y where x is 0, 1 or a numeral, unless y is as well!
+ ***)
+
+(*reorientation simprules using ==, for the following simproc*)
+val meta_zero_approx_reorient = zero_approx_reorient RS eq_reflection;
+val meta_one_approx_reorient = one_approx_reorient RS eq_reflection;
+val meta_number_of_approx_reorient = number_of_approx_reorient RS eq_reflection;
+
+(*reorientation simplification procedure: reorients (polymorphic)
+ 0 = x, 1 = x, nnn = x provided x isn't 0, 1 or a numeral.*)
+fun reorient_proc sg _ (_ $ t $ u) =
+ case u of
+ Const("0", _) => None
+ | Const("1", _) => None
+ | Const("Numeral.number_of", _) $ _ => None
+ | _ => Some (case t of
+ Const("0", _) => meta_zero_approx_reorient
+ | Const("1", _) => meta_one_approx_reorient
+ | Const("Numeral.number_of", _) $ _ =>
+ meta_number_of_approx_reorient);
+
+val approx_reorient_simproc =
+ Bin_Simprocs.prep_simproc ("reorient_simproc",
+ Bin_Simprocs.prep_pats ["0@=x", "1@=x", "number_of w @= x"],
+ reorient_proc);
+
+Addsimprocs [approx_reorient_simproc];
+
Goal "(x-y : Infinitesimal) = (x @= y)";
by (auto_tac (claset(),
@@ -704,36 +770,36 @@
approx_hypreal_of_real_HFinite,HFinite_hypreal_of_real]) 1);
qed "approx_mult_hypreal_of_real";
-Goal "[| a: Reals; a ~= Numeral0; a*x @= Numeral0 |] ==> x @= Numeral0";
+Goal "[| a: Reals; a ~= 0; a*x @= 0 |] ==> x @= 0";
by (dtac (SReal_inverse RS (SReal_subset_HFinite RS subsetD)) 1);
by (auto_tac (claset() addDs [approx_mult2],
simpset() addsimps [hypreal_mult_assoc RS sym]));
qed "approx_SReal_mult_cancel_zero";
(* REM comments: newly added *)
-Goal "[| a: Reals; x @= Numeral0 |] ==> x*a @= Numeral0";
+Goal "[| a: Reals; x @= 0 |] ==> x*a @= 0";
by (auto_tac (claset() addDs [(SReal_subset_HFinite RS subsetD),
approx_mult1], simpset()));
qed "approx_mult_SReal1";
-Goal "[| a: Reals; x @= Numeral0 |] ==> a*x @= Numeral0";
+Goal "[| a: Reals; x @= 0 |] ==> a*x @= 0";
by (auto_tac (claset() addDs [(SReal_subset_HFinite RS subsetD),
approx_mult2], simpset()));
qed "approx_mult_SReal2";
-Goal "[|a : Reals; a ~= Numeral0 |] ==> (a*x @= Numeral0) = (x @= Numeral0)";
+Goal "[|a : Reals; a ~= 0 |] ==> (a*x @= 0) = (x @= 0)";
by (blast_tac (claset() addIs [approx_SReal_mult_cancel_zero,
approx_mult_SReal2]) 1);
qed "approx_mult_SReal_zero_cancel_iff";
Addsimps [approx_mult_SReal_zero_cancel_iff];
-Goal "[| a: Reals; a ~= Numeral0; a* w @= a*z |] ==> w @= z";
+Goal "[| a: Reals; a ~= 0; a* w @= a*z |] ==> w @= z";
by (dtac (SReal_inverse RS (SReal_subset_HFinite RS subsetD)) 1);
by (auto_tac (claset() addDs [approx_mult2],
simpset() addsimps [hypreal_mult_assoc RS sym]));
qed "approx_SReal_mult_cancel";
-Goal "[| a: Reals; a ~= Numeral0|] ==> (a* w @= a*z) = (w @= z)";
+Goal "[| a: Reals; a ~= 0|] ==> (a* w @= a*z) = (w @= z)";
by (auto_tac (claset() addSIs [approx_mult2,SReal_subset_HFinite RS subsetD]
addIs [approx_SReal_mult_cancel], simpset()));
qed "approx_SReal_mult_cancel_iff1";
@@ -754,63 +820,71 @@
-----------------------------------------------------------------*)
Goalw [Infinitesimal_def]
- "[| x: Reals; y: Infinitesimal; Numeral0 < x |] ==> y < x";
+ "[| x: Reals; y: Infinitesimal; 0 < x |] ==> y < x";
by (rtac (hrabs_ge_self RS order_le_less_trans) 1);
by Auto_tac;
qed "Infinitesimal_less_SReal";
-Goal "y: Infinitesimal ==> ALL r: Reals. Numeral0 < r --> y < r";
+Goal "y: Infinitesimal ==> ALL r: Reals. 0 < r --> y < r";
by (blast_tac (claset() addIs [Infinitesimal_less_SReal]) 1);
qed "Infinitesimal_less_SReal2";
Goalw [Infinitesimal_def]
- "[| Numeral0 < y; y: Reals|] ==> y ~: Infinitesimal";
+ "[| 0 < y; y: Reals|] ==> y ~: Infinitesimal";
by (auto_tac (claset(), simpset() addsimps [hrabs_def]));
qed "SReal_not_Infinitesimal";
-Goal "[| y < Numeral0; y : Reals |] ==> y ~: Infinitesimal";
+Goal "[| y < 0; y : Reals |] ==> y ~: Infinitesimal";
by (stac (Infinitesimal_minus_iff RS sym) 1);
by (rtac SReal_not_Infinitesimal 1);
by Auto_tac;
qed "SReal_minus_not_Infinitesimal";
-Goal "Reals Int Infinitesimal = {Numeral0}";
+Goal "Reals Int Infinitesimal = {0}";
by Auto_tac;
-by (cut_inst_tac [("x","x"),("y","Numeral0")] hypreal_linear 1);
+by (cut_inst_tac [("x","x"),("y","0")] hypreal_linear 1);
by (blast_tac (claset() addDs [SReal_not_Infinitesimal,
SReal_minus_not_Infinitesimal]) 1);
qed "SReal_Int_Infinitesimal_zero";
-Goal "[| x: Reals; x: Infinitesimal|] ==> x = Numeral0";
+Goal "[| x: Reals; x: Infinitesimal|] ==> x = 0";
by (cut_facts_tac [SReal_Int_Infinitesimal_zero] 1);
by (Blast_tac 1);
qed "SReal_Infinitesimal_zero";
-Goal "[| x : Reals; x ~= Numeral0 |] ==> x : HFinite - Infinitesimal";
+Goal "[| x : Reals; x ~= 0 |] ==> x : HFinite - Infinitesimal";
by (auto_tac (claset() addDs [SReal_Infinitesimal_zero,
SReal_subset_HFinite RS subsetD],
simpset()));
qed "SReal_HFinite_diff_Infinitesimal";
-Goal "hypreal_of_real x ~= Numeral0 ==> hypreal_of_real x : HFinite - Infinitesimal";
+Goal "hypreal_of_real x ~= 0 ==> hypreal_of_real x : HFinite - Infinitesimal";
by (rtac SReal_HFinite_diff_Infinitesimal 1);
by Auto_tac;
qed "hypreal_of_real_HFinite_diff_Infinitesimal";
-Goal "(hypreal_of_real x : Infinitesimal) = (x=Numeral0)";
-by (auto_tac (claset(), simpset() addsimps [hypreal_of_real_zero]));
+Goal "(hypreal_of_real x : Infinitesimal) = (x=0)";
+by Auto_tac;
by (rtac ccontr 1);
by (rtac (hypreal_of_real_HFinite_diff_Infinitesimal RS DiffD2) 1);
-by Auto_tac;
+by Auto_tac;
qed "hypreal_of_real_Infinitesimal_iff_0";
AddIffs [hypreal_of_real_Infinitesimal_iff_0];
-Goal "number_of w ~= (Numeral0::hypreal) ==> number_of w ~: Infinitesimal";
+Goal "number_of w ~= (0::hypreal) ==> number_of w ~: Infinitesimal";
by (fast_tac (claset() addDs [SReal_number_of RS SReal_Infinitesimal_zero]) 1);
qed "number_of_not_Infinitesimal";
Addsimps [number_of_not_Infinitesimal];
-Goal "[| y: Reals; x @= y; y~= Numeral0 |] ==> x ~= Numeral0";
+(*again: 1 is a special case, but not 0 this time*)
+Goal "1 ~: Infinitesimal";
+by (stac (hypreal_numeral_1_eq_1 RS sym) 1);
+by (rtac number_of_not_Infinitesimal 1);
+by (Simp_tac 1);
+qed "one_not_Infinitesimal";
+Addsimps [one_not_Infinitesimal];
+
+Goal "[| y: Reals; x @= y; y~= 0 |] ==> x ~= 0";
by (cut_inst_tac [("x","y")] hypreal_trichotomy 1);
by (Asm_full_simp_tac 1);
by (blast_tac (claset() addDs
@@ -828,7 +902,7 @@
(*The premise y~=0 is essential; otherwise x/y =0 and we lose the
HFinite premise.*)
-Goal "[| y ~= Numeral0; y: Infinitesimal; x/y : HFinite |] ==> x : Infinitesimal";
+Goal "[| y ~= 0; y: Infinitesimal; x/y : HFinite |] ==> x : Infinitesimal";
by (dtac Infinitesimal_HFinite_mult2 1);
by (assume_tac 1);
by (asm_full_simp_tac
@@ -859,6 +933,27 @@
qed "number_of_approx_iff";
Addsimps [number_of_approx_iff];
+(*And also for 0 @= #nn and 1 @= #nn, #nn @= 0 and #nn @= 1.*)
+Addsimps
+ (map (simplify (simpset()))
+ [inst "v" "Pls" number_of_approx_iff,
+ inst "v" "Pls BIT True" number_of_approx_iff,
+ inst "w" "Pls" number_of_approx_iff,
+ inst "w" "Pls BIT True" number_of_approx_iff]);
+
+
+Goal "~ (0 @= 1)";
+by (stac SReal_approx_iff 1);
+by Auto_tac;
+qed "not_0_approx_1";
+Addsimps [not_0_approx_1];
+
+Goal "~ (1 @= 0)";
+by (stac SReal_approx_iff 1);
+by Auto_tac;
+qed "not_1_approx_0";
+Addsimps [not_1_approx_0];
+
Goal "(hypreal_of_real k @= hypreal_of_real m) = (k = m)";
by Auto_tac;
by (rtac (inj_hypreal_of_real RS injD) 1);
@@ -873,6 +968,12 @@
qed "hypreal_of_real_approx_number_of_iff";
Addsimps [hypreal_of_real_approx_number_of_iff];
+(*And also for 0 and 1.*)
+Addsimps
+ (map (simplify (simpset()))
+ [inst "w" "Pls" hypreal_of_real_approx_number_of_iff,
+ inst "w" "Pls BIT True" hypreal_of_real_approx_number_of_iff]);
+
Goal "[| r: Reals; s: Reals; r @= x; s @= x|] ==> r = s";
by (blast_tac (claset() addIs [(SReal_approx_iff RS iffD1),
approx_trans2]) 1);
@@ -912,7 +1013,7 @@
lemma_st_part_nonempty, lemma_st_part_subset]) 1);
qed "lemma_st_part_lub";
-Goal "((t::hypreal) + r <= t) = (r <= Numeral0)";
+Goal "((t::hypreal) + r <= t) = (r <= 0)";
by (Step_tac 1);
by (dres_inst_tac [("x","-t")] hypreal_add_left_le_mono1 1);
by (dres_inst_tac [("x","t")] hypreal_add_left_le_mono1 2);
@@ -920,7 +1021,7 @@
qed "lemma_hypreal_le_left_cancel";
Goal "[| x: HFinite; isLub Reals {s. s: Reals & s < x} t; \
-\ r: Reals; Numeral0 < r |] ==> x <= t + r";
+\ r: Reals; 0 < r |] ==> x <= t + r";
by (forward_tac [isLubD1a] 1);
by (rtac ccontr 1 THEN dtac (linorder_not_le RS iffD2) 1);
by (dres_inst_tac [("x","t")] SReal_add 1 THEN assume_tac 1);
@@ -945,14 +1046,14 @@
addIs [order_less_imp_le] addSIs [isUbI,setleI], simpset()));
qed "lemma_st_part_gt_ub";
-Goal "t <= t + -r ==> r <= (Numeral0::hypreal)";
+Goal "t <= t + -r ==> r <= (0::hypreal)";
by (dres_inst_tac [("x","-t")] hypreal_add_left_le_mono1 1);
by (auto_tac (claset(), simpset() addsimps [hypreal_add_assoc RS sym]));
qed "lemma_minus_le_zero";
Goal "[| x: HFinite; \
\ isLub Reals {s. s: Reals & s < x} t; \
-\ r: Reals; Numeral0 < r |] \
+\ r: Reals; 0 < r |] \
\ ==> t + -r <= x";
by (forward_tac [isLubD1a] 1);
by (rtac ccontr 1 THEN dtac not_hypreal_leE 1);
@@ -970,7 +1071,7 @@
Goal "[| x: HFinite; \
\ isLub Reals {s. s: Reals & s < x} t; \
-\ r: Reals; Numeral0 < r |] \
+\ r: Reals; 0 < r |] \
\ ==> x + -t <= r";
by (blast_tac (claset() addSIs [lemma_hypreal_le_swap RS iffD1,
lemma_st_part_le1]) 1);
@@ -982,7 +1083,7 @@
Goal "[| x: HFinite; \
\ isLub Reals {s. s: Reals & s < x} t; \
-\ r: Reals; Numeral0 < r |] \
+\ r: Reals; 0 < r |] \
\ ==> -(x + -t) <= r";
by (blast_tac (claset() addSIs [lemma_hypreal_le_swap2 RS iffD1,
lemma_st_part_le2]) 1);
@@ -1004,7 +1105,7 @@
Goal "[| x: HFinite; \
\ isLub Reals {s. s: Reals & s < x} t; \
-\ r: Reals; Numeral0 < r |] \
+\ r: Reals; 0 < r |] \
\ ==> x + -t ~= r";
by Auto_tac;
by (forward_tac [isLubD1a RS SReal_minus] 1);
@@ -1016,7 +1117,7 @@
Goal "[| x: HFinite; \
\ isLub Reals {s. s: Reals & s < x} t; \
-\ r: Reals; Numeral0 < r |] \
+\ r: Reals; 0 < r |] \
\ ==> -(x + -t) ~= r";
by (auto_tac (claset(), simpset() addsimps [hypreal_minus_add_distrib]));
by (forward_tac [isLubD1a] 1);
@@ -1030,7 +1131,7 @@
Goal "[| x: HFinite; \
\ isLub Reals {s. s: Reals & s < x} t; \
-\ r: Reals; Numeral0 < r |] \
+\ r: Reals; 0 < r |] \
\ ==> abs (x + -t) < r";
by (forward_tac [lemma_st_part1a] 1);
by (forward_tac [lemma_st_part2a] 4);
@@ -1042,7 +1143,7 @@
Goal "[| x: HFinite; \
\ isLub Reals {s. s: Reals & s < x} t |] \
-\ ==> ALL r: Reals. Numeral0 < r --> abs (x + -t) < r";
+\ ==> ALL r: Reals. 0 < r --> abs (x + -t) < r";
by (blast_tac (claset() addSDs [lemma_st_part_major]) 1);
qed "lemma_st_part_major2";
@@ -1050,7 +1151,7 @@
Existence of real and Standard Part Theorem
----------------------------------------------*)
Goal "x: HFinite ==> \
-\ EX t: Reals. ALL r: Reals. Numeral0 < r --> abs (x + -t) < r";
+\ EX t: Reals. ALL r: Reals. 0 < r --> abs (x + -t) < r";
by (forward_tac [lemma_st_part_lub] 1 THEN Step_tac 1);
by (forward_tac [isLubD1a] 1);
by (blast_tac (claset() addDs [lemma_st_part_major2]) 1);
@@ -1089,7 +1190,7 @@
Goalw [HInfinite_def, HFinite_def] "x~: HFinite ==> x: HInfinite";
by Auto_tac;
-by (dres_inst_tac [("x","r + Numeral1")] bspec 1);
+by (dres_inst_tac [("x","r + 1")] bspec 1);
by (auto_tac (claset(), simpset() addsimps [SReal_add]));
qed "not_HFinite_HInfinite";
@@ -1241,16 +1342,16 @@
by Auto_tac;
qed "mem_monad_iff";
-Goalw [monad_def] "(x:Infinitesimal) = (x:monad Numeral0)";
+Goalw [monad_def] "(x:Infinitesimal) = (x:monad 0)";
by (auto_tac (claset() addIs [approx_sym],
simpset() addsimps [mem_infmal_iff]));
qed "Infinitesimal_monad_zero_iff";
-Goal "(x:monad Numeral0) = (-x:monad Numeral0)";
+Goal "(x:monad 0) = (-x:monad 0)";
by (simp_tac (simpset() addsimps [Infinitesimal_monad_zero_iff RS sym]) 1);
qed "monad_zero_minus_iff";
-Goal "(x:monad Numeral0) = (abs x:monad Numeral0)";
+Goal "(x:monad 0) = (abs x:monad 0)";
by (res_inst_tac [("x1","x")] (hrabs_disj RS disjE) 1);
by (auto_tac (claset(), simpset() addsimps [monad_zero_minus_iff RS sym]));
qed "monad_zero_hrabs_iff";
@@ -1286,7 +1387,7 @@
by (blast_tac (claset() addSIs [approx_sym]) 1);
qed "approx_mem_monad2";
-Goal "[| x @= y;x:monad Numeral0 |] ==> y:monad Numeral0";
+Goal "[| x @= y;x:monad 0 |] ==> y:monad 0";
by (dtac mem_monad_approx 1);
by (fast_tac (claset() addIs [approx_mem_monad,approx_trans]) 1);
qed "approx_mem_monad_zero";
@@ -1297,7 +1398,7 @@
monad_zero_hrabs_iff RS iffD1, mem_monad_approx, approx_trans3]) 1);
qed "Infinitesimal_approx_hrabs";
-Goal "[| Numeral0 < x; x ~:Infinitesimal; e :Infinitesimal |] ==> e < x";
+Goal "[| 0 < x; x ~:Infinitesimal; e :Infinitesimal |] ==> e < x";
by (rtac ccontr 1);
by (auto_tac (claset()
addIs [Infinitesimal_zero RSN (2, Infinitesimal_interval)]
@@ -1305,38 +1406,38 @@
simpset()));
qed "less_Infinitesimal_less";
-Goal "[| Numeral0 < x; x ~: Infinitesimal; u: monad x |] ==> Numeral0 < u";
+Goal "[| 0 < x; x ~: Infinitesimal; u: monad x |] ==> 0 < u";
by (dtac (mem_monad_approx RS approx_sym) 1);
by (etac (bex_Infinitesimal_iff2 RS iffD2 RS bexE) 1);
by (dres_inst_tac [("e","-xa")] less_Infinitesimal_less 1);
by Auto_tac;
qed "Ball_mem_monad_gt_zero";
-Goal "[| x < Numeral0; x ~: Infinitesimal; u: monad x |] ==> u < Numeral0";
+Goal "[| x < 0; x ~: Infinitesimal; u: monad x |] ==> u < 0";
by (dtac (mem_monad_approx RS approx_sym) 1);
by (etac (bex_Infinitesimal_iff RS iffD2 RS bexE) 1);
by (cut_inst_tac [("x","-x"),("e","xa")] less_Infinitesimal_less 1);
by Auto_tac;
qed "Ball_mem_monad_less_zero";
-Goal "[|Numeral0 < x; x ~: Infinitesimal; x @= y|] ==> Numeral0 < y";
+Goal "[|0 < x; x ~: Infinitesimal; x @= y|] ==> 0 < y";
by (blast_tac (claset() addDs [Ball_mem_monad_gt_zero,
approx_subset_monad]) 1);
qed "lemma_approx_gt_zero";
-Goal "[|x < Numeral0; x ~: Infinitesimal; x @= y|] ==> y < Numeral0";
+Goal "[|x < 0; x ~: Infinitesimal; x @= y|] ==> y < 0";
by (blast_tac (claset() addDs [Ball_mem_monad_less_zero,
approx_subset_monad]) 1);
qed "lemma_approx_less_zero";
-Goal "[| x @= y; x < Numeral0; x ~: Infinitesimal |] ==> abs x @= abs y";
+Goal "[| x @= y; x < 0; x ~: Infinitesimal |] ==> abs x @= abs y";
by (forward_tac [lemma_approx_less_zero] 1);
by (REPEAT(assume_tac 1));
by (REPEAT(dtac hrabs_minus_eqI2 1));
by Auto_tac;
qed "approx_hrabs1";
-Goal "[| x @= y; Numeral0 < x; x ~: Infinitesimal |] ==> abs x @= abs y";
+Goal "[| x @= y; 0 < x; x ~: Infinitesimal |] ==> abs x @= abs y";
by (forward_tac [lemma_approx_gt_zero] 1);
by (REPEAT(assume_tac 1));
by (REPEAT(dtac hrabs_eqI2 1));
@@ -1345,12 +1446,12 @@
Goal "x @= y ==> abs x @= abs y";
by (res_inst_tac [("Q","x:Infinitesimal")] (excluded_middle RS disjE) 1);
-by (res_inst_tac [("x1","x"),("y1","Numeral0")] (hypreal_linear RS disjE) 1);
+by (res_inst_tac [("x1","x"),("y1","0")] (hypreal_linear RS disjE) 1);
by (auto_tac (claset() addIs [approx_hrabs1,approx_hrabs2,
Infinitesimal_approx_hrabs], simpset()));
qed "approx_hrabs";
-Goal "abs(x) @= Numeral0 ==> x @= Numeral0";
+Goal "abs(x) @= 0 ==> x @= 0";
by (cut_inst_tac [("x","x")] hrabs_disj 1);
by (auto_tac (claset() addDs [approx_minus], simpset()));
qed "approx_hrabs_zero_cancel";
@@ -1445,15 +1546,15 @@
hypreal_of_real_le_add_Infininitesimal_cancel]) 1);
qed "hypreal_of_real_le_add_Infininitesimal_cancel2";
-Goal "[| hypreal_of_real x < e; e: Infinitesimal |] ==> hypreal_of_real x <= Numeral0";
+Goal "[| hypreal_of_real x < e; e: Infinitesimal |] ==> hypreal_of_real x <= 0";
by (rtac hypreal_leI 1 THEN Step_tac 1);
by (dtac Infinitesimal_interval 1);
by (dtac (SReal_hypreal_of_real RS SReal_Infinitesimal_zero) 4);
-by (auto_tac (claset(), simpset() addsimps [hypreal_of_real_zero]));
+by Auto_tac;
qed "hypreal_of_real_less_Infinitesimal_le_zero";
(*used once, in NSDERIV_inverse*)
-Goal "[| h: Infinitesimal; x ~= Numeral0 |] ==> hypreal_of_real x + h ~= Numeral0";
+Goal "[| h: Infinitesimal; x ~= 0 |] ==> hypreal_of_real x + h ~= 0";
by Auto_tac;
qed "Infinitesimal_add_not_zero";
@@ -1468,7 +1569,7 @@
Goal "x*x + y*y : HFinite ==> x*x : HFinite";
by (rtac HFinite_bounded 1);
by (rtac hypreal_self_le_add_pos 2);
-by (rtac (rename_numerals hypreal_le_square) 2);
+by (rtac (hypreal_le_square) 2);
by (assume_tac 1);
qed "HFinite_square_cancel";
Addsimps [HFinite_square_cancel];
@@ -1489,13 +1590,13 @@
Goal "x*x + y*y + z*z : Infinitesimal ==> x*x : Infinitesimal";
by (blast_tac (claset() addIs [hypreal_self_le_add_pos2,
- Infinitesimal_interval2, rename_numerals hypreal_le_square]) 1);
+ Infinitesimal_interval2, hypreal_le_square]) 1);
qed "Infinitesimal_sum_square_cancel";
Addsimps [Infinitesimal_sum_square_cancel];
Goal "x*x + y*y + z*z : HFinite ==> x*x : HFinite";
by (blast_tac (claset() addIs [hypreal_self_le_add_pos2, HFinite_bounded,
- rename_numerals hypreal_le_square,
+ hypreal_le_square,
HFinite_number_of]) 1);
qed "HFinite_sum_square_cancel";
Addsimps [HFinite_sum_square_cancel];
@@ -1524,7 +1625,7 @@
qed "HFinite_sum_square_cancel3";
Addsimps [HFinite_sum_square_cancel3];
-Goal "[| y: monad x; Numeral0 < hypreal_of_real e |] \
+Goal "[| y: monad x; 0 < hypreal_of_real e |] \
\ ==> abs (y + -x) < hypreal_of_real e";
by (dtac (mem_monad_approx RS approx_sym) 1);
by (dtac (bex_Infinitesimal_iff RS iffD2) 1);
@@ -1614,8 +1715,7 @@
approx_sym),bex_Infinitesimal_iff2 RS iffD2]) 1);
qed "HFinite_st_Infinitesimal_add";
-Goal "[| x: HFinite; y: HFinite |] \
-\ ==> st (x + y) = st(x) + st(y)";
+Goal "[| x: HFinite; y: HFinite |] ==> st (x + y) = st(x) + st(y)";
by (forward_tac [HFinite_st_Infinitesimal_add] 1);
by (forw_inst_tac [("x","y")] HFinite_st_Infinitesimal_add 1);
by (Step_tac 1);
@@ -1635,6 +1735,11 @@
qed "st_number_of";
Addsimps [st_number_of];
+(*the theorem above for the special cases of zero and one*)
+Addsimps
+ (map (simplify (simpset()))
+ [inst "w" "Pls" st_number_of, inst "w" "Pls BIT True" st_number_of]);
+
Goal "y: HFinite ==> st(-y) = -st(y)";
by (forward_tac [HFinite_minus_iff RS iffD2] 1);
by (rtac hypreal_add_minus_eq_minus 1);
@@ -1682,18 +1787,19 @@
by (blast_tac (claset() addSIs [lemma_st_mult]) 1);
qed "st_mult";
-Goal "x: Infinitesimal ==> st x = Numeral0";
+Goal "x: Infinitesimal ==> st x = 0";
+by (stac (hypreal_numeral_0_eq_0 RS sym) 1);
by (rtac (st_number_of RS subst) 1);
by (rtac approx_st_eq 1);
by (auto_tac (claset() addIs [Infinitesimal_subset_HFinite RS subsetD],
simpset() addsimps [mem_infmal_iff RS sym]));
qed "st_Infinitesimal";
-Goal "st(x) ~= Numeral0 ==> x ~: Infinitesimal";
+Goal "st(x) ~= 0 ==> x ~: Infinitesimal";
by (fast_tac (claset() addIs [st_Infinitesimal]) 1);
qed "st_not_Infinitesimal";
-Goal "[| x: HFinite; st x ~= Numeral0 |] \
+Goal "[| x: HFinite; st x ~= 0 |] \
\ ==> st(inverse x) = inverse (st x)";
by (res_inst_tac [("c1","st x")] (hypreal_mult_left_cancel RS iffD1) 1);
by (auto_tac (claset(),
@@ -1703,7 +1809,7 @@
by Auto_tac;
qed "st_inverse";
-Goal "[| x: HFinite; y: HFinite; st y ~= Numeral0 |] \
+Goal "[| x: HFinite; y: HFinite; st y ~= 0 |] \
\ ==> st(x/y) = (st x) / (st y)";
by (auto_tac (claset(),
simpset() addsimps [hypreal_divide_def, st_mult, st_not_Infinitesimal,
@@ -1734,8 +1840,7 @@
simpset()));
qed "Infinitesimal_add_st_le_cancel";
-Goal "[| x: HFinite; y: HFinite; x <= y |] \
-\ ==> st(x) <= st(y)";
+Goal "[| x: HFinite; y: HFinite; x <= y |] ==> st(x) <= st(y)";
by (forward_tac [HFinite_st_Infinitesimal_add] 1);
by (rotate_tac 1 1);
by (forward_tac [HFinite_st_Infinitesimal_add] 1);
@@ -1743,24 +1848,25 @@
by (rtac Infinitesimal_add_st_le_cancel 1);
by (res_inst_tac [("x","ea"),("y","e")]
Infinitesimal_diff 3);
-by (auto_tac (claset(),
- simpset() addsimps [hypreal_add_assoc RS sym]));
+by (auto_tac (claset(), simpset() addsimps [hypreal_add_assoc RS sym]));
qed "st_le";
-Goal "[| Numeral0 <= x; x: HFinite |] ==> Numeral0 <= st x";
+Goal "[| 0 <= x; x: HFinite |] ==> 0 <= st x";
+by (stac (hypreal_numeral_0_eq_0 RS sym) 1);
by (rtac (st_number_of RS subst) 1);
-by (auto_tac (claset() addIs [st_le],
- simpset() delsimps [st_number_of]));
+by (rtac st_le 1);
+by Auto_tac;
qed "st_zero_le";
-Goal "[| x <= Numeral0; x: HFinite |] ==> st x <= Numeral0";
+Goal "[| x <= 0; x: HFinite |] ==> st x <= 0";
+by (stac (hypreal_numeral_0_eq_0 RS sym) 1);
by (rtac (st_number_of RS subst) 1);
-by (auto_tac (claset() addIs [st_le],
- simpset() delsimps [st_number_of]));
+by (rtac st_le 1);
+by Auto_tac;
qed "st_zero_ge";
Goal "x: HFinite ==> abs(st x) = st(abs x)";
-by (case_tac "Numeral0 <= x" 1);
+by (case_tac "0 <= x" 1);
by (auto_tac (claset() addSDs [not_hypreal_leE, order_less_imp_le],
simpset() addsimps [st_zero_le,hrabs_eqI1, hrabs_minus_eqI1,
st_zero_ge,st_minus]));
@@ -1834,7 +1940,7 @@
by Auto_tac;
qed "lemma_Int_eq1";
-Goal "{n. abs (xa n) = u} <= {n. abs (xa n) < u + (Numeral1::real)}";
+Goal "{n. abs (xa n) = u} <= {n. abs (xa n) < u + (1::real)}";
by Auto_tac;
qed "lemma_FreeUltrafilterNat_one";
@@ -1847,7 +1953,7 @@
\ |] ==> x: HFinite";
by (rtac FreeUltrafilterNat_HFinite 1);
by (res_inst_tac [("x","xa")] bexI 1);
-by (res_inst_tac [("x","u + Numeral1")] exI 1);
+by (res_inst_tac [("x","u + 1")] exI 1);
by (Ultra_tac 1 THEN assume_tac 1);
qed "FreeUltrafilterNat_const_Finite";
@@ -1915,22 +2021,22 @@
Goalw [Infinitesimal_def]
"x : Infinitesimal ==> EX X: Rep_hypreal x. \
-\ ALL u. Numeral0 < u --> {n. abs (X n) < u}: FreeUltrafilterNat";
+\ ALL u. 0 < u --> {n. abs (X n) < u}: FreeUltrafilterNat";
by (auto_tac (claset(), simpset() addsimps [hrabs_interval_iff]));
by (res_inst_tac [("z","x")] eq_Abs_hypreal 1);
by (EVERY[Auto_tac, rtac bexI 1, rtac lemma_hyprel_refl 2, Step_tac 1]);
by (dtac (hypreal_of_real_less_iff RS iffD2) 1);
by (dres_inst_tac [("x","hypreal_of_real u")] bspec 1);
-by (auto_tac (claset(), simpset() addsimps [hypreal_of_real_zero]));
+by Auto_tac;
by (auto_tac (claset(),
simpset() addsimps [hypreal_less_def, hypreal_minus,
- hypreal_of_real_def,hypreal_of_real_zero]));
+ hypreal_of_real_def]));
by (Ultra_tac 1 THEN arith_tac 1);
qed "Infinitesimal_FreeUltrafilterNat";
Goalw [Infinitesimal_def]
"EX X: Rep_hypreal x. \
-\ ALL u. Numeral0 < u --> {n. abs (X n) < u} : FreeUltrafilterNat \
+\ ALL u. 0 < u --> {n. abs (X n) < u} : FreeUltrafilterNat \
\ ==> x : Infinitesimal";
by (auto_tac (claset(),
simpset() addsimps [hrabs_interval_iff,abs_interval_iff]));
@@ -1942,7 +2048,7 @@
qed "FreeUltrafilterNat_Infinitesimal";
Goal "(x : Infinitesimal) = (EX X: Rep_hypreal x. \
-\ ALL u. Numeral0 < u --> {n. abs (X n) < u}: FreeUltrafilterNat)";
+\ ALL u. 0 < u --> {n. abs (X n) < u}: FreeUltrafilterNat)";
by (blast_tac (claset() addSIs [Infinitesimal_FreeUltrafilterNat,
FreeUltrafilterNat_Infinitesimal]) 1);
qed "Infinitesimal_FreeUltrafilterNat_iff";
@@ -1951,25 +2057,25 @@
Infinitesimals as smaller than 1/n for all n::nat (> 0)
------------------------------------------------------------------------*)
-Goal "(ALL r. Numeral0 < r --> x < r) = (ALL n. x < inverse(real (Suc n)))";
+Goal "(ALL r. 0 < r --> x < r) = (ALL n. x < inverse(real (Suc n)))";
by (auto_tac (claset(), simpset() addsimps [real_of_nat_Suc_gt_zero]));
by (blast_tac (claset() addSDs [reals_Archimedean]
addIs [order_less_trans]) 1);
qed "lemma_Infinitesimal";
-Goal "(ALL r: Reals. Numeral0 < r --> x < r) = \
+Goal "(ALL r: Reals. 0 < r --> x < r) = \
\ (ALL n. x < inverse(hypreal_of_nat (Suc n)))";
by (Step_tac 1);
by (dres_inst_tac [("x","inverse (hypreal_of_real(real (Suc n)))")]
bspec 1);
by (full_simp_tac (simpset() addsimps [SReal_inverse]) 1);
-by (rtac (real_of_nat_Suc_gt_zero RS rename_numerals real_inverse_gt_zero RS
+by (rtac (real_of_nat_Suc_gt_zero RS real_inverse_gt_0 RS
(hypreal_of_real_less_iff RS iffD2) RSN(2,impE)) 1);
by (assume_tac 2);
by (asm_full_simp_tac (simpset() addsimps
- [real_of_nat_Suc_gt_zero, hypreal_of_real_zero, hypreal_of_nat_def]) 1);
+ [real_of_nat_Suc_gt_zero, hypreal_of_nat_def]) 1);
by (auto_tac (claset() addSDs [reals_Archimedean],
- simpset() addsimps [SReal_iff,hypreal_of_real_zero RS sym]));
+ simpset() addsimps [SReal_iff]));
by (dtac (hypreal_of_real_less_iff RS iffD2) 1);
by (asm_full_simp_tac (simpset() addsimps
[real_of_nat_Suc_gt_zero, hypreal_of_nat_def])1);
@@ -2089,7 +2195,7 @@
qed "HFinite_epsilon";
Addsimps [HFinite_epsilon];
-Goal "epsilon @= Numeral0";
+Goal "epsilon @= 0";
by (simp_tac (simpset() addsimps [mem_infmal_iff RS sym]) 1);
qed "epsilon_approx_zero";
Addsimps [epsilon_approx_zero];
@@ -2109,7 +2215,7 @@
by (simp_tac (simpset() addsimps [real_of_nat_Suc_gt_zero]) 1);
qed "real_of_nat_less_inverse_iff";
-Goal "Numeral0 < u ==> finite {n. u < inverse(real(Suc n))}";
+Goal "0 < u ==> finite {n. u < inverse(real(Suc n))}";
by (asm_simp_tac (simpset() addsimps [real_of_nat_less_inverse_iff]) 1);
by (asm_simp_tac (simpset() addsimps [real_of_nat_Suc,
real_less_diff_eq RS sym]) 1);
@@ -2122,7 +2228,7 @@
simpset() addsimps [order_less_imp_le]));
qed "lemma_real_le_Un_eq2";
-Goal "(inverse (real(Suc n)) <= r) = (Numeral1 <= r * real(Suc n))";
+Goal "(inverse (real(Suc n)) <= r) = (1 <= r * real(Suc n))";
by (simp_tac (simpset() addsimps [linorder_not_less RS sym]) 1);
by (simp_tac (simpset() addsimps [real_inverse_eq_divide]) 1);
by (stac pos_real_less_divide_eq 1);
@@ -2138,18 +2244,18 @@
Goal "finite {n::nat. u = inverse(real(Suc n))}";
by (asm_simp_tac (simpset() addsimps [real_of_nat_inverse_eq_iff]) 1);
-by (cut_inst_tac [("x","inverse u - Numeral1")] lemma_finite_omega_set 1);
+by (cut_inst_tac [("x","inverse u - 1")] lemma_finite_omega_set 1);
by (asm_full_simp_tac (simpset() addsimps [real_of_nat_Suc,
real_diff_eq_eq RS sym, eq_commute]) 1);
qed "lemma_finite_omega_set2";
-Goal "Numeral0 < u ==> finite {n. u <= inverse(real(Suc n))}";
+Goal "0 < u ==> finite {n. u <= inverse(real(Suc n))}";
by (auto_tac (claset(),
simpset() addsimps [lemma_real_le_Un_eq2,lemma_finite_omega_set2,
finite_inverse_real_of_posnat_gt_real]));
qed "finite_inverse_real_of_posnat_ge_real";
-Goal "Numeral0 < u ==> \
+Goal "0 < u ==> \
\ {n. u <= inverse(real(Suc n))} ~: FreeUltrafilterNat";
by (blast_tac (claset() addSIs [FreeUltrafilterNat_finite,
finite_inverse_real_of_posnat_ge_real]) 1);
@@ -2166,7 +2272,7 @@
simpset() addsimps [not_real_leE]));
val lemma = result();
-Goal "Numeral0 < u ==> \
+Goal "0 < u ==> \
\ {n. inverse(real(Suc n)) < u} : FreeUltrafilterNat";
by (cut_inst_tac [("u","u")] inverse_real_of_posnat_ge_real_FreeUltrafilterNat 1);
by (auto_tac (claset() addDs [FreeUltrafilterNat_Compl_mem],
@@ -2187,7 +2293,7 @@
Goal "ALL n. abs(X n + -x) < inverse(real(Suc n)) \
\ ==> Abs_hypreal(hyprel``{X}) + -hypreal_of_real x : Infinitesimal";
by (auto_tac (claset() addSIs [bexI]
- addDs [rename_numerals FreeUltrafilterNat_inverse_real_of_posnat,
+ addDs [FreeUltrafilterNat_inverse_real_of_posnat,
FreeUltrafilterNat_all,FreeUltrafilterNat_Int]
addIs [order_less_trans, FreeUltrafilterNat_subset],
simpset() addsimps [hypreal_minus,
@@ -2212,8 +2318,7 @@
\ ==> Abs_hypreal(hyprel``{X}) + \
\ -Abs_hypreal(hyprel``{Y}) : Infinitesimal";
by (auto_tac (claset() addSIs [bexI]
- addDs [rename_numerals
- FreeUltrafilterNat_inverse_real_of_posnat,
+ addDs [FreeUltrafilterNat_inverse_real_of_posnat,
FreeUltrafilterNat_all,FreeUltrafilterNat_Int]
addIs [order_less_trans, FreeUltrafilterNat_subset],
simpset() addsimps
--- a/src/HOL/Hyperreal/NatStar.ML Thu Nov 01 21:12:13 2001 +0100
+++ b/src/HOL/Hyperreal/NatStar.ML Fri Nov 02 17:55:24 2001 +0100
@@ -404,7 +404,7 @@
Goal "N : HNatInfinite \
\ ==> (*fNat* (%x::nat. inverse(real x))) N = inverse(hypreal_of_hypnat N)";
by (res_inst_tac [("f1","inverse")] (starfun_stafunNat_o2 RS subst) 1);
-by (subgoal_tac "hypreal_of_hypnat N ~= Numeral0" 1);
+by (subgoal_tac "hypreal_of_hypnat N ~= 0" 1);
by (auto_tac (claset(),
simpset() addsimps [starfunNat_real_of_nat, starfun_inverse_inverse]));
qed "starfunNat_inverse_real_of_nat_eq";
--- a/src/HOL/Hyperreal/SEQ.ML Thu Nov 01 21:12:13 2001 +0100
+++ b/src/HOL/Hyperreal/SEQ.ML Fri Nov 02 17:55:24 2001 +0100
@@ -26,7 +26,7 @@
Goalw [LIMSEQ_def]
"(X ----> L) = \
-\ (ALL r. Numeral0 <r --> (EX no. ALL n. no <= n --> abs(X n + -L) < r))";
+\ (ALL r. 0 <r --> (EX no. ALL n. no <= n --> abs(X n + -L) < r))";
by (Simp_tac 1);
qed "LIMSEQ_iff";
@@ -120,7 +120,7 @@
by Auto_tac;
val lemmaLIM2 = result();
-Goal "[| Numeral0 < r; ALL n. r <= abs (X (f n) + - L); \
+Goal "[| 0 < r; ALL n. r <= abs (X (f n) + - L); \
\ (*fNat* X) (Abs_hypnat (hypnatrel `` {f})) + \
\ - hypreal_of_real L @= 0 |] ==> False";
by (auto_tac (claset(),simpset() addsimps [starfunNat,
@@ -147,7 +147,7 @@
by (dtac (approx_minus_iff RS iffD1) 2);
by (fold_tac [real_le_def]);
by (blast_tac (claset() addIs [HNatInfinite_NSLIMSEQ]) 1);
-by (blast_tac (claset() addIs [rename_numerals lemmaLIM3]) 1);
+by (blast_tac (claset() addIs [lemmaLIM3]) 1);
qed "NSLIMSEQ_LIMSEQ";
(* Now the all important result is trivially proved! *)
@@ -234,7 +234,7 @@
Proof is like that of NSLIM_inverse.
--------------------------------------------------------------*)
Goalw [NSLIMSEQ_def]
- "[| X ----NS> a; a ~= Numeral0 |] ==> (%n. inverse(X n)) ----NS> inverse(a)";
+ "[| X ----NS> a; a ~= 0 |] ==> (%n. inverse(X n)) ----NS> inverse(a)";
by (Clarify_tac 1);
by (dtac bspec 1);
by (auto_tac (claset(),
@@ -244,18 +244,18 @@
(*------ Standard version of theorem -------*)
-Goal "[| X ----> a; a ~= Numeral0 |] ==> (%n. inverse(X n)) ----> inverse(a)";
+Goal "[| X ----> a; a ~= 0 |] ==> (%n. inverse(X n)) ----> inverse(a)";
by (asm_full_simp_tac (simpset() addsimps [NSLIMSEQ_inverse,
LIMSEQ_NSLIMSEQ_iff]) 1);
qed "LIMSEQ_inverse";
-Goal "[| X ----NS> a; Y ----NS> b; b ~= Numeral0 |] \
+Goal "[| X ----NS> a; Y ----NS> b; b ~= 0 |] \
\ ==> (%n. X n / Y n) ----NS> a/b";
by (asm_full_simp_tac (simpset() addsimps [NSLIMSEQ_mult, NSLIMSEQ_inverse,
real_divide_def]) 1);
qed "NSLIMSEQ_mult_inverse";
-Goal "[| X ----> a; Y ----> b; b ~= Numeral0 |] ==> (%n. X n / Y n) ----> a/b";
+Goal "[| X ----> a; Y ----> b; b ~= 0 |] ==> (%n. X n / Y n) ----> a/b";
by (asm_full_simp_tac (simpset() addsimps [LIMSEQ_mult, LIMSEQ_inverse,
real_divide_def]) 1);
qed "LIMSEQ_divide";
@@ -376,16 +376,16 @@
Bounded Sequence
------------------------------------------------------------------*)
Goalw [Bseq_def]
- "Bseq X ==> EX K. Numeral0 < K & (ALL n. abs(X n) <= K)";
+ "Bseq X ==> EX K. 0 < K & (ALL n. abs(X n) <= K)";
by (assume_tac 1);
qed "BseqD";
Goalw [Bseq_def]
- "[| Numeral0 < K; ALL n. abs(X n) <= K |] ==> Bseq X";
+ "[| 0 < K; ALL n. abs(X n) <= K |] ==> Bseq X";
by (Blast_tac 1);
qed "BseqI";
-Goal "(EX K. Numeral0 < K & (ALL n. abs(X n) <= K)) = \
+Goal "(EX K. 0 < K & (ALL n. abs(X n) <= K)) = \
\ (EX N. ALL n. abs(X n) <= real(Suc N))";
by Auto_tac;
by (cut_inst_tac [("x","K")] reals_Archimedean2 1);
@@ -401,7 +401,7 @@
by (simp_tac (simpset() addsimps [lemma_NBseq_def]) 1);
qed "Bseq_iff";
-Goal "(EX K. Numeral0 < K & (ALL n. abs(X n) <= K)) = \
+Goal "(EX K. 0 < K & (ALL n. abs(X n) <= K)) = \
\ (EX N. ALL n. abs(X n) < real(Suc N))";
by (stac lemma_NBseq_def 1);
by Auto_tac;
@@ -444,7 +444,7 @@
HNatInfinite_FreeUltrafilterNat_iff]));
by (EVERY[rtac bexI 1, rtac lemma_hyprel_refl 2]);
by (dres_inst_tac [("f","Xa")] lemma_Bseq 1);
-by (res_inst_tac [("x","K+Numeral1")] exI 1);
+by (res_inst_tac [("x","K+1")] exI 1);
by (rotate_tac 2 1 THEN dtac FreeUltrafilterNat_all 1);
by (Ultra_tac 1);
qed "Bseq_NSBseq";
@@ -461,14 +461,14 @@
is not what we want (read useless!)
-------------------------------------------------------------------*)
-Goal "ALL K. Numeral0 < K --> (EX n. K < abs (X n)) \
+Goal "ALL K. 0 < K --> (EX n. K < abs (X n)) \
\ ==> ALL N. EX n. real(Suc N) < abs (X n)";
by (Step_tac 1);
by (cut_inst_tac [("n","N")] real_of_nat_Suc_gt_zero 1);
by (Blast_tac 1);
val lemmaNSBseq = result();
-Goal "ALL K. Numeral0 < K --> (EX n. K < abs (X n)) \
+Goal "ALL K. 0 < K --> (EX n. K < abs (X n)) \
\ ==> EX f. ALL N. real(Suc N) < abs (X (f N))";
by (dtac lemmaNSBseq 1);
by (dtac choice 1);
@@ -652,7 +652,7 @@
Goal "!!(X::nat=> real). \
\ [| ALL m. X m ~= U; \
\ isLub UNIV {x. EX n. X n = x} U; \
-\ Numeral0 < T; \
+\ 0 < T; \
\ U + - T < U \
\ |] ==> EX m. U + -T < X m & X m < U";
by (dtac lemma_converg2 1 THEN assume_tac 1);
@@ -722,7 +722,7 @@
(***--- alternative formulation for boundedness---***)
Goalw [Bseq_def]
- "Bseq X = (EX k x. Numeral0 < k & (ALL n. abs(X(n) + -x) <= k))";
+ "Bseq X = (EX k x. 0 < k & (ALL n. abs(X(n) + -x) <= k))";
by (Step_tac 1);
by (res_inst_tac [("x","k + abs(x)")] exI 2);
by (res_inst_tac [("x","K")] exI 1);
@@ -733,7 +733,7 @@
qed "Bseq_iff2";
(***--- alternative formulation for boundedness ---***)
-Goal "Bseq X = (EX k N. Numeral0 < k & (ALL n. abs(X(n) + -X(N)) <= k))";
+Goal "Bseq X = (EX k N. 0 < k & (ALL n. abs(X(n) + -X(N)) <= k))";
by (Step_tac 1);
by (asm_full_simp_tac (simpset() addsimps [Bseq_def]) 1);
by (Step_tac 1);
@@ -748,7 +748,7 @@
qed "Bseq_iff3";
Goalw [Bseq_def] "(ALL n. k <= f n & f n <= K) ==> Bseq f";
-by (res_inst_tac [("x","(abs(k) + abs(K)) + Numeral1")] exI 1);
+by (res_inst_tac [("x","(abs(k) + abs(K)) + 1")] exI 1);
by (Auto_tac);
by (dres_inst_tac [("x","n")] spec 2);
by (ALLGOALS arith_tac);
@@ -841,8 +841,8 @@
-------------------------------------------------------*)
(***------------- VARIOUS LEMMAS --------------***)
-Goal "ALL n. M <= n --> abs (X M + - X n) < (Numeral1::real) \
-\ ==> ALL n. M <= n --> abs(X n) < Numeral1 + abs(X M)";
+Goal "ALL n. M <= n --> abs (X M + - X n) < (1::real) \
+\ ==> ALL n. M <= n --> abs(X n) < 1 + abs(X M)";
by (Step_tac 1);
by (dtac spec 1 THEN Auto_tac);
by (arith_tac 1);
@@ -911,8 +911,8 @@
outlines sketched by various authors would suggest
---------------------------------------------------------*)
Goalw [Cauchy_def,Bseq_def] "Cauchy X ==> Bseq X";
-by (dres_inst_tac [("x","Numeral1")] spec 1);
-by (etac (rename_numerals real_zero_less_one RSN (2,impE)) 1);
+by (dres_inst_tac [("x","1")] spec 1);
+by (etac (real_zero_less_one RSN (2,impE)) 1);
by (Step_tac 1);
by (dres_inst_tac [("x","M")] spec 1);
by (Asm_full_simp_tac 1);
@@ -920,7 +920,7 @@
by (cut_inst_tac [("M","M"),("X","X")] SUP_rabs_subseq 1);
by (Step_tac 1);
by (cut_inst_tac [("R1.0","abs(X m)"),
- ("R2.0","Numeral1 + abs(X M)")] real_linear 1);
+ ("R2.0","1 + abs(X M)")] real_linear 1);
by (Step_tac 1);
by (dtac lemma_trans1 1 THEN assume_tac 1);
by (dtac lemma_trans2 3 THEN assume_tac 3);
@@ -928,8 +928,8 @@
by (dtac (abs_add_one_gt_zero RS order_less_trans) 3);
by (dtac lemma_trans4 1);
by (dtac lemma_trans4 2);
-by (res_inst_tac [("x","Numeral1 + abs(X M)")] exI 1);
-by (res_inst_tac [("x","Numeral1 + abs(X M)")] exI 2);
+by (res_inst_tac [("x","1 + abs(X M)")] exI 1);
+by (res_inst_tac [("x","1 + abs(X M)")] exI 2);
by (res_inst_tac [("x","abs(X m)")] exI 3);
by (auto_tac (claset() addSEs [lemma_Nat_covered],
simpset()));
@@ -1082,7 +1082,7 @@
----------------------------------------------------*)
(* we can prove this directly since proof is trivial *)
Goalw [LIMSEQ_def]
- "((%n. abs(f n)) ----> Numeral0) = (f ----> Numeral0)";
+ "((%n. abs(f n)) ----> 0) = (f ----> 0)";
by (simp_tac (simpset() addsimps [abs_idempotent]) 1);
qed "LIMSEQ_rabs_zero";
@@ -1092,7 +1092,7 @@
(* than the direct standard one above! *)
(*-----------------------------------------------------*)
-Goal "((%n. abs(f n)) ----NS> Numeral0) = (f ----NS> Numeral0)";
+Goal "((%n. abs(f n)) ----NS> 0) = (f ----NS> 0)";
by (simp_tac (simpset() addsimps [LIMSEQ_NSLIMSEQ_iff RS sym,
LIMSEQ_rabs_zero]) 1);
qed "NSLIMSEQ_rabs_zero";
@@ -1119,7 +1119,7 @@
(* standard proof seems easier *)
Goalw [LIMSEQ_def]
"ALL y. EX N. ALL n. N <= n --> y < f(n) \
-\ ==> (%n. inverse(f n)) ----> Numeral0";
+\ ==> (%n. inverse(f n)) ----> 0";
by (Step_tac 1 THEN Asm_full_simp_tac 1);
by (dres_inst_tac [("x","inverse r")] spec 1 THEN Step_tac 1);
by (res_inst_tac [("x","N")] exI 1 THEN Step_tac 1);
@@ -1134,7 +1134,7 @@
qed "LIMSEQ_inverse_zero";
Goal "ALL y. EX N. ALL n. N <= n --> y < f(n) \
-\ ==> (%n. inverse(f n)) ----NS> Numeral0";
+\ ==> (%n. inverse(f n)) ----NS> 0";
by (asm_simp_tac (simpset() addsimps [LIMSEQ_NSLIMSEQ_iff RS sym,
LIMSEQ_inverse_zero]) 1);
qed "NSLIMSEQ_inverse_zero";
@@ -1143,7 +1143,7 @@
Sequence 1/n --> 0 as n --> infinity
-------------------------------------------------------------*)
-Goal "(%n. inverse(real(Suc n))) ----> Numeral0";
+Goal "(%n. inverse(real(Suc n))) ----> 0";
by (rtac LIMSEQ_inverse_zero 1 THEN Step_tac 1);
by (cut_inst_tac [("x","y")] reals_Archimedean2 1);
by (Step_tac 1 THEN res_inst_tac [("x","n")] exI 1);
@@ -1153,7 +1153,7 @@
by (blast_tac (claset() addIs [order_less_le_trans]) 1);
qed "LIMSEQ_inverse_real_of_nat";
-Goal "(%n. inverse(real(Suc n))) ----NS> Numeral0";
+Goal "(%n. inverse(real(Suc n))) ----NS> 0";
by (simp_tac (simpset() addsimps [LIMSEQ_NSLIMSEQ_iff RS sym,
LIMSEQ_inverse_real_of_nat]) 1);
qed "NSLIMSEQ_inverse_real_of_nat";
@@ -1188,13 +1188,13 @@
LIMSEQ_inverse_real_of_posnat_add_minus]) 1);
qed "NSLIMSEQ_inverse_real_of_posnat_add_minus";
-Goal "(%n. r*( Numeral1 + -inverse(real(Suc n)))) ----> r";
-by (cut_inst_tac [("b","Numeral1")] ([LIMSEQ_const,
+Goal "(%n. r*( 1 + -inverse(real(Suc n)))) ----> r";
+by (cut_inst_tac [("b","1")] ([LIMSEQ_const,
LIMSEQ_inverse_real_of_posnat_add_minus] MRS LIMSEQ_mult) 1);
by (Auto_tac);
qed "LIMSEQ_inverse_real_of_posnat_add_minus_mult";
-Goal "(%n. r*( Numeral1 + -inverse(real(Suc n)))) ----NS> r";
+Goal "(%n. r*( 1 + -inverse(real(Suc n)))) ----NS> r";
by (simp_tac (simpset() addsimps [LIMSEQ_NSLIMSEQ_iff RS sym,
LIMSEQ_inverse_real_of_posnat_add_minus_mult]) 1);
qed "NSLIMSEQ_inverse_real_of_posnat_add_minus_mult";
@@ -1214,22 +1214,22 @@
qed "LIMSEQ_pow";
(*----------------------------------------------------------------
- 0 <= x < Numeral1 ==> (x ^ n ----> 0)
+ 0 <= x < 1 ==> (x ^ n ----> 0)
Proof will use (NS) Cauchy equivalence for convergence and
also fact that bounded and monotonic sequence converges.
---------------------------------------------------------------*)
-Goalw [Bseq_def] "[| Numeral0 <= x; x < Numeral1 |] ==> Bseq (%n. x ^ n)";
-by (res_inst_tac [("x","Numeral1")] exI 1);
+Goalw [Bseq_def] "[| 0 <= x; x < 1 |] ==> Bseq (%n. x ^ n)";
+by (res_inst_tac [("x","1")] exI 1);
by (auto_tac (claset() addDs [conjI RS realpow_le]
addIs [order_less_imp_le],
simpset() addsimps [abs_eqI1, realpow_abs RS sym] ));
qed "Bseq_realpow";
-Goal "[| Numeral0 <= x; x < Numeral1 |] ==> monoseq (%n. x ^ n)";
+Goal "[| 0 <= x; x < 1 |] ==> monoseq (%n. x ^ n)";
by (blast_tac (claset() addSIs [mono_SucI2,realpow_Suc_le3]) 1);
qed "monoseq_realpow";
-Goal "[| Numeral0 <= x; x < Numeral1 |] ==> convergent (%n. x ^ n)";
+Goal "[| 0 <= x; x < 1 |] ==> convergent (%n. x ^ n)";
by (blast_tac (claset() addSIs [Bseq_monoseq_convergent,
Bseq_realpow,monoseq_realpow]) 1);
qed "convergent_realpow";
@@ -1238,7 +1238,7 @@
Goalw [NSLIMSEQ_def]
- "[| Numeral0 <= x; x < Numeral1 |] ==> (%n. x ^ n) ----NS> Numeral0";
+ "[| 0 <= x; x < 1 |] ==> (%n. x ^ n) ----NS> 0";
by (auto_tac (claset() addSDs [convergent_realpow],
simpset() addsimps [convergent_NSconvergent_iff]));
by (forward_tac [NSconvergentD] 1);
@@ -1258,12 +1258,12 @@
qed "NSLIMSEQ_realpow_zero";
(*--------------- standard version ---------------*)
-Goal "[| Numeral0 <= x; x < Numeral1 |] ==> (%n. x ^ n) ----> Numeral0";
+Goal "[| 0 <= x; x < 1 |] ==> (%n. x ^ n) ----> 0";
by (asm_simp_tac (simpset() addsimps [NSLIMSEQ_realpow_zero,
LIMSEQ_NSLIMSEQ_iff]) 1);
qed "LIMSEQ_realpow_zero";
-Goal "Numeral1 < x ==> (%n. a / (x ^ n)) ----> Numeral0";
+Goal "1 < x ==> (%n. a / (x ^ n)) ----> 0";
by (cut_inst_tac [("a","a"),("x1","inverse x")]
([LIMSEQ_const, LIMSEQ_realpow_zero] MRS LIMSEQ_mult) 1);
by (auto_tac (claset(),
@@ -1275,22 +1275,22 @@
(*----------------------------------------------------------------
Limit of c^n for |c| < 1
---------------------------------------------------------------*)
-Goal "abs(c) < Numeral1 ==> (%n. abs(c) ^ n) ----> Numeral0";
+Goal "abs(c) < 1 ==> (%n. abs(c) ^ n) ----> 0";
by (blast_tac (claset() addSIs [LIMSEQ_realpow_zero,abs_ge_zero]) 1);
qed "LIMSEQ_rabs_realpow_zero";
-Goal "abs(c) < Numeral1 ==> (%n. abs(c) ^ n) ----NS> Numeral0";
+Goal "abs(c) < 1 ==> (%n. abs(c) ^ n) ----NS> 0";
by (asm_full_simp_tac (simpset() addsimps [LIMSEQ_rabs_realpow_zero,
LIMSEQ_NSLIMSEQ_iff RS sym]) 1);
qed "NSLIMSEQ_rabs_realpow_zero";
-Goal "abs(c) < Numeral1 ==> (%n. c ^ n) ----> Numeral0";
+Goal "abs(c) < 1 ==> (%n. c ^ n) ----> 0";
by (rtac (LIMSEQ_rabs_zero RS iffD1) 1);
by (auto_tac (claset() addIs [LIMSEQ_rabs_realpow_zero],
simpset() addsimps [realpow_abs RS sym]));
qed "LIMSEQ_rabs_realpow_zero2";
-Goal "abs(c) < Numeral1 ==> (%n. c ^ n) ----NS> Numeral0";
+Goal "abs(c) < 1 ==> (%n. c ^ n) ----NS> 0";
by (asm_full_simp_tac (simpset() addsimps [LIMSEQ_rabs_realpow_zero2,
LIMSEQ_NSLIMSEQ_iff RS sym]) 1);
qed "NSLIMSEQ_rabs_realpow_zero2";
@@ -1308,12 +1308,12 @@
(*** A sequence converging to zero defines an infinitesimal ***)
Goalw [NSLIMSEQ_def]
- "X ----NS> Numeral0 ==> Abs_hypreal(hyprel``{X}) : Infinitesimal";
+ "X ----NS> 0 ==> Abs_hypreal(hyprel``{X}) : Infinitesimal";
by (dres_inst_tac [("x","whn")] bspec 1);
by (simp_tac (simpset() addsimps [HNatInfinite_whn]) 1);
by (auto_tac (claset(),
simpset() addsimps [hypnat_omega_def, mem_infmal_iff RS sym,
- starfunNat,hypreal_of_real_zero]));
+ starfunNat]));
qed "NSLIMSEQ_zero_Infinitesimal_hypreal";
(***---------------------------------------------------------------
--- a/src/HOL/Hyperreal/SEQ.thy Thu Nov 01 21:12:13 2001 +0100
+++ b/src/HOL/Hyperreal/SEQ.thy Fri Nov 02 17:55:24 2001 +0100
@@ -10,7 +10,7 @@
(* Standard definition of convergence of sequence *)
LIMSEQ :: [nat=>real,real] => bool ("((_)/ ----> (_))" [60, 60] 60)
- "X ----> L == (ALL r. Numeral0 < r --> (EX no. ALL n. no <= n --> abs (X n + -L) < r))"
+ "X ----> L == (ALL r. 0 < r --> (EX no. ALL n. no <= n --> abs (X n + -L) < r))"
(* Nonstandard definition of convergence of sequence *)
NSLIMSEQ :: [nat=>real,real] => bool ("((_)/ ----NS> (_))" [60, 60] 60)
@@ -33,7 +33,7 @@
(* Standard definition for bounded sequence *)
Bseq :: (nat => real) => bool
- "Bseq X == (EX K. (Numeral0 < K & (ALL n. abs(X n) <= K)))"
+ "Bseq X == (EX K. (0 < K & (ALL n. abs(X n) <= K)))"
(* Nonstandard definition for bounded sequence *)
NSBseq :: (nat=>real) => bool
@@ -52,7 +52,7 @@
(* Standard definition *)
Cauchy :: (nat => real) => bool
- "Cauchy X == (ALL e. (Numeral0 < e -->
+ "Cauchy X == (ALL e. (0 < e -->
(EX M. (ALL m n. M <= m & M <= n
--> abs((X m) + -(X n)) < e))))"
--- a/src/HOL/Hyperreal/Series.ML Thu Nov 01 21:12:13 2001 +0100
+++ b/src/HOL/Hyperreal/Series.ML Fri Nov 02 17:55:24 2001 +0100
@@ -5,13 +5,13 @@
Description : Finite summation and infinite series
*)
-Goal "sumr (Suc n) n f = Numeral0";
+Goal "sumr (Suc n) n f = 0";
by (induct_tac "n" 1);
by (Auto_tac);
qed "sumr_Suc_zero";
Addsimps [sumr_Suc_zero];
-Goal "sumr m m f = Numeral0";
+Goal "sumr m m f = 0";
by (induct_tac "m" 1);
by (Auto_tac);
qed "sumr_eq_bounds";
@@ -22,7 +22,7 @@
qed "sumr_Suc_eq";
Addsimps [sumr_Suc_eq];
-Goal "sumr (m+k) k f = Numeral0";
+Goal "sumr (m+k) k f = 0";
by (induct_tac "k" 1);
by (Auto_tac);
qed "sumr_add_lbound_zero";
@@ -83,7 +83,7 @@
by (full_simp_tac (simpset() addsimps [sumr_add_mult_const]) 1);
qed "sumr_diff_mult_const";
-Goal "n < m --> sumr m n f = Numeral0";
+Goal "n < m --> sumr m n f = 0";
by (induct_tac "n" 1);
by (auto_tac (claset() addDs [less_imp_le], simpset()));
qed_spec_mp "sumr_less_bounds_zero";
@@ -101,7 +101,7 @@
by (Auto_tac);
qed "sumr_shift_bounds";
-Goal "sumr 0 (2*n) (%i. (-1) ^ Suc i) = Numeral0";
+Goal "sumr 0 (2*n) (%i. (-1) ^ Suc i) = 0";
by (induct_tac "n" 1);
by (Auto_tac);
qed "sumr_minus_one_realpow_zero";
@@ -137,7 +137,7 @@
real_of_nat_Suc]) 1);
qed_spec_mp "sumr_interval_const2";
-Goal "(ALL n. m <= n --> Numeral0 <= f n) & m < k --> sumr 0 m f <= sumr 0 k f";
+Goal "(ALL n. m <= n --> 0 <= f n) & m < k --> sumr 0 m f <= sumr 0 k f";
by (induct_tac "k" 1);
by (Step_tac 1);
by (ALLGOALS(asm_full_simp_tac (simpset() addsimps [less_Suc_eq_le])));
@@ -156,21 +156,21 @@
simpset() addsimps [le_def]));
qed_spec_mp "sumr_le2";
-Goal "(ALL n. Numeral0 <= f n) --> Numeral0 <= sumr m n f";
+Goal "(ALL n. 0 <= f n) --> 0 <= sumr m n f";
by (induct_tac "n" 1);
by Auto_tac;
by (dres_inst_tac [("x","n")] spec 1);
by (arith_tac 1);
qed_spec_mp "sumr_ge_zero";
-Goal "(ALL n. m <= n --> Numeral0 <= f n) --> Numeral0 <= sumr m n f";
+Goal "(ALL n. m <= n --> 0 <= f n) --> 0 <= sumr m n f";
by (induct_tac "n" 1);
by Auto_tac;
by (dres_inst_tac [("x","n")] spec 1);
by (arith_tac 1);
qed_spec_mp "sumr_ge_zero2";
-Goal "Numeral0 <= sumr m n (%n. abs (f n))";
+Goal "0 <= sumr m n (%n. abs (f n))";
by (induct_tac "n" 1);
by Auto_tac;
by (arith_tac 1);
@@ -184,21 +184,21 @@
qed "rabs_sumr_rabs_cancel";
Addsimps [rabs_sumr_rabs_cancel];
-Goal "ALL n. N <= n --> f n = Numeral0 \
-\ ==> ALL m n. N <= m --> sumr m n f = Numeral0";
+Goal "ALL n. N <= n --> f n = 0 \
+\ ==> ALL m n. N <= m --> sumr m n f = 0";
by (Step_tac 1);
by (induct_tac "n" 1);
by (Auto_tac);
qed "sumr_zero";
-Goal "ALL n. N <= n --> f (Suc n) = Numeral0 \
-\ ==> ALL m n. Suc N <= m --> sumr m n f = Numeral0";
+Goal "ALL n. N <= n --> f (Suc n) = 0 \
+\ ==> ALL m n. Suc N <= m --> sumr m n f = 0";
by (rtac sumr_zero 1 THEN Step_tac 1);
by (case_tac "n" 1);
by Auto_tac;
qed "Suc_le_imp_diff_ge2";
-Goal "sumr (Suc 0) n (%n. f(n) * Numeral0 ^ n) = Numeral0";
+Goal "sumr (Suc 0) n (%n. f(n) * 0 ^ n) = 0";
by (induct_tac "n" 1);
by (case_tac "n" 2);
by Auto_tac;
@@ -269,7 +269,7 @@
(*
Goalw [sums_def,LIMSEQ_def]
- "(ALL m. n <= Suc m --> f(m) = Numeral0) ==> f sums (sumr 0 n f)";
+ "(ALL m. n <= Suc m --> f(m) = 0) ==> f sums (sumr 0 n f)";
by (Step_tac 1);
by (res_inst_tac [("x","n")] exI 1);
by (Step_tac 1 THEN forward_tac [le_imp_less_or_eq] 1);
@@ -283,7 +283,7 @@
**********************)
Goalw [sums_def,LIMSEQ_def]
- "(ALL m. n <= m --> f(m) = Numeral0) ==> f sums (sumr 0 n f)";
+ "(ALL m. n <= m --> f(m) = 0) ==> f sums (sumr 0 n f)";
by (Step_tac 1);
by (res_inst_tac [("x","n")] exI 1);
by (Step_tac 1 THEN forward_tac [le_imp_less_or_eq] 1);
@@ -341,7 +341,7 @@
by (Auto_tac);
qed "sums_group";
-Goal "[|summable f; ALL d. Numeral0 < (f(n + (Suc (Suc 0) * d))) + f(n + ((Suc (Suc 0) * d) + 1))|] \
+Goal "[|summable f; ALL d. 0 < (f(n + (Suc (Suc 0) * d))) + f(n + ((Suc (Suc 0) * d) + 1))|] \
\ ==> sumr 0 n f < suminf f";
by (dtac summable_sums 1);
by (auto_tac (claset(),simpset() addsimps [sums_def,LIMSEQ_def]));
@@ -369,8 +369,8 @@
by (assume_tac 3);
by (dres_inst_tac [("x","0")] spec 2);
by (Asm_full_simp_tac 2);
-by (subgoal_tac "Numeral0 <= sumr 0 (Suc (Suc 0) * Suc no + n) f + - suminf f" 1);
-by (dtac (rename_numerals abs_eqI1) 1 );
+by (subgoal_tac "0 <= sumr 0 (Suc (Suc 0) * Suc no + n) f + - suminf f" 1);
+by (dtac (abs_eqI1) 1 );
by (Asm_full_simp_tac 1);
by (auto_tac (claset(),simpset() addsimps [real_le_def]));
qed "sumr_pos_lt_pair";
@@ -379,7 +379,7 @@
Summable series of positive terms has limit >= any partial sum
----------------------------------------------------------------*)
Goal
- "[| summable f; ALL m. n <= m --> Numeral0 <= f(m) |] \
+ "[| summable f; ALL m. n <= m --> 0 <= f(m) |] \
\ ==> sumr 0 n f <= suminf f";
by (dtac summable_sums 1);
by (rewtac sums_def);
@@ -390,7 +390,7 @@
by (auto_tac (claset() addIs [sumr_le], simpset()));
qed "series_pos_le";
-Goal "[| summable f; ALL m. n <= m --> Numeral0 < f(m) |] \
+Goal "[| summable f; ALL m. n <= m --> 0 < f(m) |] \
\ ==> sumr 0 n f < suminf f";
by (res_inst_tac [("y","sumr 0 (Suc n) f")] order_less_le_trans 1);
by (rtac series_pos_le 2);
@@ -403,10 +403,10 @@
sum of geometric progression
-------------------------------------------------------------------*)
-Goal "x ~= Numeral1 ==> sumr 0 n (%n. x ^ n) = (x ^ n - Numeral1) / (x - Numeral1)";
+Goal "x ~= 1 ==> sumr 0 n (%n. x ^ n) = (x ^ n - 1) / (x - 1)";
by (induct_tac "n" 1);
by (Auto_tac);
-by (res_inst_tac [("c1","x - Numeral1")] (real_mult_right_cancel RS iffD1) 1);
+by (res_inst_tac [("c1","x - 1")] (real_mult_right_cancel RS iffD1) 1);
by (auto_tac (claset(),
simpset() addsimps [real_mult_assoc, real_add_mult_distrib]));
by (auto_tac (claset(),
@@ -414,14 +414,14 @@
real_diff_def, real_mult_commute]));
qed "sumr_geometric";
-Goal "abs(x) < Numeral1 ==> (%n. x ^ n) sums (Numeral1/(Numeral1 - x))";
-by (case_tac "x = Numeral1" 1);
+Goal "abs(x) < 1 ==> (%n. x ^ n) sums (1/(1 - x))";
+by (case_tac "x = 1" 1);
by (auto_tac (claset() addSDs [LIMSEQ_rabs_realpow_zero2],
simpset() addsimps [sumr_geometric ,sums_def,
real_diff_def, real_add_divide_distrib]));
-by (subgoal_tac "Numeral1 / (Numeral1 + - x) = Numeral0/(x-Numeral1) + - Numeral1/(x-Numeral1)" 1);
+by (subgoal_tac "1 / (1 + -x) = 0/(x - 1) + - 1/(x - 1)" 1);
by (asm_full_simp_tac (simpset() addsimps [real_divide_eq_cancel1,
- real_divide_minus_eq RS sym, real_diff_def]) 2);
+ real_divide_minus_eq RS sym, real_diff_def]) 2);
by (etac ssubst 1);
by (rtac LIMSEQ_add 1 THEN rtac LIMSEQ_divide 1);
by (auto_tac (claset() addIs [LIMSEQ_const],
@@ -437,7 +437,7 @@
qed "summable_convergent_sumr_iff";
Goal "summable f = \
-\ (ALL e. Numeral0 < e --> (EX N. ALL m n. N <= m --> abs(sumr m n f) < e))";
+\ (ALL e. 0 < e --> (EX N. ALL m n. N <= m --> abs(sumr m n f) < e))";
by (auto_tac (claset(),simpset() addsimps [summable_convergent_sumr_iff,
Cauchy_convergent_iff RS sym,Cauchy_def]));
by (ALLGOALS(dtac spec) THEN Auto_tac);
@@ -455,7 +455,7 @@
(*-------------------------------------------------------------------
Terms of a convergent series tend to zero
- > Goalw [LIMSEQ_def] "summable f ==> f ----> Numeral0";
+ > Goalw [LIMSEQ_def] "summable f ==> f ----> 0";
Proved easily in HSeries after proving nonstandard case.
-------------------------------------------------------------------*)
(*-------------------------------------------------------------------
@@ -527,10 +527,10 @@
The ratio test
-------------------------------------------------------------------*)
-Goal "[| c <= Numeral0; abs x <= c * abs y |] ==> x = (Numeral0::real)";
+Goal "[| c <= 0; abs x <= c * abs y |] ==> x = (0::real)";
by (dtac order_le_imp_less_or_eq 1);
by Auto_tac;
-by (subgoal_tac "Numeral0 <= c * abs y" 1);
+by (subgoal_tac "0 <= c * abs y" 1);
by (arith_tac 2);
by (asm_full_simp_tac (simpset() addsimps [real_0_le_mult_iff]) 1);
qed "rabs_ratiotest_lemma";
@@ -546,19 +546,19 @@
by (auto_tac (claset(),simpset() addsimps [le_Suc_ex]));
qed "le_Suc_ex_iff";
-(*All this trouble just to get Numeral0<c *)
+(*All this trouble just to get 0<c *)
Goal "[| ALL n. N <= n --> abs(f(Suc n)) <= c*abs(f n) |] \
-\ ==> Numeral0 < c | summable f";
+\ ==> 0 < c | summable f";
by (simp_tac (simpset() addsimps [linorder_not_le RS sym]) 1);
by (asm_full_simp_tac (simpset() addsimps [summable_Cauchy]) 1);
-by (Step_tac 1 THEN subgoal_tac "ALL n. N <= n --> f (Suc n) = Numeral0" 1);
+by (Step_tac 1 THEN subgoal_tac "ALL n. N <= n --> f (Suc n) = 0" 1);
by (blast_tac (claset() addIs [rabs_ratiotest_lemma]) 2);
by (res_inst_tac [("x","Suc N")] exI 1);
by (Clarify_tac 1);
by (dtac Suc_le_imp_diff_ge2 1 THEN Auto_tac);
qed "ratio_test_lemma2";
-Goal "[| c < Numeral1; ALL n. N <= n --> abs(f(Suc n)) <= c*abs(f n) |] \
+Goal "[| c < 1; ALL n. N <= n --> abs(f(Suc n)) <= c*abs(f n) |] \
\ ==> summable f";
by (forward_tac [ratio_test_lemma2] 1);
by Auto_tac;
@@ -566,14 +566,13 @@
summable_comparison_test 1);
by (res_inst_tac [("x","N")] exI 1 THEN Step_tac 1);
by (dtac (le_Suc_ex_iff RS iffD1) 1);
-by (auto_tac (claset(),simpset() addsimps [realpow_add,
- rename_numerals realpow_not_zero]));
+by (auto_tac (claset(),simpset() addsimps [realpow_add, realpow_not_zero]));
by (induct_tac "na" 1 THEN Auto_tac);
by (res_inst_tac [("y","c*abs(f(N + n))")] order_trans 1);
by (auto_tac (claset() addIs [real_mult_le_mono1],
simpset() addsimps [summable_def]));
by (asm_full_simp_tac (simpset() addsimps real_mult_ac) 1);
-by (res_inst_tac [("x","abs(f N) * (Numeral1/(Numeral1 - c)) / (c ^ N)")] exI 1);
+by (res_inst_tac [("x","abs(f N) * (1/(1 - c)) / (c ^ N)")] exI 1);
by (rtac sums_divide 1);
by (rtac sums_mult 1);
by (auto_tac (claset() addSIs [sums_mult,geometric_sums],
--- a/src/HOL/Hyperreal/Series.thy Thu Nov 01 21:12:13 2001 +0100
+++ b/src/HOL/Hyperreal/Series.thy Fri Nov 02 17:55:24 2001 +0100
@@ -9,8 +9,8 @@
consts sumr :: "[nat,nat,(nat=>real)] => real"
primrec
- sumr_0 "sumr m 0 f = Numeral0"
- sumr_Suc "sumr m (Suc n) f = (if n < m then Numeral0
+ sumr_0 "sumr m 0 f = 0"
+ sumr_Suc "sumr m (Suc n) f = (if n < m then 0
else sumr m n f + f(n))"
constdefs
--- a/src/HOL/Hyperreal/Star.ML Thu Nov 01 21:12:13 2001 +0100
+++ b/src/HOL/Hyperreal/Star.ML Fri Nov 02 17:55:24 2001 +0100
@@ -176,8 +176,7 @@
by (rtac bexI 1 THEN rtac lemma_hyprel_refl 2);
by (rtac bexI 1 THEN rtac lemma_hyprel_refl 2);
by (auto_tac (claset() addSDs [spec],
- simpset() addsimps [hypreal_minus,hrabs_def,
- rename_numerals hypreal_zero_def,
+ simpset() addsimps [hypreal_minus,hrabs_def, hypreal_zero_def,
hypreal_le_def, hypreal_less_def]));
by (TRYALL(Ultra_tac));
by (TRYALL(arith_tac));
--- a/src/HOL/Hyperreal/hypreal_arith0.ML Thu Nov 01 21:12:13 2001 +0100
+++ b/src/HOL/Hyperreal/hypreal_arith0.ML Fri Nov 02 17:55:24 2001 +0100
@@ -9,24 +9,23 @@
local
(* reduce contradictory <= to False *)
-val simps =
- [order_less_irrefl, zero_eq_numeral_0, one_eq_numeral_1,
+val add_rules =
+ [order_less_irrefl, hypreal_numeral_0_eq_0, hypreal_numeral_1_eq_1,
add_hypreal_number_of, minus_hypreal_number_of, diff_hypreal_number_of,
mult_hypreal_number_of, eq_hypreal_number_of, less_hypreal_number_of,
le_hypreal_number_of_eq_not_less, hypreal_diff_def,
- hypreal_minus_add_distrib, hypreal_minus_minus, hypreal_mult_assoc];
-
-val add_rules =
- map rename_numerals
- [hypreal_add_zero_left, hypreal_add_zero_right,
- hypreal_add_minus, hypreal_add_minus_left,
- hypreal_mult_0, hypreal_mult_0_right,
- hypreal_mult_1, hypreal_mult_1_right,
- hypreal_mult_minus_1, hypreal_mult_minus_1_right];
+ hypreal_minus_add_distrib, hypreal_minus_minus, hypreal_mult_assoc,
+ hypreal_minus_zero,
+ hypreal_add_zero_left, hypreal_add_zero_right,
+ hypreal_add_minus, hypreal_add_minus_left,
+ hypreal_mult_0, hypreal_mult_0_right,
+ hypreal_mult_1, hypreal_mult_1_right,
+ hypreal_mult_minus_1, hypreal_mult_minus_1_right];
val simprocs = [Hyperreal_Times_Assoc.conv,
Hyperreal_Numeral_Simprocs.combine_numerals]@
- Hyperreal_Numeral_Simprocs.cancel_numerals;
+ Hyperreal_Numeral_Simprocs.cancel_numerals @
+ Hyperreal_Numeral_Simprocs.eval_numerals;
val mono_ss = simpset() addsimps
[hypreal_add_le_mono,hypreal_add_less_mono,
@@ -69,8 +68,7 @@
mult_mono_thms = mult_mono_thms @ hypreal_mult_mono_thms,
inj_thms = inj_thms, (*FIXME: add hypreal*)
lessD = lessD, (*We don't change LA_Data_Ref.lessD because the hypreal ordering is dense!*)
- simpset = simpset addsimps (add_rules @ simps)
- addsimprocs simprocs}),
+ simpset = simpset addsimps add_rules addsimprocs simprocs}),
arith_discrete ("HyperDef.hypreal",false),
Simplifier.change_simpset_of (op addsimprocs) [fast_hypreal_arith_simproc]];
--- a/src/HOL/Integ/Bin.ML Thu Nov 01 21:12:13 2001 +0100
+++ b/src/HOL/Integ/Bin.ML Fri Nov 02 17:55:24 2001 +0100
@@ -189,6 +189,12 @@
Addsimps (map (inst "y" "number_of ?v")
[zminus_zless, zminus_zle, zminus_equation]);
+(*Equations and inequations involving 1*)
+Addsimps (map (simplify (simpset()) o inst "x" "1")
+ [zless_zminus, zle_zminus, equation_zminus]);
+Addsimps (map (simplify (simpset()) o inst "y" "1")
+ [zminus_zless, zminus_zle, zminus_equation]);
+
(*Moving negation out of products*)
Addsimps [zmult_zminus, zmult_zminus_right];
--- a/src/HOL/Integ/IntArith.ML Thu Nov 01 21:12:13 2001 +0100
+++ b/src/HOL/Integ/IntArith.ML Fri Nov 02 17:55:24 2001 +0100
@@ -4,6 +4,11 @@
*)
+Goal "x - - y = x + (y::int)";
+by (Simp_tac 1);
+qed "int_diff_minus_eq";
+Addsimps [int_diff_minus_eq];
+
Goal "abs(abs(x::int)) = abs(x)";
by(arith_tac 1);
qed "abs_abs";
--- a/src/HOL/Integ/int_arith1.ML Thu Nov 01 21:12:13 2001 +0100
+++ b/src/HOL/Integ/int_arith1.ML Fri Nov 02 17:55:24 2001 +0100
@@ -407,7 +407,7 @@
val add_rules =
simp_thms @ bin_arith_simps @ bin_rel_simps @
[int_numeral_0_eq_0, int_numeral_1_eq_1,
- zadd_0, zadd_0_right, zdiff_def,
+ zminus_0, zadd_0, zadd_0_right, zdiff_def,
zadd_zminus_inverse, zadd_zminus_inverse2,
zmult_0, zmult_0_right,
zmult_1, zmult_1_right,
--- a/src/HOL/Integ/int_factor_simprocs.ML Thu Nov 01 21:12:13 2001 +0100
+++ b/src/HOL/Integ/int_factor_simprocs.ML Fri Nov 02 17:55:24 2001 +0100
@@ -44,10 +44,10 @@
val dest_coeff = dest_coeff 1
val trans_tac = trans_tac
val norm_tac =
- ALLGOALS (simp_tac (HOL_ss addsimps mult_1s))
+ ALLGOALS (simp_tac (HOL_ss addsimps int_minus_from_mult_simps@mult_1s))
THEN ALLGOALS (simp_tac (HOL_ss addsimps bin_simps@int_mult_minus_simps))
THEN ALLGOALS
- (simp_tac (HOL_ss addsimps int_minus_from_mult_simps@zmult_ac))
+ (simp_tac (HOL_ss addsimps zmult_ac))
val numeral_simp_tac = ALLGOALS (simp_tac (HOL_ss addsimps bin_simps))
val simplify_meta_eq = simplify_meta_eq mult_1s
end
--- a/src/HOL/Real/HahnBanach/Aux.thy Thu Nov 01 21:12:13 2001 +0100
+++ b/src/HOL/Real/HahnBanach/Aux.thy Fri Nov 02 17:55:24 2001 +0100
@@ -40,20 +40,20 @@
text {* \medskip Some lemmas for the reals. *}
-lemma real_add_minus_eq: "x - y = (Numeral0::real) \<Longrightarrow> x = y"
+lemma real_add_minus_eq: "x - y = (0::real) \<Longrightarrow> x = y"
by simp
-lemma abs_minus_one: "abs (- (Numeral1::real)) = Numeral1"
+lemma abs_minus_one: "abs (- (1::real)) = 1"
by simp
lemma real_mult_le_le_mono1a:
- "(Numeral0::real) \<le> z \<Longrightarrow> x \<le> y \<Longrightarrow> z * x \<le> z * y"
+ "(0::real) \<le> z \<Longrightarrow> x \<le> y \<Longrightarrow> z * x \<le> z * y"
by (simp add: real_mult_le_mono2)
lemma real_mult_le_le_mono2:
- "(Numeral0::real) \<le> z \<Longrightarrow> x \<le> y \<Longrightarrow> x * z \<le> y * z"
+ "(0::real) \<le> z \<Longrightarrow> x \<le> y \<Longrightarrow> x * z \<le> y * z"
proof -
- assume "(Numeral0::real) \<le> z" "x \<le> y"
+ assume "(0::real) \<le> z" "x \<le> y"
hence "x < y \<or> x = y" by (auto simp add: order_le_less)
thus ?thesis
proof
@@ -66,11 +66,11 @@
qed
lemma real_mult_less_le_anti:
- "z < (Numeral0::real) \<Longrightarrow> x \<le> y \<Longrightarrow> z * y \<le> z * x"
+ "z < (0::real) \<Longrightarrow> x \<le> y \<Longrightarrow> z * y \<le> z * x"
proof -
- assume "z < Numeral0" "x \<le> y"
- hence "Numeral0 < - z" by simp
- hence "Numeral0 \<le> - z" by (rule order_less_imp_le)
+ assume "z < 0" "x \<le> y"
+ hence "0 < - z" by simp
+ hence "0 \<le> - z" by (rule order_less_imp_le)
hence "x * (- z) \<le> y * (- z)"
by (rule real_mult_le_le_mono2)
hence "- (x * z) \<le> - (y * z)"
@@ -79,31 +79,23 @@
qed
lemma real_mult_less_le_mono:
- "(Numeral0::real) < z \<Longrightarrow> x \<le> y \<Longrightarrow> z * x \<le> z * y"
+ "(0::real) < z \<Longrightarrow> x \<le> y \<Longrightarrow> z * x \<le> z * y"
proof -
- assume "Numeral0 < z" "x \<le> y"
- have "Numeral0 \<le> z" by (rule order_less_imp_le)
+ assume "0 < z" "x \<le> y"
+ have "0 \<le> z" by (rule order_less_imp_le)
hence "x * z \<le> y * z"
by (rule real_mult_le_le_mono2)
thus ?thesis by (simp only: real_mult_commute)
qed
-lemma real_inverse_gt_zero1: "Numeral0 < (x::real) \<Longrightarrow> Numeral0 < inverse x"
-proof -
- assume "Numeral0 < x"
- have "0 < x" by simp
- hence "0 < inverse x" by (rule real_inverse_gt_zero)
- thus ?thesis by simp
-qed
-
-lemma real_mult_inv_right1: "(x::real) \<noteq> Numeral0 \<Longrightarrow> x * inverse x = Numeral1"
+lemma real_mult_inv_right1: "(x::real) \<noteq> 0 \<Longrightarrow> x * inverse x = 1"
by simp
-lemma real_mult_inv_left1: "(x::real) \<noteq> Numeral0 \<Longrightarrow> inverse x * x = Numeral1"
+lemma real_mult_inv_left1: "(x::real) \<noteq> 0 \<Longrightarrow> inverse x * x = 1"
by simp
lemma real_le_mult_order1a:
- "(Numeral0::real) \<le> x \<Longrightarrow> Numeral0 \<le> y \<Longrightarrow> Numeral0 \<le> x * y"
+ "(0::real) \<le> x \<Longrightarrow> 0 \<le> y \<Longrightarrow> 0 \<le> x * y"
by (simp add: real_0_le_mult_iff)
lemma real_mult_diff_distrib:
--- a/src/HOL/Real/HahnBanach/FunctionNorm.thy Thu Nov 01 21:12:13 2001 +0100
+++ b/src/HOL/Real/HahnBanach/FunctionNorm.thy Fri Nov 02 17:55:24 2001 +0100
@@ -73,7 +73,7 @@
constdefs
B :: "'a set \<Rightarrow> ('a \<Rightarrow> real) \<Rightarrow> ('a::{plus, minus, zero} \<Rightarrow> real) \<Rightarrow> real set"
"B V norm f \<equiv>
- {Numeral0} \<union> {\<bar>f x\<bar> * inverse (norm x) | x. x \<noteq> 0 \<and> x \<in> V}"
+ {0} \<union> {\<bar>f x\<bar> * inverse (norm x) | x. x \<noteq> 0 \<and> x \<in> V}"
text {*
@{text n} is the function norm of @{text f}, iff @{text n} is the
@@ -97,7 +97,7 @@
syntax
function_norm :: "('a \<Rightarrow> real) \<Rightarrow> 'a set \<Rightarrow> ('a \<Rightarrow> real) \<Rightarrow> real" ("\<parallel>_\<parallel>_,_")
-lemma B_not_empty: "Numeral0 \<in> B V norm f"
+lemma B_not_empty: "0 \<in> B V norm f"
by (unfold B_def) blast
text {*
@@ -125,7 +125,7 @@
show "\<exists>X. X \<in> B V norm f"
proof
- show "Numeral0 \<in> (B V norm f)" by (unfold B_def) blast
+ show "0 \<in> (B V norm f)" by (unfold B_def) blast
qed
txt {* Then we have to show that @{text B} is bounded: *}
@@ -136,7 +136,7 @@
txt {* We know that @{text f} is bounded by some value @{text c}. *}
fix c assume a: "\<forall>x \<in> V. \<bar>f x\<bar> \<le> c * norm x"
- def b \<equiv> "max c Numeral0"
+ def b \<equiv> "max c 0"
show "?thesis"
proof (intro exI isUbI setleI ballI, unfold B_def,
@@ -148,7 +148,7 @@
two cases for @{text "y \<in> B"}. If @{text "y = 0"} then
@{text "y \<le> max c 0"}: *}
- fix y assume "y = (Numeral0::real)"
+ fix y assume "y = (0::real)"
show "y \<le> b" by (simp! add: le_maxI2)
txt {* The second case is @{text "y = \<bar>f x\<bar> / \<parallel>x\<parallel>"} for some
@@ -164,16 +164,16 @@
assume "y = \<bar>f x\<bar> * inverse (norm x)"
also have "... \<le> c * norm x * inverse (norm x)"
proof (rule real_mult_le_le_mono2)
- show "Numeral0 \<le> inverse (norm x)"
- by (rule order_less_imp_le, rule real_inverse_gt_zero1,
+ show "0 \<le> inverse (norm x)"
+ by (rule order_less_imp_le, rule real_inverse_gt_0,
rule normed_vs_norm_gt_zero)
from a show "\<bar>f x\<bar> \<le> c * norm x" ..
qed
also have "... = c * (norm x * inverse (norm x))"
by (rule real_mult_assoc)
- also have "(norm x * inverse (norm x)) = (Numeral1::real)"
+ also have "(norm x * inverse (norm x)) = (1::real)"
proof (rule real_mult_inv_right1)
- show nz: "norm x \<noteq> Numeral0"
+ show nz: "norm x \<noteq> 0"
by (rule not_sym, rule lt_imp_not_eq,
rule normed_vs_norm_gt_zero)
qed
@@ -188,7 +188,7 @@
lemma fnorm_ge_zero [intro?]:
"is_continuous V norm f \<Longrightarrow> is_normed_vectorspace V norm
- \<Longrightarrow> Numeral0 \<le> \<parallel>f\<parallel>V,norm"
+ \<Longrightarrow> 0 \<le> \<parallel>f\<parallel>V,norm"
proof -
assume c: "is_continuous V norm f"
and n: "is_normed_vectorspace V norm"
@@ -200,23 +200,23 @@
show ?thesis
proof (unfold function_norm_def, rule sup_ub1)
- show "\<forall>x \<in> (B V norm f). Numeral0 \<le> x"
+ show "\<forall>x \<in> (B V norm f). 0 \<le> x"
proof (intro ballI, unfold B_def,
elim UnE singletonE CollectE exE conjE)
fix x r
assume "x \<in> V" "x \<noteq> 0"
and r: "r = \<bar>f x\<bar> * inverse (norm x)"
- have ge: "Numeral0 \<le> \<bar>f x\<bar>" by (simp! only: abs_ge_zero)
- have "Numeral0 \<le> inverse (norm x)"
- by (rule order_less_imp_le, rule real_inverse_gt_zero1, rule)(***
+ have ge: "0 \<le> \<bar>f x\<bar>" by (simp! only: abs_ge_zero)
+ have "0 \<le> inverse (norm x)"
+ by (rule order_less_imp_le, rule real_inverse_gt_0, rule)(***
proof (rule order_less_imp_le);
- show "Numeral0 < inverse (norm x)";
+ show "0 < inverse (norm x)";
proof (rule real_inverse_gt_zero);
- show "Numeral0 < norm x"; ..;
+ show "0 < norm x"; ..;
qed;
qed; ***)
- with ge show "Numeral0 \<le> r"
+ with ge show "0 \<le> r"
by (simp only: r, rule real_le_mult_order1a)
qed (simp!)
@@ -228,7 +228,7 @@
txt {* @{text B} is non-empty by construction: *}
- show "Numeral0 \<in> B V norm f" by (rule B_not_empty)
+ show "0 \<in> B V norm f" by (rule B_not_empty)
qed
qed
@@ -258,20 +258,20 @@
assume "x = 0"
have "\<bar>f x\<bar> = \<bar>f 0\<bar>" by (simp!)
- also from v continuous_linearform have "f 0 = Numeral0" ..
+ also from v continuous_linearform have "f 0 = 0" ..
also note abs_zero
- also have "Numeral0 \<le> \<parallel>f\<parallel>V,norm * norm x"
+ also have "0 \<le> \<parallel>f\<parallel>V,norm * norm x"
proof (rule real_le_mult_order1a)
- show "Numeral0 \<le> \<parallel>f\<parallel>V,norm" ..
- show "Numeral0 \<le> norm x" ..
+ show "0 \<le> \<parallel>f\<parallel>V,norm" ..
+ show "0 \<le> norm x" ..
qed
finally
show "\<bar>f x\<bar> \<le> \<parallel>f\<parallel>V,norm * norm x" .
next
assume "x \<noteq> 0"
- have n: "Numeral0 < norm x" ..
- hence nz: "norm x \<noteq> Numeral0"
+ have n: "0 < norm x" ..
+ hence nz: "norm x \<noteq> 0"
by (simp only: lt_imp_not_eq)
txt {* For the case @{text "x \<noteq> 0"} we derive the following fact
@@ -289,8 +289,8 @@
txt {* The thesis now follows by a short calculation: *}
- have "\<bar>f x\<bar> = \<bar>f x\<bar> * Numeral1" by (simp!)
- also from nz have "Numeral1 = inverse (norm x) * norm x"
+ have "\<bar>f x\<bar> = \<bar>f x\<bar> * 1" by (simp!)
+ also from nz have "1 = inverse (norm x) * norm x"
by (simp add: real_mult_inv_left1)
also have "\<bar>f x\<bar> * ... = \<bar>f x\<bar> * inverse (norm x) * norm x"
by (simp! add: real_mult_assoc)
@@ -310,13 +310,13 @@
lemma fnorm_le_ub:
"is_continuous V norm f \<Longrightarrow> is_normed_vectorspace V norm \<Longrightarrow>
- \<forall>x \<in> V. \<bar>f x\<bar> \<le> c * norm x \<Longrightarrow> Numeral0 \<le> c
+ \<forall>x \<in> V. \<bar>f x\<bar> \<le> c * norm x \<Longrightarrow> 0 \<le> c
\<Longrightarrow> \<parallel>f\<parallel>V,norm \<le> c"
proof (unfold function_norm_def)
assume "is_normed_vectorspace V norm"
assume c: "is_continuous V norm f"
assume fb: "\<forall>x \<in> V. \<bar>f x\<bar> \<le> c * norm x"
- and "Numeral0 \<le> c"
+ and "0 \<le> c"
txt {* Suppose the inequation holds for some @{text "c \<ge> 0"}. If
@{text c} is an upper bound of @{text B}, then @{text c} is greater
@@ -340,7 +340,7 @@
txt {* The first case for @{text "y \<in> B"} is @{text "y = 0"}. *}
- assume "y = Numeral0"
+ assume "y = 0"
show "y \<le> c" by (blast!)
txt{* The second case is @{text "y = \<bar>f x\<bar> / \<parallel>x\<parallel>"} for some
@@ -350,18 +350,18 @@
fix x
assume "x \<in> V" "x \<noteq> 0"
- have lz: "Numeral0 < norm x"
+ have lz: "0 < norm x"
by (simp! add: normed_vs_norm_gt_zero)
- have nz: "norm x \<noteq> Numeral0"
+ have nz: "norm x \<noteq> 0"
proof (rule not_sym)
- from lz show "Numeral0 \<noteq> norm x"
+ from lz show "0 \<noteq> norm x"
by (simp! add: order_less_imp_not_eq)
qed
- from lz have "Numeral0 < inverse (norm x)"
- by (simp! add: real_inverse_gt_zero1)
- hence inverse_gez: "Numeral0 \<le> inverse (norm x)"
+ from lz have "0 < inverse (norm x)"
+ by (simp! add: real_inverse_gt_0)
+ hence inverse_gez: "0 \<le> inverse (norm x)"
by (rule order_less_imp_le)
assume "y = \<bar>f x\<bar> * inverse (norm x)"
--- a/src/HOL/Real/HahnBanach/HahnBanach.thy Thu Nov 01 21:12:13 2001 +0100
+++ b/src/HOL/Real/HahnBanach/HahnBanach.thy Fri Nov 02 17:55:24 2001 +0100
@@ -201,7 +201,7 @@
proof (rule graph_extI)
fix t assume "t \<in> H"
have "(SOME (y, a). t = y + a \<cdot> x' \<and> y \<in> H)
- = (t, Numeral0)"
+ = (t, 0)"
by (rule decomp_H'_H) (assumption+, rule x')
thus "h t = h' t" by (simp! add: Let_def)
next
@@ -255,12 +255,12 @@
proof (rule graph_extI)
fix x assume "x \<in> F"
have "f x = h x" ..
- also have " ... = h x + Numeral0 * xi" by simp
+ also have " ... = h x + 0 * xi" by simp
also
- have "... = (let (y, a) = (x, Numeral0) in h y + a * xi)"
+ have "... = (let (y, a) = (x, 0) in h y + a * xi)"
by (simp add: Let_def)
also have
- "(x, Numeral0) = (SOME (y, a). x = y + a \<cdot> x' \<and> y \<in> H)"
+ "(x, 0) = (SOME (y, a). x = y + a \<cdot> x' \<and> y \<in> H)"
by (rule decomp_H'_H [symmetric]) (simp! add: x')+
also have
"(let (y, a) = (SOME (y, a). x = y + a \<cdot> x' \<and> y \<in> H)
@@ -372,10 +372,10 @@
txt {* @{text p} is positive definite: *}
-show "Numeral0 \<le> p x"
+show "0 \<le> p x"
proof (unfold p_def, rule real_le_mult_order1a)
- from f_cont f_norm show "Numeral0 \<le> \<parallel>f\<parallel>F,norm" ..
- show "Numeral0 \<le> norm x" ..
+ from f_cont f_norm show "0 \<le> \<parallel>f\<parallel>F,norm" ..
+ show "0 \<le> norm x" ..
qed
txt {* @{text p} is absolutely homogenous: *}
@@ -402,7 +402,7 @@
also
have "... \<le> \<parallel>f\<parallel>F,norm * (norm x + norm y)"
proof (rule real_mult_le_le_mono1a)
- from f_cont f_norm show "Numeral0 \<le> \<parallel>f\<parallel>F,norm" ..
+ from f_cont f_norm show "0 \<le> \<parallel>f\<parallel>F,norm" ..
show "norm (x + y) \<le> norm x + norm y" ..
qed
also have "... = \<parallel>f\<parallel>F,norm * norm x
@@ -489,7 +489,7 @@
with g_cont e_norm show "?L \<le> ?R"
proof (rule fnorm_le_ub)
- from f_cont f_norm show "Numeral0 \<le> \<parallel>f\<parallel>F,norm" ..
+ from f_cont f_norm show "0 \<le> \<parallel>f\<parallel>F,norm" ..
qed
txt{* The other direction is achieved by a similar
@@ -509,7 +509,7 @@
qed
thus "?R \<le> ?L"
proof (rule fnorm_le_ub [OF f_cont f_norm])
- from g_cont show "Numeral0 \<le> \<parallel>g\<parallel>E,norm" ..
+ from g_cont show "0 \<le> \<parallel>g\<parallel>E,norm" ..
qed
qed
qed
--- a/src/HOL/Real/HahnBanach/HahnBanachExtLemmas.thy Thu Nov 01 21:12:13 2001 +0100
+++ b/src/HOL/Real/HahnBanach/HahnBanachExtLemmas.thy Fri Nov 02 17:55:24 2001 +0100
@@ -279,14 +279,14 @@
also have "... \<le> p (y + a \<cdot> x0)"
proof (rule linorder_cases)
- assume z: "a = Numeral0"
+ assume z: "a = 0"
with vs y a show ?thesis by simp
txt {* In the case @{text "a < 0"}, we use @{text "a\<^sub>1"}
with @{text ya} taken as @{text "y / a"}: *}
next
- assume lz: "a < Numeral0" hence nz: "a \<noteq> Numeral0" by simp
+ assume lz: "a < 0" hence nz: "a \<noteq> 0" by simp
from a1
have "- p (inverse a \<cdot> y + x0) - h (inverse a \<cdot> y) \<le> xi"
by (rule bspec) (simp!)
@@ -315,7 +315,7 @@
with @{text ya} taken as @{text "y / a"}: *}
next
- assume gz: "Numeral0 < a" hence nz: "a \<noteq> Numeral0" by simp
+ assume gz: "0 < a" hence nz: "a \<noteq> 0" by simp
from a2 have "xi \<le> p (inverse a \<cdot> y + x0) - h (inverse a \<cdot> y)"
by (rule bspec) (simp!)
--- a/src/HOL/Real/HahnBanach/Linearform.thy Thu Nov 01 21:12:13 2001 +0100
+++ b/src/HOL/Real/HahnBanach/Linearform.thy Fri Nov 02 17:55:24 2001 +0100
@@ -37,8 +37,8 @@
\<Longrightarrow> f (- x) = - f x"
proof -
assume "is_linearform V f" "is_vectorspace V" "x \<in> V"
- have "f (- x) = f ((- Numeral1) \<cdot> x)" by (simp! add: negate_eq1)
- also have "... = (- Numeral1) * (f x)" by (rule linearform_mult)
+ have "f (- x) = f ((- 1) \<cdot> x)" by (simp! add: negate_eq1)
+ also have "... = (- 1) * (f x)" by (rule linearform_mult)
also have "... = - (f x)" by (simp!)
finally show ?thesis .
qed
@@ -58,14 +58,14 @@
text {* Every linear form yields @{text 0} for the @{text 0} vector. *}
lemma linearform_zero [intro?, simp]:
- "is_vectorspace V \<Longrightarrow> is_linearform V f \<Longrightarrow> f 0 = Numeral0"
+ "is_vectorspace V \<Longrightarrow> is_linearform V f \<Longrightarrow> f 0 = 0"
proof -
assume "is_vectorspace V" "is_linearform V f"
have "f 0 = f (0 - 0)" by (simp!)
also have "... = f 0 - f 0"
by (rule linearform_diff) (simp!)+
- also have "... = Numeral0" by simp
- finally show "f 0 = Numeral0" .
+ also have "... = 0" by simp
+ finally show "f 0 = 0" .
qed
end
--- a/src/HOL/Real/HahnBanach/NormedSpace.thy Thu Nov 01 21:12:13 2001 +0100
+++ b/src/HOL/Real/HahnBanach/NormedSpace.thy Fri Nov 02 17:55:24 2001 +0100
@@ -18,19 +18,19 @@
constdefs
is_seminorm :: "'a::{plus, minus, zero} set \<Rightarrow> ('a \<Rightarrow> real) \<Rightarrow> bool"
"is_seminorm V norm \<equiv> \<forall>x \<in> V. \<forall>y \<in> V. \<forall>a.
- Numeral0 \<le> norm x
+ 0 \<le> norm x
\<and> norm (a \<cdot> x) = \<bar>a\<bar> * norm x
\<and> norm (x + y) \<le> norm x + norm y"
lemma is_seminormI [intro]:
- "(\<And>x y a. x \<in> V \<Longrightarrow> y \<in> V \<Longrightarrow> Numeral0 \<le> norm x) \<Longrightarrow>
+ "(\<And>x y a. x \<in> V \<Longrightarrow> y \<in> V \<Longrightarrow> 0 \<le> norm x) \<Longrightarrow>
(\<And>x a. x \<in> V \<Longrightarrow> norm (a \<cdot> x) = \<bar>a\<bar> * norm x) \<Longrightarrow>
(\<And>x y. x \<in> V \<Longrightarrow> y \<in> V \<Longrightarrow> norm (x + y) \<le> norm x + norm y)
\<Longrightarrow> is_seminorm V norm"
by (unfold is_seminorm_def) auto
lemma seminorm_ge_zero [intro?]:
- "is_seminorm V norm \<Longrightarrow> x \<in> V \<Longrightarrow> Numeral0 \<le> norm x"
+ "is_seminorm V norm \<Longrightarrow> x \<in> V \<Longrightarrow> 0 \<le> norm x"
by (unfold is_seminorm_def) blast
lemma seminorm_abs_homogenous:
@@ -48,13 +48,13 @@
\<Longrightarrow> norm (x - y) \<le> norm x + norm y"
proof -
assume "is_seminorm V norm" "x \<in> V" "y \<in> V" "is_vectorspace V"
- have "norm (x - y) = norm (x + - Numeral1 \<cdot> y)"
+ have "norm (x - y) = norm (x + - 1 \<cdot> y)"
by (simp! add: diff_eq2 negate_eq2a)
- also have "... \<le> norm x + norm (- Numeral1 \<cdot> y)"
+ also have "... \<le> norm x + norm (- 1 \<cdot> y)"
by (simp! add: seminorm_subadditive)
- also have "norm (- Numeral1 \<cdot> y) = \<bar>- Numeral1\<bar> * norm y"
+ also have "norm (- 1 \<cdot> y) = \<bar>- 1\<bar> * norm y"
by (rule seminorm_abs_homogenous)
- also have "\<bar>- Numeral1\<bar> = (Numeral1::real)" by (rule abs_minus_one)
+ also have "\<bar>- 1\<bar> = (1::real)" by (rule abs_minus_one)
finally show "norm (x - y) \<le> norm x + norm y" by simp
qed
@@ -63,10 +63,10 @@
\<Longrightarrow> norm (- x) = norm x"
proof -
assume "is_seminorm V norm" "x \<in> V" "is_vectorspace V"
- have "norm (- x) = norm (- Numeral1 \<cdot> x)" by (simp! only: negate_eq1)
- also have "... = \<bar>- Numeral1\<bar> * norm x"
+ have "norm (- x) = norm (- 1 \<cdot> x)" by (simp! only: negate_eq1)
+ also have "... = \<bar>- 1\<bar> * norm x"
by (rule seminorm_abs_homogenous)
- also have "\<bar>- Numeral1\<bar> = (Numeral1::real)" by (rule abs_minus_one)
+ also have "\<bar>- 1\<bar> = (1::real)" by (rule abs_minus_one)
finally show "norm (- x) = norm x" by simp
qed
@@ -81,10 +81,10 @@
constdefs
is_norm :: "'a::{plus, minus, zero} set \<Rightarrow> ('a \<Rightarrow> real) \<Rightarrow> bool"
"is_norm V norm \<equiv> \<forall>x \<in> V. is_seminorm V norm
- \<and> (norm x = Numeral0) = (x = 0)"
+ \<and> (norm x = 0) = (x = 0)"
lemma is_normI [intro]:
- "\<forall>x \<in> V. is_seminorm V norm \<and> (norm x = Numeral0) = (x = 0)
+ "\<forall>x \<in> V. is_seminorm V norm \<and> (norm x = 0) = (x = 0)
\<Longrightarrow> is_norm V norm" by (simp only: is_norm_def)
lemma norm_is_seminorm [intro?]:
@@ -92,11 +92,11 @@
by (unfold is_norm_def) blast
lemma norm_zero_iff:
- "is_norm V norm \<Longrightarrow> x \<in> V \<Longrightarrow> (norm x = Numeral0) = (x = 0)"
+ "is_norm V norm \<Longrightarrow> x \<in> V \<Longrightarrow> (norm x = 0) = (x = 0)"
by (unfold is_norm_def) blast
lemma norm_ge_zero [intro?]:
- "is_norm V norm \<Longrightarrow> x \<in> V \<Longrightarrow> Numeral0 \<le> norm x"
+ "is_norm V norm \<Longrightarrow> x \<in> V \<Longrightarrow> 0 \<le> norm x"
by (unfold is_norm_def is_seminorm_def) blast
@@ -125,22 +125,22 @@
by (unfold is_normed_vectorspace_def) blast
lemma normed_vs_norm_ge_zero [intro?]:
- "is_normed_vectorspace V norm \<Longrightarrow> x \<in> V \<Longrightarrow> Numeral0 \<le> norm x"
+ "is_normed_vectorspace V norm \<Longrightarrow> x \<in> V \<Longrightarrow> 0 \<le> norm x"
by (unfold is_normed_vectorspace_def) (fast elim: norm_ge_zero)
lemma normed_vs_norm_gt_zero [intro?]:
- "is_normed_vectorspace V norm \<Longrightarrow> x \<in> V \<Longrightarrow> x \<noteq> 0 \<Longrightarrow> Numeral0 < norm x"
+ "is_normed_vectorspace V norm \<Longrightarrow> x \<in> V \<Longrightarrow> x \<noteq> 0 \<Longrightarrow> 0 < norm x"
proof (unfold is_normed_vectorspace_def, elim conjE)
assume "x \<in> V" "x \<noteq> 0" "is_vectorspace V" "is_norm V norm"
- have "Numeral0 \<le> norm x" ..
- also have "Numeral0 \<noteq> norm x"
+ have "0 \<le> norm x" ..
+ also have "0 \<noteq> norm x"
proof
- presume "norm x = Numeral0"
+ presume "norm x = 0"
also have "?this = (x = 0)" by (rule norm_zero_iff)
finally have "x = 0" .
thus "False" by contradiction
qed (rule sym)
- finally show "Numeral0 < norm x" .
+ finally show "0 < norm x" .
qed
lemma normed_vs_norm_abs_homogenous [intro?]:
@@ -170,14 +170,14 @@
show "is_seminorm F norm"
proof
fix x y a presume "x \<in> E"
- show "Numeral0 \<le> norm x" ..
+ show "0 \<le> norm x" ..
show "norm (a \<cdot> x) = \<bar>a\<bar> * norm x" ..
presume "y \<in> E"
show "norm (x + y) \<le> norm x + norm y" ..
qed (simp!)+
fix x assume "x \<in> F"
- show "(norm x = Numeral0) = (x = 0)"
+ show "(norm x = 0) = (x = 0)"
proof (rule norm_zero_iff)
show "is_norm E norm" ..
qed (simp!)
--- a/src/HOL/Real/HahnBanach/Subspace.thy Thu Nov 01 21:12:13 2001 +0100
+++ b/src/HOL/Real/HahnBanach/Subspace.thy Fri Nov 02 17:55:24 2001 +0100
@@ -87,7 +87,7 @@
show "0 \<in> U" ..
show "\<forall>x \<in> U. \<forall>a. a \<cdot> x \<in> U" by (simp!)
show "\<forall>x \<in> U. \<forall>y \<in> U. x + y \<in> U" by (simp!)
- show "\<forall>x \<in> U. - x = -Numeral1 \<cdot> x" by (simp! add: negate_eq1)
+ show "\<forall>x \<in> U. - x = - 1 \<cdot> x" by (simp! add: negate_eq1)
show "\<forall>x \<in> U. \<forall>y \<in> U. x - y = x + - y"
by (simp! add: diff_eq1)
qed (simp! add: vs_add_mult_distrib1 vs_add_mult_distrib2)+
@@ -154,7 +154,7 @@
lemma x_lin_x: "is_vectorspace V \<Longrightarrow> x \<in> V \<Longrightarrow> x \<in> lin x"
proof (unfold lin_def, intro CollectI exI conjI)
assume "is_vectorspace V" "x \<in> V"
- show "x = Numeral1 \<cdot> x" by (simp!)
+ show "x = 1 \<cdot> x" by (simp!)
qed simp
text {* Any linear closure is a subspace. *}
@@ -165,7 +165,7 @@
assume "is_vectorspace V" "x \<in> V"
show "0 \<in> lin x"
proof (unfold lin_def, intro CollectI exI conjI)
- show "0 = (Numeral0::real) \<cdot> x" by (simp!)
+ show "0 = (0::real) \<cdot> x" by (simp!)
qed simp
show "lin x \<subseteq> V"
@@ -383,9 +383,9 @@
fix a assume "x = a \<cdot> x'"
show ?thesis
proof cases
- assume "a = (Numeral0::real)" show ?thesis by (simp!)
+ assume "a = (0::real)" show ?thesis by (simp!)
next
- assume "a \<noteq> (Numeral0::real)"
+ assume "a \<noteq> (0::real)"
from h have "inverse a \<cdot> a \<cdot> x' \<in> H"
by (rule subspace_mult_closed) (simp!)
also have "inverse a \<cdot> a \<cdot> x' = x'" by (simp!)
@@ -425,15 +425,15 @@
lemma decomp_H'_H:
"is_vectorspace E \<Longrightarrow> is_subspace H E \<Longrightarrow> t \<in> H \<Longrightarrow> x' \<notin> H \<Longrightarrow> x' \<in> E
\<Longrightarrow> x' \<noteq> 0
- \<Longrightarrow> (SOME (y, a). t = y + a \<cdot> x' \<and> y \<in> H) = (t, (Numeral0::real))"
+ \<Longrightarrow> (SOME (y, a). t = y + a \<cdot> x' \<and> y \<in> H) = (t, (0::real))"
proof (rule, unfold split_tupled_all)
assume "is_vectorspace E" "is_subspace H E" "t \<in> H" "x' \<notin> H" "x' \<in> E"
"x' \<noteq> 0"
have h: "is_vectorspace H" ..
fix y a presume t1: "t = y + a \<cdot> x'" and "y \<in> H"
- have "y = t \<and> a = (Numeral0::real)"
+ have "y = t \<and> a = (0::real)"
by (rule decomp_H') (auto!)
- thus "(y, a) = (t, (Numeral0::real))" by (simp!)
+ thus "(y, a) = (t, (0::real))" by (simp!)
qed (simp_all!)
text {*
--- a/src/HOL/Real/HahnBanach/VectorSpace.thy Thu Nov 01 21:12:13 2001 +0100
+++ b/src/HOL/Real/HahnBanach/VectorSpace.thy Fri Nov 02 17:55:24 2001 +0100
@@ -31,7 +31,7 @@
associative and commutative; @{text "- x"} is the inverse of @{text
x} w.~r.~t.~addition and @{text 0} is the neutral element of
addition. Addition and multiplication are distributive; scalar
- multiplication is associative and the real number @{text "Numeral1"} is
+ multiplication is associative and the real number @{text "1"} is
the neutral element of scalar multiplication.
*}
@@ -48,8 +48,8 @@
\<and> a \<cdot> (x + y) = a \<cdot> x + a \<cdot> y
\<and> (a + b) \<cdot> x = a \<cdot> x + b \<cdot> x
\<and> (a * b) \<cdot> x = a \<cdot> b \<cdot> x
- \<and> Numeral1 \<cdot> x = x
- \<and> - x = (- Numeral1) \<cdot> x
+ \<and> 1 \<cdot> x = x
+ \<and> - x = (- 1) \<cdot> x
\<and> x - y = x + - y)"
@@ -66,15 +66,15 @@
\<forall>x \<in> V. \<forall>y \<in> V. \<forall>a. a \<cdot> (x + y) = a \<cdot> x + a \<cdot> y \<Longrightarrow>
\<forall>x \<in> V. \<forall>a b. (a + b) \<cdot> x = a \<cdot> x + b \<cdot> x \<Longrightarrow>
\<forall>x \<in> V. \<forall>a b. (a * b) \<cdot> x = a \<cdot> b \<cdot> x \<Longrightarrow>
- \<forall>x \<in> V. Numeral1 \<cdot> x = x \<Longrightarrow>
- \<forall>x \<in> V. - x = (- Numeral1) \<cdot> x \<Longrightarrow>
+ \<forall>x \<in> V. 1 \<cdot> x = x \<Longrightarrow>
+ \<forall>x \<in> V. - x = (- 1) \<cdot> x \<Longrightarrow>
\<forall>x \<in> V. \<forall>y \<in> V. x - y = x + - y \<Longrightarrow> is_vectorspace V"
by (unfold is_vectorspace_def) auto
text {* \medskip The corresponding destruction rules are: *}
lemma negate_eq1:
- "is_vectorspace V \<Longrightarrow> x \<in> V \<Longrightarrow> - x = (- Numeral1) \<cdot> x"
+ "is_vectorspace V \<Longrightarrow> x \<in> V \<Longrightarrow> - x = (- 1) \<cdot> x"
by (unfold is_vectorspace_def) simp
lemma diff_eq1:
@@ -82,7 +82,7 @@
by (unfold is_vectorspace_def) simp
lemma negate_eq2:
- "is_vectorspace V \<Longrightarrow> x \<in> V \<Longrightarrow> (- Numeral1) \<cdot> x = - x"
+ "is_vectorspace V \<Longrightarrow> x \<in> V \<Longrightarrow> (- 1) \<cdot> x = - x"
by (unfold is_vectorspace_def) simp
lemma negate_eq2a:
@@ -184,7 +184,7 @@
by (simp only: vs_mult_assoc)
lemma vs_mult_1 [simp]:
- "is_vectorspace V \<Longrightarrow> x \<in> V \<Longrightarrow> Numeral1 \<cdot> x = x"
+ "is_vectorspace V \<Longrightarrow> x \<in> V \<Longrightarrow> 1 \<cdot> x = x"
by (unfold is_vectorspace_def) simp
lemma vs_diff_mult_distrib1:
@@ -212,14 +212,14 @@
text {* \medskip Further derived laws: *}
lemma vs_mult_zero_left [simp]:
- "is_vectorspace V \<Longrightarrow> x \<in> V \<Longrightarrow> Numeral0 \<cdot> x = 0"
+ "is_vectorspace V \<Longrightarrow> x \<in> V \<Longrightarrow> 0 \<cdot> x = 0"
proof -
assume "is_vectorspace V" "x \<in> V"
- have "Numeral0 \<cdot> x = (Numeral1 - Numeral1) \<cdot> x" by simp
- also have "... = (Numeral1 + - Numeral1) \<cdot> x" by simp
- also have "... = Numeral1 \<cdot> x + (- Numeral1) \<cdot> x"
+ have "0 \<cdot> x = (1 - 1) \<cdot> x" by simp
+ also have "... = (1 + - 1) \<cdot> x" by simp
+ also have "... = 1 \<cdot> x + (- 1) \<cdot> x"
by (rule vs_add_mult_distrib2)
- also have "... = x + (- Numeral1) \<cdot> x" by (simp!)
+ also have "... = x + (- 1) \<cdot> x" by (simp!)
also have "... = x + - x" by (simp! add: negate_eq2a)
also have "... = x - x" by (simp! add: diff_eq2)
also have "... = 0" by (simp!)
@@ -349,12 +349,12 @@
qed
lemma vs_mult_left_cancel:
- "is_vectorspace V \<Longrightarrow> x \<in> V \<Longrightarrow> y \<in> V \<Longrightarrow> a \<noteq> Numeral0 \<Longrightarrow>
+ "is_vectorspace V \<Longrightarrow> x \<in> V \<Longrightarrow> y \<in> V \<Longrightarrow> a \<noteq> 0 \<Longrightarrow>
(a \<cdot> x = a \<cdot> y) = (x = y)"
(concl is "?L = ?R")
proof
- assume "is_vectorspace V" "x \<in> V" "y \<in> V" "a \<noteq> Numeral0"
- have "x = Numeral1 \<cdot> x" by (simp!)
+ assume "is_vectorspace V" "x \<in> V" "y \<in> V" "a \<noteq> 0"
+ have "x = 1 \<cdot> x" by (simp!)
also have "... = (inverse a * a) \<cdot> x" by (simp!)
also have "... = inverse a \<cdot> (a \<cdot> x)"
by (simp! only: vs_mult_assoc)
--- a/src/HOL/Real/PNat.ML Thu Nov 01 21:12:13 2001 +0100
+++ b/src/HOL/Real/PNat.ML Fri Nov 02 17:55:24 2001 +0100
@@ -38,7 +38,7 @@
qed "PNat_induct";
val prems = Goalw [pnat_one_def,pnat_Suc_def]
- "[| P(1p); \
+ "[| P(1); \
\ !!n. P(n) ==> P(pSuc n) |] ==> P(n)";
by (rtac (Rep_pnat_inverse RS subst) 1);
by (rtac (Rep_pnat RS PNat_induct) 1);
@@ -51,8 +51,8 @@
induct_thm_tac pnat_induct a i THEN rename_last_tac a [""] (i+1);
val prems = Goal
- "[| !!x. P x 1p; \
-\ !!y. P 1p (pSuc y); \
+ "[| !!x. P x 1; \
+\ !!y. P 1 (pSuc y); \
\ !!x y. [| P x y |] ==> P (pSuc x) (pSuc y) \
\ |] ==> P m n";
by (res_inst_tac [("x","m")] spec 1);
@@ -64,8 +64,8 @@
(*Case analysis on the natural numbers*)
val prems = Goal
- "[| n=1p ==> P; !!x. n = pSuc(x) ==> P |] ==> P";
-by (subgoal_tac "n=1p | (EX x. n = pSuc(x))" 1);
+ "[| n=1 ==> P; !!x. n = pSuc(x) ==> P |] ==> P";
+by (subgoal_tac "n=1 | (EX x. n = pSuc(x))" 1);
by (fast_tac (claset() addSEs prems) 1);
by (pnat_ind_tac "n" 1);
by (rtac (refl RS disjI1) 1);
@@ -131,7 +131,7 @@
(*** Distinctness of constructors ***)
-Goalw [pnat_one_def,pnat_Suc_def] "pSuc(m) ~= 1p";
+Goalw [pnat_one_def,pnat_Suc_def] "pSuc(m) ~= 1";
by (rtac (inj_on_Abs_pnat RS inj_on_contraD) 1);
by (rtac (Rep_pnat_gt_zero RS Suc_mono RS less_not_refl2) 1);
by (REPEAT (resolve_tac [Rep_pnat RS pnat_Suc_RepI, one_RepI] 1));
@@ -169,12 +169,12 @@
bind_thm ("pSuc_n_not_n", n_not_pSuc_n RS not_sym);
-Goal "n ~= 1p ==> EX m. n = pSuc m";
+Goal "n ~= 1 ==> EX m. n = pSuc m";
by (rtac pnatE 1);
by (REPEAT (Blast_tac 1));
-qed "not1p_implies_pSuc";
+qed "not1_implies_pSuc";
-Goal "pSuc m = m + 1p";
+Goal "pSuc m = m + 1";
by (auto_tac (claset(),simpset() addsimps [pnat_Suc_def,
pnat_one_def,Abs_pnat_inverse,pnat_add_def]));
qed "pSuc_is_plus_one";
@@ -485,12 +485,12 @@
(** embedding of naturals in positive naturals **)
(* pnat_one_eq! *)
-Goalw [pnat_of_nat_def,pnat_one_def]"1p = pnat_of_nat 0";
+Goalw [pnat_of_nat_def,pnat_one_def]"1 = pnat_of_nat 0";
by (Full_simp_tac 1);
qed "pnat_one_iff";
Goalw [pnat_of_nat_def,pnat_one_def,
- pnat_add_def] "1p + 1p = pnat_of_nat 1";
+ pnat_add_def] "1 + 1 = pnat_of_nat 1";
by (res_inst_tac [("f","Abs_pnat")] arg_cong 1);
by (auto_tac (claset() addIs [(gt_0_mem_pnat RS Abs_pnat_inverse RS ssubst)],
simpset()));
@@ -514,7 +514,7 @@
(* this worked with one call to auto_tac before! *)
Goalw [pnat_add_def,pnat_of_nat_def,pnat_one_def]
"pnat_of_nat n1 + pnat_of_nat n2 = \
-\ pnat_of_nat (n1 + n2) + 1p";
+\ pnat_of_nat (n1 + n2) + 1";
by (res_inst_tac [("f","Abs_pnat")] arg_cong 1);
by (rtac (gt_0_mem_pnat RS Abs_pnat_inverse RS ssubst) 1);
by (rtac (gt_0_mem_pnat RS Abs_pnat_inverse RS ssubst) 2);
--- a/src/HOL/Real/PNat.thy Thu Nov 01 21:12:13 2001 +0100
+++ b/src/HOL/Real/PNat.thy Fri Nov 02 17:55:24 2001 +0100
@@ -12,12 +12,11 @@
pnat = "lfp(%X. {Suc 0} Un Suc`X)" (lfp_def)
instance
- pnat :: {ord, plus, times}
+ pnat :: {ord, one, plus, times}
consts
pSuc :: pnat => pnat
- "1p" :: pnat ("1p")
constdefs
@@ -27,7 +26,7 @@
defs
pnat_one_def
- "1p == Abs_pnat(Suc 0)"
+ "1 == Abs_pnat(Suc 0)"
pnat_Suc_def
"pSuc == (%n. Abs_pnat(Suc(Rep_pnat(n))))"
--- a/src/HOL/Real/PRat.ML Thu Nov 01 21:12:13 2001 +0100
+++ b/src/HOL/Real/PRat.ML Fri Nov 02 17:55:24 2001 +0100
@@ -524,7 +524,7 @@
qed "prat_qinv_gt_1";
Goalw [pnat_one_def]
- "q1 < prat_of_pnat 1p ==> prat_of_pnat 1p < qinv(q1)";
+ "q1 < prat_of_pnat 1 ==> prat_of_pnat 1 < qinv(q1)";
by (etac prat_qinv_gt_1 1);
qed "prat_qinv_is_gt_1";
@@ -562,7 +562,7 @@
by (simp_tac (simpset() addsimps [prat_self_less_add_right]) 1);
qed "prat_self_less_add_left";
-Goalw [prat_less_def] "prat_of_pnat 1p < y ==> (x::prat) < x * y";
+Goalw [prat_less_def] "prat_of_pnat 1 < y ==> (x::prat) < x * y";
by (auto_tac (claset(),simpset() addsimps [pnat_one_def,
prat_add_mult_distrib2]));
qed "prat_self_less_mult_right";
--- a/src/HOL/Real/PReal.ML Thu Nov 01 21:12:13 2001 +0100
+++ b/src/HOL/Real/PReal.ML Fri Nov 02 17:55:24 2001 +0100
@@ -7,8 +7,6 @@
[Gleason- p. 121] provides some of the definitions.
*)
-claset_ref() := claset() delWrapper "bspec";
-
Goal "inj_on Abs_preal preal";
by (rtac inj_on_inverseI 1);
by (etac Abs_preal_inverse 1);
@@ -627,7 +625,7 @@
(******)
(*** FIXME: long! ***)
-Goal "prat_of_pnat 1p < x ==> EX r: Rep_preal(A). r*x ~: Rep_preal(A)";
+Goal "prat_of_pnat 1 < x ==> EX r: Rep_preal(A). r*x ~: Rep_preal(A)";
by (res_inst_tac [("X1","A")] (mem_Rep_preal_Ex RS exE) 1);
by (res_inst_tac [("Q","xa*x : Rep_preal(A)")] (excluded_middle RS disjE) 1);
by (Fast_tac 1);
--- a/src/HOL/Real/PReal.thy Thu Nov 01 21:12:13 2001 +0100
+++ b/src/HOL/Real/PReal.thy Fri Nov 02 17:55:24 2001 +0100
@@ -28,16 +28,16 @@
defs
preal_add_def
- "R + S == Abs_preal({w. ? x: Rep_preal(R). ? y: Rep_preal(S). w = x + y})"
+ "R + S == Abs_preal({w. ? x: Rep_preal(R). ? y: Rep_preal(S). w = x + y})"
preal_mult_def
- "R * S == Abs_preal({w. ? x: Rep_preal(R). ? y: Rep_preal(S). w = x * y})"
+ "R * S == Abs_preal({w. ? x: Rep_preal(R). ? y: Rep_preal(S). w = x * y})"
preal_less_def
- "R < (S::preal) == Rep_preal(R) < Rep_preal(S)"
+ "R < (S::preal) == Rep_preal(R) < Rep_preal(S)"
preal_le_def
- "R <= (S::preal) == Rep_preal(R) <= Rep_preal(S)"
+ "R <= (S::preal) == Rep_preal(R) <= Rep_preal(S)"
end
--- a/src/HOL/Real/RComplete.ML Thu Nov 01 21:12:13 2001 +0100
+++ b/src/HOL/Real/RComplete.ML Fri Nov 02 17:55:24 2001 +0100
@@ -6,8 +6,6 @@
Completeness theorems for positive reals and reals.
*)
-claset_ref() := claset() delWrapper "bspec";
-
Goal "x/2 + x/2 = (x::real)";
by (Simp_tac 1);
qed "real_sum_of_halves";
@@ -18,29 +16,25 @@
previously in Real.ML.
---------------------------------------------------------*)
(*a few lemmas*)
-Goal "ALL x:P. Numeral0 < x ==> \
+Goal "ALL x:P. 0 < x ==> \
\ ((EX x:P. y < x) = (EX X. real_of_preal X : P & \
\ y < real_of_preal X))";
by (blast_tac (claset() addSDs [bspec,
- rename_numerals real_gt_zero_preal_Ex RS iffD1]) 1);
+ real_gt_zero_preal_Ex RS iffD1]) 1);
qed "real_sup_lemma1";
-Goal "[| ALL x:P. Numeral0 < x; EX x. x: P; EX y. ALL x: P. x < y |] \
+Goal "[| ALL x:P. 0 < x; a: P; ALL x: P. x < y |] \
\ ==> (EX X. X: {w. real_of_preal w : P}) & \
\ (EX Y. ALL X: {w. real_of_preal w : P}. X < Y)";
by (rtac conjI 1);
by (blast_tac (claset() addDs [bspec,
- rename_numerals real_gt_zero_preal_Ex RS iffD1]) 1);
+ real_gt_zero_preal_Ex RS iffD1]) 1);
by Auto_tac;
by (dtac bspec 1 THEN assume_tac 1);
by (ftac bspec 1 THEN assume_tac 1);
-by (dtac order_less_trans 1 THEN assume_tac 1);
-by (dtac ((rename_numerals real_gt_zero_preal_Ex) RS iffD1) 1
- THEN etac exE 1);
-by (res_inst_tac [("x","ya")] exI 1);
-by Auto_tac;
-by (dres_inst_tac [("x","real_of_preal X")] bspec 1 THEN assume_tac 1);
-by (etac real_of_preal_lessD 1);
+by (dtac order_less_trans 1 THEN assume_tac 1);
+by (dtac (real_gt_zero_preal_Ex RS iffD1) 1);
+by (Force_tac 1);
qed "real_sup_lemma2";
(*-------------------------------------------------------------
@@ -49,33 +43,33 @@
(**
Supremum property for the set of positive reals
- FIXME: long proof - should be improved - need
- only have one case split
+ FIXME: long proof - should be improved
**)
-Goal "[| ALL x:P. (Numeral0::real) < x; EX x. x: P; EX y. ALL x: P. x < y |] \
-\ ==> (EX S. ALL y. (EX x: P. y < x) = (y < S))";
+(*Let P be a non-empty set of positive reals, with an upper bound y.
+ Then P has a least upper bound (written S).*)
+Goal "[| ALL x:P. (0::real) < x; EX x. x: P; EX y. ALL x: P. x<y |] \
+\ ==> (EX S. ALL y. (EX x: P. y < x) = (y < S))";
by (res_inst_tac
[("x","real_of_preal (psup({w. real_of_preal w : P}))")] exI 1);
-by Auto_tac;
-by (ftac real_sup_lemma2 1 THEN Auto_tac);
-by (case_tac "Numeral0 < ya" 1);
-by (dtac ((rename_numerals real_gt_zero_preal_Ex) RS iffD1) 1);
-by (dtac (rename_numerals real_less_all_real2) 2);
+by (Clarify_tac 1);
+by (case_tac "0 < ya" 1);
+by Auto_tac;
+by (ftac real_sup_lemma2 1 THEN REPEAT (assume_tac 1));
+by (dtac (real_gt_zero_preal_Ex RS iffD1) 1);
+by (dtac (real_less_all_real2) 3);
by Auto_tac;
by (rtac (preal_complete RS spec RS iffD1) 1);
by Auto_tac;
by (ftac real_gt_preal_preal_Ex 1);
-by Auto_tac;
+by (Force_tac 1);
(* second part *)
by (rtac (real_sup_lemma1 RS iffD2) 1 THEN assume_tac 1);
-by (case_tac "Numeral0 < ya" 1);
-by (auto_tac (claset() addSDs (map rename_numerals
- [real_less_all_real2,
- real_gt_zero_preal_Ex RS iffD1]),
+by (auto_tac (claset() addSDs [real_less_all_real2,
+ real_gt_zero_preal_Ex RS iffD1],
simpset()));
-by (ftac real_sup_lemma2 2 THEN Auto_tac);
-by (ftac real_sup_lemma2 1 THEN Auto_tac);
+by (ftac real_sup_lemma2 2 THEN REPEAT (assume_tac 1));
+by (ftac real_sup_lemma2 1 THEN REPEAT (assume_tac 1));
by (rtac (preal_complete RS spec RS iffD2 RS bexE) 1);
by (Blast_tac 3);
by (ALLGOALS(Blast_tac));
@@ -100,7 +94,7 @@
Completeness theorem for the positive reals(again)
----------------------------------------------------------------*)
-Goal "[| ALL x: S. Numeral0 < x; \
+Goal "[| ALL x: S. 0 < x; \
\ EX x. x: S; \
\ EX u. isUb (UNIV::real set) S u \
\ |] ==> EX t. isLub (UNIV::real set) S t";
@@ -109,10 +103,10 @@
by (auto_tac (claset(), simpset() addsimps
[isLub_def,leastP_def,isUb_def]));
by (auto_tac (claset() addSIs [setleI,setgeI]
- addSDs [(rename_numerals real_gt_zero_preal_Ex) RS iffD1],
+ addSDs [(real_gt_zero_preal_Ex) RS iffD1],
simpset()));
by (forw_inst_tac [("x","y")] bspec 1 THEN assume_tac 1);
-by (dtac ((rename_numerals real_gt_zero_preal_Ex) RS iffD1) 1);
+by (dtac ((real_gt_zero_preal_Ex) RS iffD1) 1);
by (auto_tac (claset(), simpset() addsimps [real_of_preal_le_iff]));
by (rtac preal_psup_leI2a 1);
by (forw_inst_tac [("y","real_of_preal ya")] setleD 1 THEN assume_tac 1);
@@ -122,7 +116,7 @@
by (blast_tac (claset() addSDs [setleD] addIs [real_of_preal_le_iff RS iffD1]) 1);
by (forw_inst_tac [("x","x")] bspec 1 THEN assume_tac 1);
by (ftac isUbD2 1);
-by (dtac ((rename_numerals real_gt_zero_preal_Ex) RS iffD1) 1);
+by (dtac ((real_gt_zero_preal_Ex) RS iffD1) 1);
by (auto_tac (claset() addSDs [isUbD, real_ge_preal_preal_Ex],
simpset() addsimps [real_of_preal_le_iff]));
by (blast_tac (claset() addSDs [setleD] addSIs [psup_le_ub1]
@@ -133,16 +127,17 @@
(*-------------------------------
Lemmas
-------------------------------*)
-Goal "ALL y : {z. EX x: P. z = x + (-xa) + Numeral1} Int {x. Numeral0 < x}. Numeral0 < y";
+Goal "ALL y : {z. EX x: P. z = x + (-xa) + 1} Int {x. 0 < x}. 0 < y";
by Auto_tac;
qed "real_sup_lemma3";
Goal "(xa <= S + X + (-Z)) = (xa + (-X) + Z <= (S::real))";
-by (Auto_tac);
+by Auto_tac;
qed "lemma_le_swap2";
-Goal "[| (x::real) + (-X) + Numeral1 <= S; xa <= x |] ==> xa <= S + X + (-Numeral1)";
-by (Auto_tac);
+Goal "[| (x::real) + (-X) + 1 <= S; xa <= x |] ==> xa <= S + X + (- 1)";
+by (arith_tac 1);
+by Auto_tac;
qed "lemma_real_complete2b";
(*----------------------------------------------------------
@@ -151,19 +146,19 @@
Goal "[| EX X. X: S; EX Y. isUb (UNIV::real set) S Y |] \
\ ==> EX t. isLub (UNIV :: real set) S t";
by (Step_tac 1);
-by (subgoal_tac "EX u. u: {z. EX x: S. z = x + (-X) + Numeral1} \
-\ Int {x. Numeral0 < x}" 1);
-by (subgoal_tac "isUb (UNIV::real set) ({z. EX x: S. z = x + (-X) + Numeral1} \
-\ Int {x. Numeral0 < x}) (Y + (-X) + Numeral1)" 1);
+by (subgoal_tac "EX u. u: {z. EX x: S. z = x + (-X) + 1} \
+\ Int {x. 0 < x}" 1);
+by (subgoal_tac "isUb (UNIV::real set) ({z. EX x: S. z = x + (-X) + 1} \
+\ Int {x. 0 < x}) (Y + (-X) + 1)" 1);
by (cut_inst_tac [("P","S"),("xa","X")] real_sup_lemma3 1);
by (EVERY1[forward_tac [exI RSN (3,posreals_complete)], Blast_tac, Blast_tac,
Step_tac]);
-by (res_inst_tac [("x","t + X + (-Numeral1)")] exI 1);
+by (res_inst_tac [("x","t + X + (- 1)")] exI 1);
by (rtac isLubI2 1);
by (rtac setgeI 2 THEN Step_tac 2);
-by (subgoal_tac "isUb (UNIV:: real set) ({z. EX x: S. z = x + (-X) + Numeral1} \
-\ Int {x. Numeral0 < x}) (y + (-X) + Numeral1)" 2);
-by (dres_inst_tac [("y","(y + (- X) + Numeral1)")] isLub_le_isUb 2
+by (subgoal_tac "isUb (UNIV:: real set) ({z. EX x: S. z = x + (-X) + 1} \
+\ Int {x. 0 < x}) (y + (-X) + 1)" 2);
+by (dres_inst_tac [("y","(y + (- X) + 1)")] isLub_le_isUb 2
THEN assume_tac 2);
by (full_simp_tac
(simpset() addsimps [real_diff_def, real_diff_le_eq RS sym] @
@@ -194,15 +189,15 @@
Related: Archimedean property of reals
----------------------------------------------------------------*)
-Goal "Numeral0 < real (Suc n)";
+Goal "0 < real (Suc n)";
by (res_inst_tac [("y","real n")] order_le_less_trans 1);
-by (rtac (rename_numerals real_of_nat_ge_zero) 1);
+by (rtac (real_of_nat_ge_zero) 1);
by (simp_tac (simpset() addsimps [real_of_nat_Suc]) 1);
qed "real_of_nat_Suc_gt_zero";
-Goal "Numeral0 < x ==> EX n. inverse (real(Suc n)) < x";
+Goal "0 < x ==> EX n. inverse (real(Suc n)) < x";
by (rtac ccontr 1);
-by (subgoal_tac "ALL n. x * real (Suc n) <= Numeral1" 1);
+by (subgoal_tac "ALL n. x * real (Suc n) <= 1" 1);
by (asm_full_simp_tac
(simpset() addsimps [linorder_not_less, real_inverse_eq_divide]) 2);
by (Clarify_tac 2);
@@ -213,7 +208,7 @@
addsimps [real_of_nat_Suc_gt_zero RS real_not_refl2 RS not_sym,
real_mult_commute]) 2);
by (subgoal_tac "isUb (UNIV::real set) \
-\ {z. EX n. z = x*(real (Suc n))} Numeral1" 1);
+\ {z. EX n. z = x*(real (Suc n))} 1" 1);
by (subgoal_tac "EX X. X : {z. EX n. z = x*(real (Suc n))}" 1);
by (dtac reals_complete 1);
by (auto_tac (claset() addIs [isUbI,setleI],simpset()));
@@ -234,17 +229,16 @@
(*There must be other proofs, e.g. Suc of the largest integer in the
cut representing x*)
Goal "EX n. (x::real) < real (n::nat)";
-by (res_inst_tac [("R1.0","x"),("R2.0","Numeral0")] real_linear_less2 1);
+by (res_inst_tac [("R1.0","x"),("R2.0","0")] real_linear_less2 1);
by (res_inst_tac [("x","0")] exI 1);
by (res_inst_tac [("x","1")] exI 2);
by (auto_tac (claset() addEs [order_less_trans],
simpset() addsimps [real_of_nat_one]));
-by (ftac ((rename_numerals real_inverse_gt_zero) RS reals_Archimedean) 1);
+by (ftac (real_inverse_gt_0 RS reals_Archimedean) 1);
by (Step_tac 1 THEN res_inst_tac [("x","Suc n")] exI 1);
-by (forw_inst_tac [("y","inverse x")]
- (rename_numerals real_mult_less_mono1) 1);
+by (forw_inst_tac [("y","inverse x")] real_mult_less_mono1 1);
by Auto_tac;
-by (dres_inst_tac [("y","Numeral1"),("z","real (Suc n)")]
+by (dres_inst_tac [("y","1"),("z","real (Suc n)")]
(rotate_prems 1 real_mult_less_mono2) 1);
by (auto_tac (claset(),
simpset() addsimps [real_of_nat_Suc_gt_zero,
--- a/src/HOL/Real/RealAbs.ML Thu Nov 01 21:12:13 2001 +0100
+++ b/src/HOL/Real/RealAbs.ML Fri Nov 02 17:55:24 2001 +0100
@@ -15,15 +15,15 @@
by (simp_tac
(simpset () addsimps
bin_arith_simps@
- [minus_real_number_of, zero_eq_numeral_0, le_real_number_of_eq_not_less,
+ [minus_real_number_of, le_real_number_of_eq_not_less,
less_real_number_of, real_of_int_le_iff]) 1);
qed "abs_nat_number_of";
Addsimps [abs_nat_number_of];
Goalw [real_abs_def]
- "P(abs (x::real)) = ((Numeral0 <= x --> P x) & (x < Numeral0 --> P(-x)))";
-by(auto_tac (claset(), simpset() addsimps [zero_eq_numeral_0]));
+ "P(abs (x::real)) = ((0 <= x --> P x) & (x < 0 --> P(-x)))";
+by Auto_tac;
qed "abs_split";
@@ -32,36 +32,37 @@
(adapted version of previously proved theorems about abs)
----------------------------------------------------------------------------*)
-Goalw [real_abs_def] "abs (r::real) = (if Numeral0<=r then r else -r)";
+Goalw [real_abs_def] "abs (r::real) = (if 0<=r then r else -r)";
by Auto_tac;
qed "abs_iff";
-Goalw [real_abs_def] "abs Numeral0 = (Numeral0::real)";
+Goalw [real_abs_def] "abs 0 = (0::real)";
by Auto_tac;
qed "abs_zero";
Addsimps [abs_zero];
-Goalw [real_abs_def] "abs (Numeral0::real) = -Numeral0";
+Goalw [real_abs_def] "abs (1::real) = 1";
by (Simp_tac 1);
-qed "abs_minus_zero";
+qed "abs_one";
+Addsimps [abs_one];
-Goalw [real_abs_def] "(Numeral0::real)<=x ==> abs x = x";
+Goalw [real_abs_def] "(0::real)<=x ==> abs x = x";
by (Asm_simp_tac 1);
qed "abs_eqI1";
-Goalw [real_abs_def] "(Numeral0::real) < x ==> abs x = x";
+Goalw [real_abs_def] "(0::real) < x ==> abs x = x";
by (Asm_simp_tac 1);
qed "abs_eqI2";
-Goalw [real_abs_def,real_le_def] "x < (Numeral0::real) ==> abs x = -x";
+Goalw [real_abs_def,real_le_def] "x < (0::real) ==> abs x = -x";
by (Asm_simp_tac 1);
qed "abs_minus_eqI2";
-Goalw [real_abs_def] "x<=(Numeral0::real) ==> abs x = -x";
+Goalw [real_abs_def] "x<=(0::real) ==> abs x = -x";
by (Asm_simp_tac 1);
qed "abs_minus_eqI1";
-Goalw [real_abs_def] "(Numeral0::real)<= abs x";
+Goalw [real_abs_def] "(0::real)<= abs x";
by (Simp_tac 1);
qed "abs_ge_zero";
@@ -70,17 +71,17 @@
qed "abs_idempotent";
Addsimps [abs_idempotent];
-Goalw [real_abs_def] "(abs x = Numeral0) = (x=(Numeral0::real))";
+Goalw [real_abs_def] "(abs x = 0) = (x=(0::real))";
by (Full_simp_tac 1);
qed "abs_zero_iff";
AddIffs [abs_zero_iff];
Goalw [real_abs_def] "x<=abs (x::real)";
-by (simp_tac (simpset() addsimps [zero_eq_numeral_0]) 1);
+by (Simp_tac 1);
qed "abs_ge_self";
Goalw [real_abs_def] "-x<=abs (x::real)";
-by (simp_tac (simpset() addsimps [zero_eq_numeral_0]) 1);
+by (Simp_tac 1);
qed "abs_ge_minus_self";
Goalw [real_abs_def] "abs (x * y) = abs x * abs (y::real)";
@@ -91,8 +92,8 @@
Goalw [real_abs_def] "abs(inverse(x::real)) = inverse(abs(x))";
by (real_div_undefined_case_tac "x=0" 1);
by (auto_tac (claset(),
- simpset() addsimps [real_minus_inverse, real_le_less] @
- (map rename_numerals [INVERSE_ZERO, real_inverse_gt_zero])));
+ simpset() addsimps [real_minus_inverse, real_le_less,
+ INVERSE_ZERO, real_inverse_gt_0]));
qed "abs_inverse";
Goal "abs (x * inverse y) = (abs x) * inverse (abs (y::real))";
@@ -130,16 +131,16 @@
qed "abs_add_minus_less";
(* lemmas manipulating terms *)
-Goal "((Numeral0::real)*x < r)=(Numeral0 < r)";
+Goal "((0::real)*x < r)=(0 < r)";
by (Simp_tac 1);
qed "real_mult_0_less";
-Goal "[| (Numeral0::real) < y; x < r; y*r < t*s |] ==> y*x < t*s";
-by (blast_tac (claset() addSIs [rename_numerals real_mult_less_mono2]
+Goal "[| (0::real) < y; x < r; y*r < t*s |] ==> y*x < t*s";
+by (blast_tac (claset() addSIs [real_mult_less_mono2]
addIs [order_less_trans]) 1);
qed "real_mult_less_trans";
-Goal "[| (Numeral0::real)<=y; x < r; y*r < t*s; Numeral0 < t*s|] ==> y*x < t*s";
+Goal "[| (0::real)<=y; x < r; y*r < t*s; 0 < t*s|] ==> y*x < t*s";
by (dtac order_le_imp_less_or_eq 1);
by (fast_tac (HOL_cs addEs [real_mult_0_less RS iffD2,
real_mult_less_trans]) 1);
@@ -150,7 +151,7 @@
by (rtac real_mult_le_less_trans 1);
by (rtac abs_ge_zero 1);
by (assume_tac 1);
-by (rtac (rename_numerals real_mult_order) 2);
+by (rtac real_mult_order 2);
by (auto_tac (claset() addSIs [real_mult_less_mono1, abs_ge_zero]
addIs [order_le_less_trans],
simpset()));
@@ -161,11 +162,11 @@
simpset() addsimps [abs_mult RS sym]));
qed "abs_mult_less2";
-Goal "abs(x) < r ==> (Numeral0::real) < r";
+Goal "abs(x) < r ==> (0::real) < r";
by (blast_tac (claset() addSIs [order_le_less_trans,abs_ge_zero]) 1);
qed "abs_less_gt_zero";
-Goalw [real_abs_def] "abs (-Numeral1) = (Numeral1::real)";
+Goalw [real_abs_def] "abs (-1) = (1::real)";
by (Simp_tac 1);
qed "abs_minus_one";
Addsimps [abs_minus_one];
@@ -182,17 +183,16 @@
by Auto_tac;
qed "abs_le_interval_iff";
-Goalw [real_abs_def] "(Numeral0::real) < k ==> Numeral0 < k + abs(x)";
+Goalw [real_abs_def] "(0::real) < k ==> 0 < k + abs(x)";
by Auto_tac;
qed "abs_add_pos_gt_zero";
-Goalw [real_abs_def] "(Numeral0::real) < Numeral1 + abs(x)";
+Goalw [real_abs_def] "(0::real) < 1 + abs(x)";
by Auto_tac;
qed "abs_add_one_gt_zero";
Addsimps [abs_add_one_gt_zero];
-(* 05/2000 *)
-Goalw [real_abs_def] "~ abs x < (Numeral0::real)";
+Goalw [real_abs_def] "~ abs x < (0::real)";
by Auto_tac;
qed "abs_not_less_zero";
Addsimps [abs_not_less_zero];
@@ -202,12 +202,12 @@
simpset()));
qed "abs_circle";
-Goalw [real_abs_def] "(abs x <= (Numeral0::real)) = (x = Numeral0)";
+Goalw [real_abs_def] "(abs x <= (0::real)) = (x = 0)";
by Auto_tac;
qed "abs_le_zero_iff";
Addsimps [abs_le_zero_iff];
-Goal "((Numeral0::real) < abs x) = (x ~= 0)";
+Goal "((0::real) < abs x) = (x ~= 0)";
by (simp_tac (simpset() addsimps [real_abs_def]) 1);
by (arith_tac 1);
qed "real_0_less_abs_iff";
@@ -215,11 +215,11 @@
Goal "abs (real x) = real (x::nat)";
by (auto_tac (claset() addIs [abs_eqI1],
- simpset() addsimps [rename_numerals real_of_nat_ge_zero]));
+ simpset() addsimps [real_of_nat_ge_zero]));
qed "abs_real_of_nat_cancel";
Addsimps [abs_real_of_nat_cancel];
-Goal "~ abs(x) + (Numeral1::real) < x";
+Goal "~ abs(x) + (1::real) < x";
by (rtac real_leD 1);
by (auto_tac (claset() addIs [abs_ge_self RS order_trans], simpset()));
qed "abs_add_one_not_less_self";
@@ -232,21 +232,21 @@
simpset() addsimps [real_add_assoc]));
qed "abs_triangle_ineq_three";
-Goalw [real_abs_def] "abs(x - y) < y ==> (Numeral0::real) < y";
-by (case_tac "Numeral0 <= x - y" 1);
+Goalw [real_abs_def] "abs(x - y) < y ==> (0::real) < y";
+by (case_tac "0 <= x - y" 1);
by Auto_tac;
qed "abs_diff_less_imp_gt_zero";
-Goalw [real_abs_def] "abs(x - y) < x ==> (Numeral0::real) < x";
-by (case_tac "Numeral0 <= x - y" 1);
+Goalw [real_abs_def] "abs(x - y) < x ==> (0::real) < x";
+by (case_tac "0 <= x - y" 1);
by Auto_tac;
qed "abs_diff_less_imp_gt_zero2";
-Goal "abs(x - y) < y ==> (Numeral0::real) < x";
+Goal "abs(x - y) < y ==> (0::real) < x";
by (auto_tac (claset(),simpset() addsimps [abs_interval_iff]));
qed "abs_diff_less_imp_gt_zero3";
-Goal "abs(x - y) < -y ==> x < (Numeral0::real)";
+Goal "abs(x - y) < -y ==> x < (0::real)";
by (auto_tac (claset(),simpset() addsimps [abs_interval_iff]));
qed "abs_diff_less_imp_gt_zero4";
--- a/src/HOL/Real/RealAbs.thy Thu Nov 01 21:12:13 2001 +0100
+++ b/src/HOL/Real/RealAbs.thy Fri Nov 02 17:55:24 2001 +0100
@@ -9,6 +9,6 @@
defs
- real_abs_def "abs r == (if (Numeral0::real) <= r then r else -r)"
+ real_abs_def "abs r == (if (0::real) <= r then r else -r)"
end
--- a/src/HOL/Real/RealArith0.ML Thu Nov 01 21:12:13 2001 +0100
+++ b/src/HOL/Real/RealArith0.ML Fri Nov 02 17:55:24 2001 +0100
@@ -8,85 +8,88 @@
Also, common factor cancellation
*)
+Goal "x - - y = x + (y::real)";
+by (Simp_tac 1);
+qed "real_diff_minus_eq";
+Addsimps [real_diff_minus_eq];
+
(** Division and inverse **)
-Goal "Numeral0/x = (Numeral0::real)";
+Goal "0/x = (0::real)";
by (simp_tac (simpset() addsimps [real_divide_def]) 1);
qed "real_0_divide";
Addsimps [real_0_divide];
-Goal "((Numeral0::real) < inverse x) = (Numeral0 < x)";
-by (case_tac "x=Numeral0" 1);
-by (asm_simp_tac (HOL_ss addsimps [rename_numerals INVERSE_ZERO]) 1);
-by (auto_tac (claset() addDs [rename_numerals real_inverse_less_zero],
- simpset() addsimps [linorder_neq_iff,
- rename_numerals real_inverse_gt_zero]));
+Goal "((0::real) < inverse x) = (0 < x)";
+by (case_tac "x=0" 1);
+by (asm_simp_tac (HOL_ss addsimps [INVERSE_ZERO]) 1);
+by (auto_tac (claset() addDs [real_inverse_less_0],
+ simpset() addsimps [linorder_neq_iff, real_inverse_gt_0]));
qed "real_0_less_inverse_iff";
Addsimps [real_0_less_inverse_iff];
-Goal "(inverse x < (Numeral0::real)) = (x < Numeral0)";
-by (case_tac "x=Numeral0" 1);
-by (asm_simp_tac (HOL_ss addsimps [rename_numerals INVERSE_ZERO]) 1);
-by (auto_tac (claset() addDs [rename_numerals real_inverse_less_zero],
- simpset() addsimps [linorder_neq_iff,
- rename_numerals real_inverse_gt_zero]));
+Goal "(inverse x < (0::real)) = (x < 0)";
+by (case_tac "x=0" 1);
+by (asm_simp_tac (HOL_ss addsimps [INVERSE_ZERO]) 1);
+by (auto_tac (claset() addDs [real_inverse_less_0],
+ simpset() addsimps [linorder_neq_iff, real_inverse_gt_0]));
qed "real_inverse_less_0_iff";
Addsimps [real_inverse_less_0_iff];
-Goal "((Numeral0::real) <= inverse x) = (Numeral0 <= x)";
+Goal "((0::real) <= inverse x) = (0 <= x)";
by (simp_tac (simpset() addsimps [linorder_not_less RS sym]) 1);
qed "real_0_le_inverse_iff";
Addsimps [real_0_le_inverse_iff];
-Goal "(inverse x <= (Numeral0::real)) = (x <= Numeral0)";
+Goal "(inverse x <= (0::real)) = (x <= 0)";
by (simp_tac (simpset() addsimps [linorder_not_less RS sym]) 1);
qed "real_inverse_le_0_iff";
Addsimps [real_inverse_le_0_iff];
-Goalw [real_divide_def] "x/(Numeral0::real) = Numeral0";
-by (stac (rename_numerals INVERSE_ZERO) 1);
+Goalw [real_divide_def] "x/(0::real) = 0";
+by (stac INVERSE_ZERO 1);
by (Simp_tac 1);
qed "REAL_DIVIDE_ZERO";
-Goal "inverse (x::real) = Numeral1/x";
+Goal "inverse (x::real) = 1/x";
by (simp_tac (simpset() addsimps [real_divide_def]) 1);
qed "real_inverse_eq_divide";
-Goal "((Numeral0::real) < x/y) = (Numeral0 < x & Numeral0 < y | x < Numeral0 & y < Numeral0)";
+Goal "((0::real) < x/y) = (0 < x & 0 < y | x < 0 & y < 0)";
by (simp_tac (simpset() addsimps [real_divide_def, real_0_less_mult_iff]) 1);
qed "real_0_less_divide_iff";
Addsimps [inst "x" "number_of ?w" real_0_less_divide_iff];
-Goal "(x/y < (Numeral0::real)) = (Numeral0 < x & y < Numeral0 | x < Numeral0 & Numeral0 < y)";
+Goal "(x/y < (0::real)) = (0 < x & y < 0 | x < 0 & 0 < y)";
by (simp_tac (simpset() addsimps [real_divide_def, real_mult_less_0_iff]) 1);
qed "real_divide_less_0_iff";
Addsimps [inst "x" "number_of ?w" real_divide_less_0_iff];
-Goal "((Numeral0::real) <= x/y) = ((x <= Numeral0 | Numeral0 <= y) & (Numeral0 <= x | y <= Numeral0))";
+Goal "((0::real) <= x/y) = ((x <= 0 | 0 <= y) & (0 <= x | y <= 0))";
by (simp_tac (simpset() addsimps [real_divide_def, real_0_le_mult_iff]) 1);
by Auto_tac;
qed "real_0_le_divide_iff";
Addsimps [inst "x" "number_of ?w" real_0_le_divide_iff];
-Goal "(x/y <= (Numeral0::real)) = ((x <= Numeral0 | y <= Numeral0) & (Numeral0 <= x | Numeral0 <= y))";
+Goal "(x/y <= (0::real)) = ((x <= 0 | y <= 0) & (0 <= x | 0 <= y))";
by (simp_tac (simpset() addsimps [real_divide_def, real_mult_le_0_iff]) 1);
by Auto_tac;
qed "real_divide_le_0_iff";
Addsimps [inst "x" "number_of ?w" real_divide_le_0_iff];
-Goal "(inverse(x::real) = Numeral0) = (x = Numeral0)";
-by (auto_tac (claset(), simpset() addsimps [rename_numerals INVERSE_ZERO]));
+Goal "(inverse(x::real) = 0) = (x = 0)";
+by (auto_tac (claset(), simpset() addsimps [INVERSE_ZERO]));
by (rtac ccontr 1);
-by (blast_tac (claset() addDs [rename_numerals real_inverse_not_zero]) 1);
+by (blast_tac (claset() addDs [real_inverse_not_zero]) 1);
qed "real_inverse_zero_iff";
Addsimps [real_inverse_zero_iff];
-Goal "(x/y = Numeral0) = (x=Numeral0 | y=(Numeral0::real))";
+Goal "(x/y = 0) = (x=0 | y=(0::real))";
by (auto_tac (claset(), simpset() addsimps [real_divide_def]));
qed "real_divide_eq_0_iff";
Addsimps [real_divide_eq_0_iff];
-Goal "h ~= (Numeral0::real) ==> h/h = Numeral1";
+Goal "h ~= (0::real) ==> h/h = 1";
by (asm_simp_tac (simpset() addsimps [real_divide_def, real_mult_inv_left]) 1);
qed "real_divide_self_eq";
Addsimps [real_divide_self_eq];
@@ -128,7 +131,7 @@
by (ALLGOALS (asm_full_simp_tac (simpset() addsimps [real_mult_commute])));
qed "real_mult_le_mono2_neg";
-Goal "(m*k < n*k) = (((Numeral0::real) < k & m<n) | (k < Numeral0 & n<m))";
+Goal "(m*k < n*k) = (((0::real) < k & m<n) | (k < 0 & n<m))";
by (case_tac "k = (0::real)" 1);
by (auto_tac (claset(),
simpset() addsimps [linorder_neq_iff,
@@ -143,32 +146,32 @@
real_mult_le_mono1_neg]));
qed "real_mult_less_cancel2";
-Goal "(m*k <= n*k) = (((Numeral0::real) < k --> m<=n) & (k < Numeral0 --> n<=m))";
+Goal "(m*k <= n*k) = (((0::real) < k --> m<=n) & (k < 0 --> n<=m))";
by (simp_tac (simpset() addsimps [linorder_not_less RS sym,
real_mult_less_cancel2]) 1);
qed "real_mult_le_cancel2";
-Goal "(k*m < k*n) = (((Numeral0::real) < k & m<n) | (k < Numeral0 & n<m))";
+Goal "(k*m < k*n) = (((0::real) < k & m<n) | (k < 0 & n<m))";
by (simp_tac (simpset() addsimps [inst "z" "k" real_mult_commute,
real_mult_less_cancel2]) 1);
qed "real_mult_less_cancel1";
-Goal "!!k::real. (k*m <= k*n) = ((Numeral0 < k --> m<=n) & (k < Numeral0 --> n<=m))";
+Goal "!!k::real. (k*m <= k*n) = ((0 < k --> m<=n) & (k < 0 --> n<=m))";
by (simp_tac (simpset() addsimps [linorder_not_less RS sym,
real_mult_less_cancel1]) 1);
qed "real_mult_le_cancel1";
-Goal "!!k::real. (k*m = k*n) = (k = Numeral0 | m=n)";
+Goal "!!k::real. (k*m = k*n) = (k = 0 | m=n)";
by (case_tac "k=0" 1);
by (auto_tac (claset(), simpset() addsimps [real_mult_left_cancel]));
qed "real_mult_eq_cancel1";
-Goal "!!k::real. (m*k = n*k) = (k = Numeral0 | m=n)";
+Goal "!!k::real. (m*k = n*k) = (k = 0 | m=n)";
by (case_tac "k=0" 1);
by (auto_tac (claset(), simpset() addsimps [real_mult_right_cancel]));
qed "real_mult_eq_cancel2";
-Goal "!!k::real. k~=Numeral0 ==> (k*m) / (k*n) = (m/n)";
+Goal "!!k::real. k~=0 ==> (k*m) / (k*n) = (m/n)";
by (asm_simp_tac
(simpset() addsimps [real_divide_def, real_inverse_distrib]) 1);
by (subgoal_tac "k * m * (inverse k * inverse n) = \
@@ -178,7 +181,7 @@
qed "real_mult_div_cancel1";
(*For ExtractCommonTerm*)
-Goal "(k*m) / (k*n) = (if k = (Numeral0::real) then Numeral0 else m/n)";
+Goal "(k*m) / (k*n) = (if k = (0::real) then 0 else m/n)";
by (simp_tac (simpset() addsimps [real_mult_div_cancel1]) 1);
qed "real_mult_div_cancel_disj";
@@ -188,20 +191,17 @@
in
val rel_real_number_of = [eq_real_number_of, less_real_number_of,
- le_real_number_of_eq_not_less];
+ le_real_number_of_eq_not_less]
structure CancelNumeralFactorCommon =
struct
val mk_coeff = mk_coeff
val dest_coeff = dest_coeff 1
val trans_tac = trans_tac
- val norm_tac = ALLGOALS (simp_tac (HOL_ss addsimps mult_plus_1s))
+ val norm_tac =
+ ALLGOALS (simp_tac (HOL_ss addsimps real_minus_from_mult_simps @ mult_1s))
THEN ALLGOALS (simp_tac (HOL_ss addsimps bin_simps@real_mult_minus_simps))
- THEN ALLGOALS
- (simp_tac
- (HOL_ss addsimps [eq_real_number_of, mult_real_number_of,
- real_mult_number_of_left]@
- real_minus_from_mult_simps @ real_mult_ac))
+ THEN ALLGOALS (simp_tac (HOL_ss addsimps real_mult_ac))
val numeral_simp_tac =
ALLGOALS (simp_tac (HOL_ss addsimps rel_real_number_of@bin_simps))
val simplify_meta_eq = simplify_meta_eq
@@ -253,17 +253,17 @@
LessCancelNumeralFactor.proc),
("realle_cancel_numeral_factor",
prep_pats ["(l::real) * m <= n", "(l::real) <= m * n"],
- LeCancelNumeralFactor.proc)];
+ LeCancelNumeralFactor.proc)]
val real_cancel_numeral_factors_divide = prep_simproc
("realdiv_cancel_numeral_factor",
prep_pats ["((l::real) * m) / n", "(l::real) / (m * n)",
"((number_of v)::real) / (number_of w)"],
- DivCancelNumeralFactor.proc);
+ DivCancelNumeralFactor.proc)
val real_cancel_numeral_factors =
real_cancel_numeral_factors_relations @
- [real_cancel_numeral_factors_divide];
+ [real_cancel_numeral_factors_divide]
end;
@@ -276,7 +276,7 @@
set trace_simp;
fun test s = (Goal s; by (Simp_tac 1));
-test "Numeral0 <= (y::real) * -2";
+test "0 <= (y::real) * -2";
test "9*x = 12 * (y::real)";
test "(9*x) / (12 * (y::real)) = z";
test "9*x < 12 * (y::real)";
@@ -292,6 +292,7 @@
test "999*x < -396 * (y::real)";
test "999*x <= -396 * (y::real)";
+test "(- ((2::real) * x) <= 2 * y)";
test "-99*x = -81 * (y::real)";
test "(-99*x) / (-81 * (y::real)) = z";
test "-99*x <= -81 * (y::real)";
@@ -379,7 +380,7 @@
(*** Simplification of inequalities involving literal divisors ***)
-Goal "Numeral0<z ==> ((x::real) <= y/z) = (x*z <= y)";
+Goal "0<z ==> ((x::real) <= y/z) = (x*z <= y)";
by (subgoal_tac "(x*z <= y) = (x*z <= (y/z)*z)" 1);
by (asm_simp_tac (simpset() addsimps [real_divide_def, real_mult_assoc]) 2);
by (etac ssubst 1);
@@ -388,7 +389,7 @@
qed "pos_real_le_divide_eq";
Addsimps [inst "z" "number_of ?w" pos_real_le_divide_eq];
-Goal "z<Numeral0 ==> ((x::real) <= y/z) = (y <= x*z)";
+Goal "z<0 ==> ((x::real) <= y/z) = (y <= x*z)";
by (subgoal_tac "(y <= x*z) = ((y/z)*z <= x*z)" 1);
by (asm_simp_tac (simpset() addsimps [real_divide_def, real_mult_assoc]) 2);
by (etac ssubst 1);
@@ -397,7 +398,7 @@
qed "neg_real_le_divide_eq";
Addsimps [inst "z" "number_of ?w" neg_real_le_divide_eq];
-Goal "Numeral0<z ==> (y/z <= (x::real)) = (y <= x*z)";
+Goal "0<z ==> (y/z <= (x::real)) = (y <= x*z)";
by (subgoal_tac "(y <= x*z) = ((y/z)*z <= x*z)" 1);
by (asm_simp_tac (simpset() addsimps [real_divide_def, real_mult_assoc]) 2);
by (etac ssubst 1);
@@ -406,7 +407,7 @@
qed "pos_real_divide_le_eq";
Addsimps [inst "z" "number_of ?w" pos_real_divide_le_eq];
-Goal "z<Numeral0 ==> (y/z <= (x::real)) = (x*z <= y)";
+Goal "z<0 ==> (y/z <= (x::real)) = (x*z <= y)";
by (subgoal_tac "(x*z <= y) = (x*z <= (y/z)*z)" 1);
by (asm_simp_tac (simpset() addsimps [real_divide_def, real_mult_assoc]) 2);
by (etac ssubst 1);
@@ -415,7 +416,7 @@
qed "neg_real_divide_le_eq";
Addsimps [inst "z" "number_of ?w" neg_real_divide_le_eq];
-Goal "Numeral0<z ==> ((x::real) < y/z) = (x*z < y)";
+Goal "0<z ==> ((x::real) < y/z) = (x*z < y)";
by (subgoal_tac "(x*z < y) = (x*z < (y/z)*z)" 1);
by (asm_simp_tac (simpset() addsimps [real_divide_def, real_mult_assoc]) 2);
by (etac ssubst 1);
@@ -424,7 +425,7 @@
qed "pos_real_less_divide_eq";
Addsimps [inst "z" "number_of ?w" pos_real_less_divide_eq];
-Goal "z<Numeral0 ==> ((x::real) < y/z) = (y < x*z)";
+Goal "z<0 ==> ((x::real) < y/z) = (y < x*z)";
by (subgoal_tac "(y < x*z) = ((y/z)*z < x*z)" 1);
by (asm_simp_tac (simpset() addsimps [real_divide_def, real_mult_assoc]) 2);
by (etac ssubst 1);
@@ -433,7 +434,7 @@
qed "neg_real_less_divide_eq";
Addsimps [inst "z" "number_of ?w" neg_real_less_divide_eq];
-Goal "Numeral0<z ==> (y/z < (x::real)) = (y < x*z)";
+Goal "0<z ==> (y/z < (x::real)) = (y < x*z)";
by (subgoal_tac "(y < x*z) = ((y/z)*z < x*z)" 1);
by (asm_simp_tac (simpset() addsimps [real_divide_def, real_mult_assoc]) 2);
by (etac ssubst 1);
@@ -442,7 +443,7 @@
qed "pos_real_divide_less_eq";
Addsimps [inst "z" "number_of ?w" pos_real_divide_less_eq];
-Goal "z<Numeral0 ==> (y/z < (x::real)) = (x*z < y)";
+Goal "z<0 ==> (y/z < (x::real)) = (x*z < y)";
by (subgoal_tac "(x*z < y) = (x*z < (y/z)*z)" 1);
by (asm_simp_tac (simpset() addsimps [real_divide_def, real_mult_assoc]) 2);
by (etac ssubst 1);
@@ -451,7 +452,7 @@
qed "neg_real_divide_less_eq";
Addsimps [inst "z" "number_of ?w" neg_real_divide_less_eq];
-Goal "z~=Numeral0 ==> ((x::real) = y/z) = (x*z = y)";
+Goal "z~=0 ==> ((x::real) = y/z) = (x*z = y)";
by (subgoal_tac "(x*z = y) = (x*z = (y/z)*z)" 1);
by (asm_simp_tac (simpset() addsimps [real_divide_def, real_mult_assoc]) 2);
by (etac ssubst 1);
@@ -460,7 +461,7 @@
qed "real_eq_divide_eq";
Addsimps [inst "z" "number_of ?w" real_eq_divide_eq];
-Goal "z~=Numeral0 ==> (y/z = (x::real)) = (y = x*z)";
+Goal "z~=0 ==> (y/z = (x::real)) = (y = x*z)";
by (subgoal_tac "(y = x*z) = ((y/z)*z = x*z)" 1);
by (asm_simp_tac (simpset() addsimps [real_divide_def, real_mult_assoc]) 2);
by (etac ssubst 1);
@@ -469,38 +470,38 @@
qed "real_divide_eq_eq";
Addsimps [inst "z" "number_of ?w" real_divide_eq_eq];
-Goal "(m/k = n/k) = (k = Numeral0 | m = (n::real))";
-by (case_tac "k=Numeral0" 1);
+Goal "(m/k = n/k) = (k = 0 | m = (n::real))";
+by (case_tac "k=0" 1);
by (asm_simp_tac (simpset() addsimps [REAL_DIVIDE_ZERO]) 1);
by (asm_simp_tac (simpset() addsimps [real_divide_eq_eq, real_eq_divide_eq,
real_mult_eq_cancel2]) 1);
qed "real_divide_eq_cancel2";
-Goal "(k/m = k/n) = (k = Numeral0 | m = (n::real))";
-by (case_tac "m=Numeral0 | n = Numeral0" 1);
+Goal "(k/m = k/n) = (k = 0 | m = (n::real))";
+by (case_tac "m=0 | n = 0" 1);
by (auto_tac (claset(),
simpset() addsimps [REAL_DIVIDE_ZERO, real_divide_eq_eq,
real_eq_divide_eq, real_mult_eq_cancel1]));
qed "real_divide_eq_cancel1";
-(*Moved from RealOrd.ML to use Numeral0 *)
-Goal "[| Numeral0 < r; Numeral0 < x|] ==> (inverse x < inverse (r::real)) = (r < x)";
+(*Moved from RealOrd.ML to use 0 *)
+Goal "[| 0 < r; 0 < x|] ==> (inverse x < inverse (r::real)) = (r < x)";
by (auto_tac (claset() addIs [real_inverse_less_swap], simpset()));
by (res_inst_tac [("t","r")] (real_inverse_inverse RS subst) 1);
by (res_inst_tac [("t","x")] (real_inverse_inverse RS subst) 1);
by (auto_tac (claset() addIs [real_inverse_less_swap],
simpset() delsimps [real_inverse_inverse]
- addsimps [real_inverse_gt_zero]));
+ addsimps [real_inverse_gt_0]));
qed "real_inverse_less_iff";
-Goal "[| Numeral0 < r; Numeral0 < x|] ==> (inverse x <= inverse r) = (r <= (x::real))";
+Goal "[| 0 < r; 0 < x|] ==> (inverse x <= inverse r) = (r <= (x::real))";
by (asm_simp_tac (simpset() addsimps [linorder_not_less RS sym,
real_inverse_less_iff]) 1);
qed "real_inverse_le_iff";
(** Division by 1, -1 **)
-Goal "(x::real)/Numeral1 = x";
+Goal "(x::real)/1 = x";
by (simp_tac (simpset() addsimps [real_divide_def]) 1);
qed "real_divide_1";
Addsimps [real_divide_1];
@@ -510,30 +511,30 @@
qed "real_divide_minus1";
Addsimps [real_divide_minus1];
-Goal "-1/(x::real) = - (Numeral1/x)";
+Goal "-1/(x::real) = - (1/x)";
by (simp_tac (simpset() addsimps [real_divide_def, real_minus_inverse]) 1);
qed "real_minus1_divide";
Addsimps [real_minus1_divide];
-Goal "[| (Numeral0::real) < d1; Numeral0 < d2 |] ==> EX e. Numeral0 < e & e < d1 & e < d2";
+Goal "[| (0::real) < d1; 0 < d2 |] ==> EX e. 0 < e & e < d1 & e < d2";
by (res_inst_tac [("x","(min d1 d2)/2")] exI 1);
by (asm_simp_tac (simpset() addsimps [min_def]) 1);
qed "real_lbound_gt_zero";
Goal "(inverse x = inverse y) = (x = (y::real))";
-by (case_tac "x=Numeral0 | y=Numeral0" 1);
+by (case_tac "x=0 | y=0" 1);
by (auto_tac (claset(),
simpset() addsimps [real_inverse_eq_divide,
- rename_numerals DIVISION_BY_ZERO]));
+ DIVISION_BY_ZERO]));
by (dres_inst_tac [("f","%u. x*y*u")] arg_cong 1);
by (Asm_full_simp_tac 1);
qed "real_inverse_eq_iff";
Addsimps [real_inverse_eq_iff];
-Goal "(z/x = z/y) = (z = Numeral0 | x = (y::real))";
-by (case_tac "x=Numeral0 | y=Numeral0" 1);
+Goal "(z/x = z/y) = (z = 0 | x = (y::real))";
+by (case_tac "x=0 | y=0" 1);
by (auto_tac (claset(),
- simpset() addsimps [rename_numerals DIVISION_BY_ZERO]));
+ simpset() addsimps [DIVISION_BY_ZERO]));
by (dres_inst_tac [("f","%u. x*y*u")] arg_cong 1);
by Auto_tac;
qed "real_divide_eq_iff";
@@ -569,7 +570,7 @@
qed "real_minus_equation";
-Goal "(x + - a = (Numeral0::real)) = (x=a)";
+Goal "(x + - a = (0::real)) = (x=a)";
by (arith_tac 1);
qed "real_add_minus_iff";
Addsimps [real_add_minus_iff];
@@ -590,45 +591,50 @@
Addsimps (map (inst "y" "number_of ?v")
[real_minus_less, real_minus_le, real_minus_equation]);
+(*Equations and inequations involving 1*)
+Addsimps (map (simplify (simpset()) o inst "x" "1")
+ [real_less_minus, real_le_minus, real_equation_minus]);
+Addsimps (map (simplify (simpset()) o inst "y" "1")
+ [real_minus_less, real_minus_le, real_minus_equation]);
-(*** Simprules combining x+y and Numeral0 ***)
+(*** Simprules combining x+y and 0 ***)
-Goal "(x+y = (Numeral0::real)) = (y = -x)";
+Goal "(x+y = (0::real)) = (y = -x)";
by Auto_tac;
qed "real_add_eq_0_iff";
AddIffs [real_add_eq_0_iff];
-Goal "(x+y < (Numeral0::real)) = (y < -x)";
+Goal "(x+y < (0::real)) = (y < -x)";
by Auto_tac;
qed "real_add_less_0_iff";
AddIffs [real_add_less_0_iff];
-Goal "((Numeral0::real) < x+y) = (-x < y)";
+Goal "((0::real) < x+y) = (-x < y)";
by Auto_tac;
qed "real_0_less_add_iff";
AddIffs [real_0_less_add_iff];
-Goal "(x+y <= (Numeral0::real)) = (y <= -x)";
+Goal "(x+y <= (0::real)) = (y <= -x)";
by Auto_tac;
qed "real_add_le_0_iff";
AddIffs [real_add_le_0_iff];
-Goal "((Numeral0::real) <= x+y) = (-x <= y)";
+Goal "((0::real) <= x+y) = (-x <= y)";
by Auto_tac;
qed "real_0_le_add_iff";
AddIffs [real_0_le_add_iff];
-(** Simprules combining x-y and Numeral0; see also real_less_iff_diff_less_0, etc.,
+(** Simprules combining x-y and 0; see also real_less_iff_diff_less_0, etc.,
in RealBin
**)
-Goal "((Numeral0::real) < x-y) = (y < x)";
+Goal "((0::real) < x-y) = (y < x)";
by Auto_tac;
qed "real_0_less_diff_iff";
AddIffs [real_0_less_diff_iff];
-Goal "((Numeral0::real) <= x-y) = (y <= x)";
+Goal "((0::real) <= x-y) = (y <= x)";
by Auto_tac;
qed "real_0_le_diff_iff";
AddIffs [real_0_le_diff_iff];
@@ -660,7 +666,7 @@
qed "real_dense";
-(*Replaces "inverse #nn" by Numeral1/#nn *)
+(*Replaces "inverse #nn" by 1/#nn *)
Addsimps [inst "x" "number_of ?w" real_inverse_eq_divide];
--- a/src/HOL/Real/RealBin.ML Thu Nov 01 21:12:13 2001 +0100
+++ b/src/HOL/Real/RealBin.ML Fri Nov 02 17:55:24 2001 +0100
@@ -13,13 +13,13 @@
qed "real_number_of";
Addsimps [real_number_of];
-Goalw [real_number_of_def] "(0::real) = Numeral0";
+Goalw [real_number_of_def] "Numeral0 = (0::real)";
by (simp_tac (simpset() addsimps [real_of_int_zero RS sym]) 1);
-qed "zero_eq_numeral_0";
+qed "real_numeral_0_eq_0";
-Goalw [real_number_of_def] "(1::real) = Numeral1";
+Goalw [real_number_of_def] "Numeral1 = (1::real)";
by (simp_tac (simpset() addsimps [real_of_int_one RS sym]) 1);
-qed "one_eq_numeral_1";
+qed "real_numeral_1_eq_1";
(** Addition **)
@@ -52,21 +52,19 @@
Goal "(number_of v :: real) * number_of v' = number_of (bin_mult v v')";
by (simp_tac
- (HOL_ss addsimps [real_number_of_def,
- real_of_int_mult, number_of_mult]) 1);
+ (HOL_ss addsimps [real_number_of_def, real_of_int_mult,
+ number_of_mult]) 1);
qed "mult_real_number_of";
Addsimps [mult_real_number_of];
-Goal "(2::real) = Numeral1 + Numeral1";
-by (Simp_tac 1);
+Goal "(2::real) = 1 + 1";
+by (simp_tac (simpset() addsimps [real_numeral_1_eq_1 RS sym]) 1);
val lemma = result();
(*For specialist use: NOT as default simprules*)
Goal "2 * z = (z+z::real)";
-by (simp_tac (simpset ()
- addsimps [lemma, real_add_mult_distrib,
- one_eq_numeral_1 RS sym]) 1);
+by (simp_tac (simpset () addsimps [lemma, real_add_mult_distrib]) 1);
qed "real_mult_2";
Goal "z * 2 = (z+z::real)";
@@ -111,66 +109,49 @@
(*** New versions of existing theorems involving 0, 1 ***)
-Goal "- Numeral1 = (-1::real)";
-by (Simp_tac 1);
-qed "minus_numeral_one";
+Goal "- 1 = (-1::real)";
+by (simp_tac (simpset() addsimps [real_numeral_1_eq_1 RS sym]) 1);
+qed "real_minus_1_eq_m1";
+
+Goal "-1 * z = -(z::real)";
+by (simp_tac (simpset() addsimps [real_minus_1_eq_m1 RS sym,
+ real_mult_minus_1]) 1);
+qed "real_mult_minus1";
+
+Goal "z * -1 = -(z::real)";
+by (stac real_mult_commute 1 THEN rtac real_mult_minus1 1);
+qed "real_mult_minus1_right";
+
+Addsimps [real_mult_minus1, real_mult_minus1_right];
(*Maps 0 to Numeral0 and 1 to Numeral1 and -(Numeral1) to -1*)
val real_numeral_ss =
- HOL_ss addsimps [zero_eq_numeral_0, one_eq_numeral_1,
- minus_numeral_one];
+ HOL_ss addsimps [real_numeral_0_eq_0 RS sym, real_numeral_1_eq_1 RS sym,
+ real_minus_1_eq_m1];
fun rename_numerals th =
asm_full_simplify real_numeral_ss (Thm.transfer (the_context ()) th);
-(*Now insert some identities previously stated for 0 and 1*)
-
-(** RealDef & Real **)
-
-Addsimps (map rename_numerals
- [real_minus_zero, real_minus_zero_iff,
- real_add_zero_left, real_add_zero_right,
- real_diff_0, real_diff_0_right,
- real_mult_0_right, real_mult_0, real_mult_1_right, real_mult_1,
- real_mult_minus_1_right, real_mult_minus_1, real_inverse_1,
- real_minus_zero_less_iff]);
-
-AddIffs
- (map rename_numerals
- [real_mult_is_0, real_of_nat_zero_iff, real_le_square]);
-
-bind_thm ("real_0_less_mult_iff",
- rename_numerals real_zero_less_mult_iff);
-bind_thm ("real_0_le_mult_iff",
- rename_numerals real_zero_le_mult_iff);
-bind_thm ("real_mult_less_0_iff",
- rename_numerals real_mult_less_zero_iff);
-bind_thm ("real_mult_le_0_iff",
- rename_numerals real_mult_le_zero_iff);
-
-bind_thm ("real_inverse_less_0", rename_numerals real_inverse_less_zero);
-bind_thm ("real_inverse_gt_0", rename_numerals real_inverse_gt_zero);
-
-Addsimps [zero_eq_numeral_0,one_eq_numeral_1];
-
-
(** real from type "nat" **)
-Goal "(Numeral0 < real (n::nat)) = (0<n)";
+Goal "(0 < real (n::nat)) = (0<n)";
by (simp_tac (HOL_ss addsimps [real_of_nat_less_iff,
- rename_numerals real_of_nat_zero RS sym]) 1);
+ real_of_nat_zero RS sym]) 1);
qed "zero_less_real_of_nat_iff";
AddIffs [zero_less_real_of_nat_iff];
-Goal "(Numeral0 <= real (n::nat)) = (0<=n)";
+Goal "(0 <= real (n::nat)) = (0<=n)";
by (simp_tac (HOL_ss addsimps [real_of_nat_le_iff,
- rename_numerals real_of_nat_zero RS sym]) 1);
+ real_of_nat_zero RS sym]) 1);
qed "zero_le_real_of_nat_iff";
AddIffs [zero_le_real_of_nat_iff];
+(*Like the ones above, for "equals"*)
+AddIffs [real_of_nat_zero_iff];
+
(** Simplification of arithmetic when nested to the right **)
Goal "number_of v + (number_of w + z) = (number_of(bin_add v w) + z::real)";
@@ -182,12 +163,12 @@
qed "real_mult_number_of_left";
Goalw [real_diff_def]
- "number_of v + (number_of w - c) = number_of(bin_add v w) - (c::real)";
+ "number_of v + (number_of w - c) = number_of(bin_add v w) - (c::real)";
by (rtac real_add_number_of_left 1);
qed "real_add_number_of_diff1";
Goal "number_of v + (c - number_of w) = \
-\ number_of (bin_add v (bin_minus w)) + (c::real)";
+\ number_of (bin_add v (bin_minus w)) + (c::real)";
by (stac (diff_real_number_of RS sym) 1);
by Auto_tac;
qed "real_add_number_of_diff2";
@@ -198,12 +179,12 @@
(*"neg" is used in rewrite rules for binary comparisons*)
Goal "real (number_of v :: nat) = \
-\ (if neg (number_of v) then Numeral0 \
-\ else (number_of v :: real))";
+\ (if neg (number_of v) then 0 \
+\ else (number_of v :: real))";
by (simp_tac
(HOL_ss addsimps [nat_number_of_def, real_of_nat_real_of_int,
- real_of_nat_neg_int, real_number_of,
- zero_eq_numeral_0]) 1);
+ real_of_nat_neg_int, real_number_of,
+ real_numeral_0_eq_0 RS sym]) 1);
qed "real_of_nat_number_of";
Addsimps [real_of_nat_number_of];
@@ -212,15 +193,15 @@
(** Combining of literal coefficients in sums of products **)
-Goal "(x < y) = (x-y < (Numeral0::real))";
+Goal "(x < y) = (x-y < (0::real))";
by (simp_tac (simpset() addsimps [real_diff_less_eq]) 1);
qed "real_less_iff_diff_less_0";
-Goal "(x = y) = (x-y = (Numeral0::real))";
+Goal "(x = y) = (x-y = (0::real))";
by (simp_tac (simpset() addsimps [real_diff_eq_eq]) 1);
qed "real_eq_iff_diff_eq_0";
-Goal "(x <= y) = (x-y <= (Numeral0::real))";
+Goal "(x <= y) = (x-y <= (0::real))";
by (simp_tac (simpset() addsimps [real_diff_le_eq]) 1);
qed "real_le_iff_diff_le_0";
@@ -272,15 +253,24 @@
qed "real_le_add_iff2";
+Addsimps [real_numeral_0_eq_0, real_numeral_1_eq_1];
+
+
structure Real_Numeral_Simprocs =
struct
+(*Maps 0 to Numeral0 and 1 to Numeral1 so that arithmetic in simprocs
+ isn't complicated by the abstract 0 and 1.*)
+val numeral_syms = [real_numeral_0_eq_0 RS sym, real_numeral_1_eq_1 RS sym];
+
(*Utilities*)
fun mk_numeral n = HOLogic.number_of_const HOLogic.realT $ HOLogic.mk_bin n;
-(*Decodes a binary real constant*)
-fun dest_numeral (Const("Numeral.number_of", _) $ w) =
+(*Decodes a binary real constant, or 0, 1*)
+fun dest_numeral (Const("0", _)) = 0
+ | dest_numeral (Const("1", _)) = 1
+ | dest_numeral (Const("Numeral.number_of", _) $ w) =
(HOLogic.dest_binum w
handle TERM _ => raise TERM("Real_Numeral_Simprocs.dest_numeral:1", [w]))
| dest_numeral t = raise TERM("Real_Numeral_Simprocs.dest_numeral:2", [t]);
@@ -295,7 +285,7 @@
val uminus_const = Const ("uminus", HOLogic.realT --> HOLogic.realT);
-(*Thus mk_sum[t] yields t+Numeral0; longer sums don't have a trailing zero*)
+(*Thus mk_sum[t] yields t+0; longer sums don't have a trailing zero*)
fun mk_sum [] = zero
| mk_sum [t,u] = mk_plus (t, u)
| mk_sum (t :: ts) = mk_plus (t, mk_sum ts);
@@ -355,24 +345,21 @@
handle TERM _ => find_first_coeff (t::past) u terms;
-(*Simplify Numeral1*n and n*Numeral1 to n*)
-val add_0s = map rename_numerals
- [real_add_zero_left, real_add_zero_right];
-val mult_plus_1s = map rename_numerals
- [real_mult_1, real_mult_1_right];
-val mult_minus_1s = map rename_numerals
- [real_mult_minus_1, real_mult_minus_1_right];
-val mult_1s = mult_plus_1s @ mult_minus_1s;
+(*Simplify Numeral0+n, n+Numeral0, Numeral1*n, n*Numeral1*)
+val add_0s = map rename_numerals [real_add_zero_left, real_add_zero_right];
+val mult_1s = map rename_numerals [real_mult_1, real_mult_1_right] @
+ [real_mult_minus1, real_mult_minus1_right];
(*To perform binary arithmetic*)
val bin_simps =
- [add_real_number_of, real_add_number_of_left, minus_real_number_of,
+ [real_numeral_0_eq_0 RS sym, real_numeral_1_eq_1 RS sym,
+ add_real_number_of, real_add_number_of_left, minus_real_number_of,
diff_real_number_of, mult_real_number_of, real_mult_number_of_left] @
bin_arith_simps @ bin_rel_simps;
(*To evaluate binary negations of coefficients*)
val real_minus_simps = NCons_simps @
- [minus_real_number_of,
+ [real_minus_1_eq_m1, minus_real_number_of,
bin_minus_1, bin_minus_0, bin_minus_Pls, bin_minus_Min,
bin_pred_1, bin_pred_0, bin_pred_Pls, bin_pred_Min];
@@ -494,16 +481,17 @@
val dest_coeff = dest_coeff 1
val left_distrib = left_real_add_mult_distrib RS trans
val prove_conv = prove_conv_nohyps "real_combine_numerals"
- val trans_tac = trans_tac
+ val trans_tac = trans_tac
val norm_tac =
- ALLGOALS (simp_tac (HOL_ss addsimps add_0s@mult_1s@diff_simps@
- real_minus_simps@real_add_ac))
+ ALLGOALS (simp_tac (HOL_ss addsimps numeral_syms@add_0s@mult_1s@
+ diff_simps@real_minus_simps@real_add_ac))
THEN ALLGOALS (simp_tac (HOL_ss addsimps bin_simps@real_mult_minus_simps))
THEN ALLGOALS (simp_tac (HOL_ss addsimps real_minus_from_mult_simps@
real_add_ac@real_mult_ac))
val numeral_simp_tac = ALLGOALS
(simp_tac (HOL_ss addsimps add_0s@bin_simps))
- val simplify_meta_eq = simplify_meta_eq
+ val simplify_meta_eq =
+ Int_Numeral_Simprocs.simplify_meta_eq (add_0s@mult_1s)
end;
structure CombineNumerals = CombineNumeralsFun(CombineNumeralsData);
@@ -533,9 +521,48 @@
[real_mult_1, real_mult_1_right]
(([th, cancel_th]) MRS trans);
+(*** Making constant folding work for 0 and 1 too ***)
+
+structure RealAbstractNumeralsData =
+ struct
+ val dest_eq = HOLogic.dest_eq o HOLogic.dest_Trueprop o concl_of
+ val is_numeral = Bin_Simprocs.is_numeral
+ val numeral_0_eq_0 = real_numeral_0_eq_0
+ val numeral_1_eq_1 = real_numeral_1_eq_1
+ val prove_conv = prove_conv_nohyps "real_abstract_numerals"
+ fun norm_tac simps = ALLGOALS (simp_tac (HOL_ss addsimps simps))
+ val simplify_meta_eq = Bin_Simprocs.simplify_meta_eq
+ end
+
+structure RealAbstractNumerals = AbstractNumeralsFun (RealAbstractNumeralsData)
+
+(*For addition, we already have rules for the operand 0.
+ Multiplication is omitted because there are already special rules for
+ both 0 and 1 as operands. Unary minus is trivial, just have - 1 = -1.
+ For the others, having three patterns is a compromise between just having
+ one (many spurious calls) and having nine (just too many!) *)
+val eval_numerals =
+ map prep_simproc
+ [("real_add_eval_numerals",
+ prep_pats ["(m::real) + 1", "(m::real) + number_of v"],
+ RealAbstractNumerals.proc add_real_number_of),
+ ("real_diff_eval_numerals",
+ prep_pats ["(m::real) - 1", "(m::real) - number_of v"],
+ RealAbstractNumerals.proc diff_real_number_of),
+ ("real_eq_eval_numerals",
+ prep_pats ["(m::real) = 0", "(m::real) = 1", "(m::real) = number_of v"],
+ RealAbstractNumerals.proc eq_real_number_of),
+ ("real_less_eval_numerals",
+ prep_pats ["(m::real) < 0", "(m::real) < 1", "(m::real) < number_of v"],
+ RealAbstractNumerals.proc less_real_number_of),
+ ("real_le_eval_numerals",
+ prep_pats ["(m::real) <= 0", "(m::real) <= 1", "(m::real) <= number_of v"],
+ RealAbstractNumerals.proc le_real_number_of_eq_not_less)]
+
end;
+Addsimprocs Real_Numeral_Simprocs.eval_numerals;
Addsimprocs Real_Numeral_Simprocs.cancel_numerals;
Addsimprocs [Real_Numeral_Simprocs.combine_numerals];
@@ -623,3 +650,5 @@
by (etac real_mult_le_mono2 1);
by (assume_tac 1);
qed "real_mult_le_mono";
+
+Addsimps [real_minus_1_eq_m1];
--- a/src/HOL/Real/RealDef.ML Thu Nov 01 21:12:13 2001 +0100
+++ b/src/HOL/Real/RealDef.ML Fri Nov 02 17:55:24 2001 +0100
@@ -11,7 +11,7 @@
(*** Proving that realrel is an equivalence relation ***)
Goal "[| (x1::preal) + y2 = x2 + y1; x2 + y3 = x3 + y2 |] \
-\ ==> x1 + y3 = x3 + y1";
+\ ==> x1 + y3 = x3 + y1";
by (res_inst_tac [("C","y2")] preal_add_right_cancel 1);
by (rotate_tac 1 1 THEN dtac sym 1);
by (asm_full_simp_tac (simpset() addsimps preal_add_ac) 1);
@@ -380,17 +380,18 @@
real_minus_mult_eq1]) 1);
qed "real_minus_mult_eq2";
+(*Pull negations out*)
Addsimps [real_minus_mult_eq1 RS sym, real_minus_mult_eq2 RS sym];
Goal "(- (1::real)) * z = -z";
by (Simp_tac 1);
qed "real_mult_minus_1";
+Addsimps [real_mult_minus_1];
Goal "z * (- (1::real)) = -z";
by (stac real_mult_commute 1);
by (Simp_tac 1);
qed "real_mult_minus_1_right";
-
Addsimps [real_mult_minus_1_right];
Goal "(-x) * (-y) = x * (y::real)";
@@ -458,12 +459,14 @@
by (cut_inst_tac [("r1.0","xa"),("r2.0","y")] preal_linear 1);
by (auto_tac (claset() addSDs [preal_less_add_left_Ex],
simpset() addsimps [real_zero_iff RS sym]));
-by (res_inst_tac [("x","Abs_REAL (realrel `` \
-\ {(preal_of_prat(prat_of_pnat 1p),pinv(D)+\
-\ preal_of_prat(prat_of_pnat 1p))})")] exI 1);
-by (res_inst_tac [("x","Abs_REAL (realrel `` \
-\ {(pinv(D)+preal_of_prat(prat_of_pnat 1p),\
-\ preal_of_prat(prat_of_pnat 1p))})")] exI 2);
+by (res_inst_tac [("x",
+ "Abs_REAL (realrel `` \
+\ {(preal_of_prat(prat_of_pnat 1), \
+\ pinv(D) + preal_of_prat(prat_of_pnat 1))})")] exI 1);
+by (res_inst_tac [("x",
+ "Abs_REAL (realrel `` \
+\ {(pinv(D) + preal_of_prat(prat_of_pnat 1),\
+\ preal_of_prat(prat_of_pnat 1))})")] exI 2);
by (auto_tac (claset(),
simpset() addsimps [real_mult,
pnat_one_def,preal_mult_1_right,preal_add_mult_distrib2,
@@ -985,7 +988,6 @@
real_of_preal_not_less_zero,real_of_preal_zero_less,
real_of_preal_minus_less_zero]));
qed "real_minus_zero_less_iff";
-
Addsimps [real_minus_zero_less_iff];
Goal "(-R < 0) = ((0::real) < R)";
@@ -995,6 +997,7 @@
real_of_preal_not_less_zero,real_of_preal_zero_less,
real_of_preal_minus_less_zero]));
qed "real_minus_zero_less_iff2";
+Addsimps [real_minus_zero_less_iff2];
(*Alternative definition for real_less*)
Goal "R < S ==> EX T::real. 0 < T & R + T = S";
@@ -1007,8 +1010,7 @@
real_of_preal_not_minus_gt_zero]));
by (res_inst_tac [("x","real_of_preal D")] exI 1);
by (res_inst_tac [("x","real_of_preal m+real_of_preal ma")] exI 2);
-by (res_inst_tac [("x","real_of_preal m")] exI 3);
-by (res_inst_tac [("x","real_of_preal D")] exI 4);
+by (res_inst_tac [("x","real_of_preal D")] exI 3);
by (auto_tac (claset(),
simpset() addsimps [real_of_preal_zero_less,
real_of_preal_sum_zero_less,real_add_assoc]));
--- a/src/HOL/Real/RealDef.thy Thu Nov 01 21:12:13 2001 +0100
+++ b/src/HOL/Real/RealDef.thy Fri Nov 02 17:55:24 2001 +0100
@@ -33,11 +33,13 @@
defs
real_zero_def
- "0 == Abs_REAL(realrel``{(preal_of_prat(prat_of_pnat 1p),
- preal_of_prat(prat_of_pnat 1p))})"
+ "0 == Abs_REAL(realrel``{(preal_of_prat(prat_of_pnat 1),
+ preal_of_prat(prat_of_pnat 1))})"
+
real_one_def
- "1 == Abs_REAL(realrel``{(preal_of_prat(prat_of_pnat 1p) +
- preal_of_prat(prat_of_pnat 1p),preal_of_prat(prat_of_pnat 1p))})"
+ "1 == Abs_REAL(realrel``
+ {(preal_of_prat(prat_of_pnat 1) + preal_of_prat(prat_of_pnat 1),
+ preal_of_prat(prat_of_pnat 1))})"
real_minus_def
"- R == Abs_REAL(UN (x,y):Rep_REAL(R). realrel``{(y,x)})"
@@ -53,12 +55,12 @@
constdefs
- (** these don't use the overloaded real because users don't see them **)
+ (** these don't use the overloaded "real" function: users don't see them **)
real_of_preal :: preal => real
"real_of_preal m ==
- Abs_REAL(realrel``{(m+preal_of_prat(prat_of_pnat 1p),
- preal_of_prat(prat_of_pnat 1p))})"
+ Abs_REAL(realrel``{(m + preal_of_prat(prat_of_pnat 1),
+ preal_of_prat(prat_of_pnat 1))})"
real_of_posnat :: nat => real
"real_of_posnat n == real_of_preal(preal_of_prat(prat_of_pnat(pnat_of_nat n)))"
--- a/src/HOL/Real/RealOrd.ML Thu Nov 01 21:12:13 2001 +0100
+++ b/src/HOL/Real/RealOrd.ML Fri Nov 02 17:55:24 2001 +0100
@@ -128,14 +128,14 @@
by (REPEAT(dtac (real_minus_zero_less_iff RS iffD2) 1));
by (dtac real_mult_order 1 THEN assume_tac 1);
by (Asm_full_simp_tac 1);
-qed "real_mult_less_zero1";
+qed "neg_real_mult_order";
Goal "[| 0 < x; y < 0 |] ==> x*y < (0::real)";
by (dtac (real_minus_zero_less_iff RS iffD2) 1);
by (dtac real_mult_order 1 THEN assume_tac 1);
by (rtac (real_minus_zero_less_iff RS iffD1) 1);
by (Asm_full_simp_tac 1);
-qed "real_mult_less_zero";
+qed "real_mult_less_0";
Goalw [real_one_def] "0 < (1::real)";
by (auto_tac (claset() addIs [real_gt_zero_preal_Ex RS iffD2],
@@ -255,7 +255,7 @@
Goal "(0::real) <= x*x";
by (res_inst_tac [("R2.0","0"),("R1.0","x")] real_linear_less2 1);
by (auto_tac (claset() addIs [real_mult_order,
- real_mult_less_zero1,order_less_imp_le],
+ neg_real_mult_order,order_less_imp_le],
simpset()));
qed "real_le_square";
Addsimps [real_le_square];
@@ -390,18 +390,18 @@
by (forward_tac [real_not_refl2 RS not_sym] 1);
by (dtac (real_not_refl2 RS not_sym RS real_inverse_not_zero) 1);
by (EVERY1[dtac order_le_imp_less_or_eq, Step_tac]);
-by (dtac real_mult_less_zero1 1 THEN assume_tac 1);
+by (dtac neg_real_mult_order 1 THEN assume_tac 1);
by (auto_tac (claset() addIs [real_zero_less_one RS real_less_asym],
simpset()));
-qed "real_inverse_gt_zero";
+qed "real_inverse_gt_0";
Goal "x < 0 ==> inverse (x::real) < 0";
by (ftac real_not_refl2 1);
by (dtac (real_minus_zero_less_iff RS iffD2) 1);
by (rtac (real_minus_zero_less_iff RS iffD1) 1);
by (stac (real_minus_inverse RS sym) 1);
-by (auto_tac (claset() addIs [real_inverse_gt_zero], simpset()));
-qed "real_inverse_less_zero";
+by (auto_tac (claset() addIs [real_inverse_gt_0], simpset()));
+qed "real_inverse_less_0";
Goal "[| (0::real) < z; x < y |] ==> x*z < y*z";
by (rotate_tac 1 1);
@@ -419,7 +419,7 @@
Goal "[| (0::real) < z; x * z < y * z |] ==> x < y";
by (forw_inst_tac [("x","x*z")]
- (real_inverse_gt_zero RS real_mult_less_mono1) 1);
+ (real_inverse_gt_0 RS real_mult_less_mono1) 1);
by (auto_tac (claset(),
simpset() addsimps [real_mult_assoc,real_not_refl2 RS not_sym]));
qed "real_mult_less_cancel1";
@@ -495,13 +495,13 @@
Goal "[| 0 < r; r < x |] ==> inverse x < inverse (r::real)";
by (ftac order_less_trans 1 THEN assume_tac 1);
-by (ftac real_inverse_gt_zero 1);
-by (forw_inst_tac [("x","x")] real_inverse_gt_zero 1);
+by (ftac real_inverse_gt_0 1);
+by (forw_inst_tac [("x","x")] real_inverse_gt_0 1);
by (forw_inst_tac [("x","r"),("z","inverse r")] real_mult_less_mono1 1);
by (assume_tac 1);
by (asm_full_simp_tac (simpset() addsimps [real_not_refl2 RS
not_sym RS real_mult_inv_right]) 1);
-by (ftac real_inverse_gt_zero 1);
+by (ftac real_inverse_gt_0 1);
by (forw_inst_tac [("x","(1::real)"),("z","inverse x")] real_mult_less_mono2 1);
by (assume_tac 1);
by (asm_full_simp_tac (simpset() addsimps [real_not_refl2 RS
@@ -556,32 +556,32 @@
Goal "((0::real) < x*y) = (0 < x & 0 < y | x < 0 & y < 0)";
by (auto_tac (claset(),
simpset() addsimps [order_le_less, linorder_not_less,
- real_mult_order, real_mult_less_zero1]));
+ real_mult_order, neg_real_mult_order]));
by (ALLGOALS (rtac ccontr));
by (auto_tac (claset(), simpset() addsimps [order_le_less, linorder_not_less]));
by (ALLGOALS (etac rev_mp));
-by (ALLGOALS (dtac real_mult_less_zero THEN' assume_tac));
+by (ALLGOALS (dtac real_mult_less_0 THEN' assume_tac));
by (auto_tac (claset() addDs [order_less_not_sym],
simpset() addsimps [real_mult_commute]));
-qed "real_zero_less_mult_iff";
+qed "real_0_less_mult_iff";
Goal "((0::real) <= x*y) = (0 <= x & 0 <= y | x <= 0 & y <= 0)";
by (auto_tac (claset(),
simpset() addsimps [order_le_less, linorder_not_less,
- real_zero_less_mult_iff]));
-qed "real_zero_le_mult_iff";
+ real_0_less_mult_iff]));
+qed "real_0_le_mult_iff";
Goal "(x*y < (0::real)) = (0 < x & y < 0 | x < 0 & 0 < y)";
by (auto_tac (claset(),
- simpset() addsimps [real_zero_le_mult_iff,
+ simpset() addsimps [real_0_le_mult_iff,
linorder_not_le RS sym]));
by (auto_tac (claset() addDs [order_less_not_sym],
simpset() addsimps [linorder_not_le]));
-qed "real_mult_less_zero_iff";
+qed "real_mult_less_0_iff";
Goal "(x*y <= (0::real)) = (0 <= x & y <= 0 | x <= 0 & 0 <= y)";
by (auto_tac (claset() addDs [order_less_not_sym],
- simpset() addsimps [real_zero_less_mult_iff,
+ simpset() addsimps [real_0_less_mult_iff,
linorder_not_less RS sym]));
-qed "real_mult_le_zero_iff";
+qed "real_mult_le_0_iff";
--- a/src/HOL/Real/RealPow.ML Thu Nov 01 21:12:13 2001 +0100
+++ b/src/HOL/Real/RealPow.ML Fri Nov 02 17:55:24 2001 +0100
@@ -7,17 +7,17 @@
bind_thm ("realpow_Suc", thm "realpow_Suc");
-Goal "(Numeral0::real) ^ (Suc n) = Numeral0";
+Goal "(0::real) ^ (Suc n) = 0";
by Auto_tac;
qed "realpow_zero";
Addsimps [realpow_zero];
-Goal "r ~= (Numeral0::real) --> r ^ n ~= Numeral0";
+Goal "r ~= (0::real) --> r ^ n ~= 0";
by (induct_tac "n" 1);
by Auto_tac;
qed_spec_mp "realpow_not_zero";
-Goal "r ^ n = (Numeral0::real) ==> r = Numeral0";
+Goal "r ^ n = (0::real) ==> r = 0";
by (rtac ccontr 1);
by (auto_tac (claset() addDs [realpow_not_zero], simpset()));
qed "realpow_zero_zero";
@@ -46,37 +46,37 @@
by (Simp_tac 1);
qed "realpow_two";
-Goal "(Numeral0::real) < r --> Numeral0 < r ^ n";
+Goal "(0::real) < r --> 0 < r ^ n";
by (induct_tac "n" 1);
-by (auto_tac (claset() addIs [rename_numerals real_mult_order],
+by (auto_tac (claset() addIs [real_mult_order],
simpset() addsimps [real_zero_less_one]));
qed_spec_mp "realpow_gt_zero";
-Goal "(Numeral0::real) <= r --> Numeral0 <= r ^ n";
+Goal "(0::real) <= r --> 0 <= r ^ n";
by (induct_tac "n" 1);
by (auto_tac (claset(), simpset() addsimps [real_0_le_mult_iff]));
qed_spec_mp "realpow_ge_zero";
-Goal "(Numeral0::real) <= x & x <= y --> x ^ n <= y ^ n";
+Goal "(0::real) <= x & x <= y --> x ^ n <= y ^ n";
by (induct_tac "n" 1);
by (auto_tac (claset() addSIs [real_mult_le_mono], simpset()));
by (asm_simp_tac (simpset() addsimps [realpow_ge_zero]) 1);
qed_spec_mp "realpow_le";
-Goal "(Numeral0::real) < x & x < y & 0 < n --> x ^ n < y ^ n";
+Goal "(0::real) < x & x < y & 0 < n --> x ^ n < y ^ n";
by (induct_tac "n" 1);
-by (auto_tac (claset() addIs [rename_numerals real_mult_less_mono, gr0I]
+by (auto_tac (claset() addIs [real_mult_less_mono, gr0I]
addDs [realpow_gt_zero],
simpset()));
qed_spec_mp "realpow_less";
-Goal "Numeral1 ^ n = (Numeral1::real)";
+Goal "1 ^ n = (1::real)";
by (induct_tac "n" 1);
by Auto_tac;
qed "realpow_eq_one";
Addsimps [realpow_eq_one];
-Goal "abs((-1) ^ n) = (Numeral1::real)";
+Goal "abs((-1) ^ n) = (1::real)";
by (induct_tac "n" 1);
by (auto_tac (claset(), simpset() addsimps [abs_mult]));
qed "abs_realpow_minus_one";
@@ -87,14 +87,14 @@
by (auto_tac (claset(),simpset() addsimps real_mult_ac));
qed "realpow_mult";
-Goal "(Numeral0::real) <= r^ Suc (Suc 0)";
-by (simp_tac (simpset() addsimps [rename_numerals real_le_square]) 1);
+Goal "(0::real) <= r^ Suc (Suc 0)";
+by (simp_tac (simpset() addsimps [real_le_square]) 1);
qed "realpow_two_le";
Addsimps [realpow_two_le];
Goal "abs((x::real)^Suc (Suc 0)) = x^Suc (Suc 0)";
by (simp_tac (simpset() addsimps [abs_eqI1,
- rename_numerals real_le_square]) 1);
+ real_le_square]) 1);
qed "abs_realpow_two";
Addsimps [abs_realpow_two];
@@ -104,31 +104,31 @@
qed "realpow_two_abs";
Addsimps [realpow_two_abs];
-Goal "(Numeral1::real) < r ==> Numeral1 < r^ (Suc (Suc 0))";
+Goal "(1::real) < r ==> 1 < r^ (Suc (Suc 0))";
by Auto_tac;
-by (cut_facts_tac [rename_numerals real_zero_less_one] 1);
-by (forw_inst_tac [("x","Numeral0")] order_less_trans 1);
+by (cut_facts_tac [real_zero_less_one] 1);
+by (forw_inst_tac [("x","0")] order_less_trans 1);
by (assume_tac 1);
-by (dres_inst_tac [("z","r"),("x","Numeral1")]
- (rename_numerals real_mult_less_mono1) 1);
+by (dres_inst_tac [("z","r"),("x","1")]
+ (real_mult_less_mono1) 1);
by (auto_tac (claset() addIs [order_less_trans], simpset()));
qed "realpow_two_gt_one";
-Goal "(Numeral1::real) < r --> Numeral1 <= r ^ n";
+Goal "(1::real) < r --> 1 <= r ^ n";
by (induct_tac "n" 1);
by Auto_tac;
-by (subgoal_tac "Numeral1*Numeral1 <= r * r^n" 1);
+by (subgoal_tac "1*1 <= r * r^n" 1);
by (rtac real_mult_le_mono 2);
by Auto_tac;
qed_spec_mp "realpow_ge_one";
-Goal "(Numeral1::real) <= r ==> Numeral1 <= r ^ n";
+Goal "(1::real) <= r ==> 1 <= r ^ n";
by (dtac order_le_imp_less_or_eq 1);
by (auto_tac (claset() addDs [realpow_ge_one], simpset()));
qed "realpow_ge_one2";
-Goal "(Numeral1::real) <= 2 ^ n";
-by (res_inst_tac [("y","Numeral1 ^ n")] order_trans 1);
+Goal "(1::real) <= 2 ^ n";
+by (res_inst_tac [("y","1 ^ n")] order_trans 1);
by (rtac realpow_le 2);
by (auto_tac (claset() addIs [order_less_imp_le], simpset()));
qed "two_realpow_ge_one";
@@ -142,142 +142,142 @@
qed "two_realpow_gt";
Addsimps [two_realpow_gt,two_realpow_ge_one];
-Goal "(-1) ^ (2*n) = (Numeral1::real)";
+Goal "(-1) ^ (2*n) = (1::real)";
by (induct_tac "n" 1);
by Auto_tac;
qed "realpow_minus_one";
Addsimps [realpow_minus_one];
-Goal "(-1) ^ Suc (2*n) = -(Numeral1::real)";
+Goal "(-1) ^ Suc (2*n) = -(1::real)";
by Auto_tac;
qed "realpow_minus_one_odd";
Addsimps [realpow_minus_one_odd];
-Goal "(-1) ^ Suc (Suc (2*n)) = (Numeral1::real)";
+Goal "(-1) ^ Suc (Suc (2*n)) = (1::real)";
by Auto_tac;
qed "realpow_minus_one_even";
Addsimps [realpow_minus_one_even];
-Goal "(Numeral0::real) < r & r < (Numeral1::real) --> r ^ Suc n < r ^ n";
+Goal "(0::real) < r & r < (1::real) --> r ^ Suc n < r ^ n";
by (induct_tac "n" 1);
by Auto_tac;
qed_spec_mp "realpow_Suc_less";
-Goal "Numeral0 <= r & r < (Numeral1::real) --> r ^ Suc n <= r ^ n";
+Goal "0 <= r & r < (1::real) --> r ^ Suc n <= r ^ n";
by (induct_tac "n" 1);
by (auto_tac (claset() addIs [order_less_imp_le]
addSDs [order_le_imp_less_or_eq],
simpset()));
qed_spec_mp "realpow_Suc_le";
-Goal "(Numeral0::real) <= Numeral0 ^ n";
+Goal "(0::real) <= 0 ^ n";
by (case_tac "n" 1);
by Auto_tac;
qed "realpow_zero_le";
Addsimps [realpow_zero_le];
-Goal "Numeral0 < r & r < (Numeral1::real) --> r ^ Suc n <= r ^ n";
+Goal "0 < r & r < (1::real) --> r ^ Suc n <= r ^ n";
by (blast_tac (claset() addSIs [order_less_imp_le,
realpow_Suc_less]) 1);
qed_spec_mp "realpow_Suc_le2";
-Goal "[| Numeral0 <= r; r < (Numeral1::real) |] ==> r ^ Suc n <= r ^ n";
+Goal "[| 0 <= r; r < (1::real) |] ==> r ^ Suc n <= r ^ n";
by (etac (order_le_imp_less_or_eq RS disjE) 1);
by (rtac realpow_Suc_le2 1);
by Auto_tac;
qed "realpow_Suc_le3";
-Goal "Numeral0 <= r & r < (Numeral1::real) & n < N --> r ^ N <= r ^ n";
+Goal "0 <= r & r < (1::real) & n < N --> r ^ N <= r ^ n";
by (induct_tac "N" 1);
by (ALLGOALS Asm_simp_tac);
by (Clarify_tac 1);
-by (subgoal_tac "r * r ^ na <= Numeral1 * r ^ n" 1);
+by (subgoal_tac "r * r ^ na <= 1 * r ^ n" 1);
by (Asm_full_simp_tac 1);
by (rtac real_mult_le_mono 1);
by (auto_tac (claset(), simpset() addsimps [realpow_ge_zero, less_Suc_eq]));
qed_spec_mp "realpow_less_le";
-Goal "[| Numeral0 <= r; r < (Numeral1::real); n <= N |] ==> r ^ N <= r ^ n";
+Goal "[| 0 <= r; r < (1::real); n <= N |] ==> r ^ N <= r ^ n";
by (dres_inst_tac [("n","N")] le_imp_less_or_eq 1);
by (auto_tac (claset() addIs [realpow_less_le],
simpset()));
qed "realpow_le_le";
-Goal "[| Numeral0 < r; r < (Numeral1::real) |] ==> r ^ Suc n <= r";
+Goal "[| 0 < r; r < (1::real) |] ==> r ^ Suc n <= r";
by (dres_inst_tac [("n","1"),("N","Suc n")]
(order_less_imp_le RS realpow_le_le) 1);
by Auto_tac;
qed "realpow_Suc_le_self";
-Goal "[| Numeral0 < r; r < (Numeral1::real) |] ==> r ^ Suc n < Numeral1";
+Goal "[| 0 < r; r < (1::real) |] ==> r ^ Suc n < 1";
by (blast_tac (claset() addIs [realpow_Suc_le_self, order_le_less_trans]) 1);
qed "realpow_Suc_less_one";
-Goal "(Numeral1::real) <= r --> r ^ n <= r ^ Suc n";
+Goal "(1::real) <= r --> r ^ n <= r ^ Suc n";
by (induct_tac "n" 1);
by Auto_tac;
qed_spec_mp "realpow_le_Suc";
-Goal "(Numeral1::real) < r --> r ^ n < r ^ Suc n";
+Goal "(1::real) < r --> r ^ n < r ^ Suc n";
by (induct_tac "n" 1);
by Auto_tac;
qed_spec_mp "realpow_less_Suc";
-Goal "(Numeral1::real) < r --> r ^ n <= r ^ Suc n";
+Goal "(1::real) < r --> r ^ n <= r ^ Suc n";
by (blast_tac (claset() addSIs [order_less_imp_le, realpow_less_Suc]) 1);
qed_spec_mp "realpow_le_Suc2";
-Goal "(Numeral1::real) < r & n < N --> r ^ n <= r ^ N";
+Goal "(1::real) < r & n < N --> r ^ n <= r ^ N";
by (induct_tac "N" 1);
by Auto_tac;
by (ALLGOALS(forw_inst_tac [("n","na")] realpow_ge_one));
-by (ALLGOALS(dtac (rename_numerals real_mult_self_le)));
+by (ALLGOALS(dtac (real_mult_self_le)));
by (assume_tac 1);
by (assume_tac 2);
by (auto_tac (claset() addIs [order_trans],
simpset() addsimps [less_Suc_eq]));
qed_spec_mp "realpow_gt_ge";
-Goal "(Numeral1::real) <= r & n < N --> r ^ n <= r ^ N";
+Goal "(1::real) <= r & n < N --> r ^ n <= r ^ N";
by (induct_tac "N" 1);
by Auto_tac;
by (ALLGOALS(forw_inst_tac [("n","na")] realpow_ge_one2));
-by (ALLGOALS(dtac (rename_numerals real_mult_self_le2)));
+by (ALLGOALS(dtac (real_mult_self_le2)));
by (assume_tac 1);
by (assume_tac 2);
by (auto_tac (claset() addIs [order_trans],
simpset() addsimps [less_Suc_eq]));
qed_spec_mp "realpow_gt_ge2";
-Goal "[| (Numeral1::real) < r; n <= N |] ==> r ^ n <= r ^ N";
+Goal "[| (1::real) < r; n <= N |] ==> r ^ n <= r ^ N";
by (dres_inst_tac [("n","N")] le_imp_less_or_eq 1);
by (auto_tac (claset() addIs [realpow_gt_ge], simpset()));
qed "realpow_ge_ge";
-Goal "[| (Numeral1::real) <= r; n <= N |] ==> r ^ n <= r ^ N";
+Goal "[| (1::real) <= r; n <= N |] ==> r ^ n <= r ^ N";
by (dres_inst_tac [("n","N")] le_imp_less_or_eq 1);
by (auto_tac (claset() addIs [realpow_gt_ge2], simpset()));
qed "realpow_ge_ge2";
-Goal "(Numeral1::real) < r ==> r <= r ^ Suc n";
+Goal "(1::real) < r ==> r <= r ^ Suc n";
by (dres_inst_tac [("n","1"),("N","Suc n")]
realpow_ge_ge 1);
by Auto_tac;
qed_spec_mp "realpow_Suc_ge_self";
-Goal "(Numeral1::real) <= r ==> r <= r ^ Suc n";
+Goal "(1::real) <= r ==> r <= r ^ Suc n";
by (dres_inst_tac [("n","1"),("N","Suc n")]
realpow_ge_ge2 1);
by Auto_tac;
qed_spec_mp "realpow_Suc_ge_self2";
-Goal "[| (Numeral1::real) < r; 0 < n |] ==> r <= r ^ n";
+Goal "[| (1::real) < r; 0 < n |] ==> r <= r ^ n";
by (dtac (less_not_refl2 RS not0_implies_Suc) 1);
by (auto_tac (claset() addSIs
[realpow_Suc_ge_self],simpset()));
qed "realpow_ge_self";
-Goal "[| (Numeral1::real) <= r; 0 < n |] ==> r <= r ^ n";
+Goal "[| (1::real) <= r; 0 < n |] ==> r <= r ^ n";
by (dtac (less_not_refl2 RS not0_implies_Suc) 1);
by (auto_tac (claset() addSIs [realpow_Suc_ge_self2],simpset()));
qed "realpow_ge_self2";
@@ -289,7 +289,7 @@
qed_spec_mp "realpow_minus_mult";
Addsimps [realpow_minus_mult];
-Goal "r ~= Numeral0 ==> r * inverse r ^Suc (Suc 0) = inverse (r::real)";
+Goal "r ~= 0 ==> r * inverse r ^Suc (Suc 0) = inverse (r::real)";
by (asm_simp_tac (simpset() addsimps [realpow_two,
real_mult_assoc RS sym]) 1);
qed "realpow_two_mult_inverse";
@@ -313,7 +313,7 @@
qed "realpow_two_disj";
(* used in Transc *)
-Goal "[|(x::real) ~= Numeral0; m <= n |] ==> x ^ (n - m) = x ^ n * inverse (x ^ m)";
+Goal "[|(x::real) ~= 0; m <= n |] ==> x ^ (n - m) = x ^ n * inverse (x ^ m)";
by (auto_tac (claset(),
simpset() addsimps [le_eq_less_or_eq, less_iff_Suc_add, realpow_add,
realpow_not_zero] @ real_mult_ac));
@@ -325,15 +325,15 @@
simpset() addsimps [real_of_nat_one, real_of_nat_mult]));
qed "realpow_real_of_nat";
-Goal "Numeral0 < real (Suc (Suc 0) ^ n)";
+Goal "0 < real (Suc (Suc 0) ^ n)";
by (induct_tac "n" 1);
by (auto_tac (claset(),
- simpset() addsimps [real_of_nat_mult, real_zero_less_mult_iff]));
+ simpset() addsimps [real_of_nat_mult, real_0_less_mult_iff]));
qed "realpow_real_of_nat_two_pos";
Addsimps [realpow_real_of_nat_two_pos];
-Goal "(Numeral0::real) <= x --> Numeral0 <= y --> x ^ Suc n <= y ^ Suc n --> x <= y";
+Goal "(0::real) <= x --> 0 <= y --> x ^ Suc n <= y ^ Suc n --> x <= y";
by (induct_tac "n" 1);
by Auto_tac;
by (asm_full_simp_tac (simpset() addsimps [linorder_not_less RS sym]) 1);
@@ -345,7 +345,7 @@
by Auto_tac;
qed_spec_mp "realpow_increasing";
-Goal "[| (Numeral0::real) <= x; Numeral0 <= y; x ^ Suc n = y ^ Suc n |] ==> x = y";
+Goal "[| (0::real) <= x; 0 <= y; x ^ Suc n = y ^ Suc n |] ==> x = y";
by (blast_tac (claset() addIs [realpow_increasing, order_antisym,
order_eq_refl, sym]) 1);
qed_spec_mp "realpow_Suc_cancel_eq";
--- a/src/HOL/Real/RealPow.thy Thu Nov 01 21:12:13 2001 +0100
+++ b/src/HOL/Real/RealPow.thy Fri Nov 02 17:55:24 2001 +0100
@@ -15,7 +15,7 @@
instance real :: power ..
primrec (realpow)
- realpow_0: "r ^ 0 = Numeral1"
+ realpow_0: "r ^ 0 = 1"
realpow_Suc: "r ^ (Suc n) = (r::real) * (r ^ n)"
end
--- a/src/HOL/Real/ex/BinEx.thy Thu Nov 01 21:12:13 2001 +0100
+++ b/src/HOL/Real/ex/BinEx.thy Fri Nov 02 17:55:24 2001 +0100
@@ -57,7 +57,7 @@
lemma "(13557456::real) < 18678654"
by simp
-lemma "(999999::real) \<le> (1000001 + Numeral1) - 2"
+lemma "(999999::real) \<le> (1000001 + 1) - 2"
by simp
lemma "(1234567::real) \<le> 1234567"
@@ -66,16 +66,16 @@
text {* \medskip Tests *}
-lemma "(x + y = x) = (y = (Numeral0::real))"
+lemma "(x + y = x) = (y = (0::real))"
by arith
-lemma "(x + y = y) = (x = (Numeral0::real))"
+lemma "(x + y = y) = (x = (0::real))"
by arith
-lemma "(x + y = (Numeral0::real)) = (x = -y)"
+lemma "(x + y = (0::real)) = (x = -y)"
by arith
-lemma "(x + y = (Numeral0::real)) = (y = -x)"
+lemma "(x + y = (0::real)) = (y = -x)"
by arith
lemma "((x + y) < (x + z)) = (y < (z::real))"
@@ -123,25 +123,25 @@
lemma "\<not>(x \<le> y \<and> y < (x::real))"
by arith
-lemma "(-x < (Numeral0::real)) = (Numeral0 < x)"
+lemma "(-x < (0::real)) = (0 < x)"
by arith
-lemma "((Numeral0::real) < -x) = (x < Numeral0)"
+lemma "((0::real) < -x) = (x < 0)"
by arith
-lemma "(-x \<le> (Numeral0::real)) = (Numeral0 \<le> x)"
+lemma "(-x \<le> (0::real)) = (0 \<le> x)"
by arith
-lemma "((Numeral0::real) \<le> -x) = (x \<le> Numeral0)"
+lemma "((0::real) \<le> -x) = (x \<le> 0)"
by arith
lemma "(x::real) = y \<or> x < y \<or> y < x"
by arith
-lemma "(x::real) = Numeral0 \<or> Numeral0 < x \<or> Numeral0 < -x"
+lemma "(x::real) = 0 \<or> 0 < x \<or> 0 < -x"
by arith
-lemma "(Numeral0::real) \<le> x \<or> Numeral0 \<le> -x"
+lemma "(0::real) \<le> x \<or> 0 \<le> -x"
by arith
lemma "((x::real) + y \<le> x + z) = (y \<le> z)"
@@ -156,16 +156,16 @@
lemma "(w::real) \<le> x \<and> y \<le> z ==> w + y \<le> x + z"
by arith
-lemma "(Numeral0::real) \<le> x \<and> Numeral0 \<le> y ==> Numeral0 \<le> x + y"
+lemma "(0::real) \<le> x \<and> 0 \<le> y ==> 0 \<le> x + y"
by arith
-lemma "(Numeral0::real) < x \<and> Numeral0 < y ==> Numeral0 < x + y"
+lemma "(0::real) < x \<and> 0 < y ==> 0 < x + y"
by arith
-lemma "(-x < y) = (Numeral0 < x + (y::real))"
+lemma "(-x < y) = (0 < x + (y::real))"
by arith
-lemma "(x < -y) = (x + y < (Numeral0::real))"
+lemma "(x < -y) = (x + y < (0::real))"
by arith
lemma "(y < x + -z) = (y + z < (x::real))"
@@ -174,7 +174,7 @@
lemma "(x + -y < z) = (x < z + (y::real))"
by arith
-lemma "x \<le> y ==> x < y + (Numeral1::real)"
+lemma "x \<le> y ==> x < y + (1::real)"
by arith
lemma "(x - y) + y = (x::real)"
@@ -183,31 +183,31 @@
lemma "y + (x - y) = (x::real)"
by arith
-lemma "x - x = (Numeral0::real)"
+lemma "x - x = (0::real)"
by arith
-lemma "(x - y = Numeral0) = (x = (y::real))"
+lemma "(x - y = 0) = (x = (y::real))"
by arith
-lemma "((Numeral0::real) \<le> x + x) = (Numeral0 \<le> x)"
+lemma "((0::real) \<le> x + x) = (0 \<le> x)"
by arith
-lemma "(-x \<le> x) = ((Numeral0::real) \<le> x)"
+lemma "(-x \<le> x) = ((0::real) \<le> x)"
by arith
-lemma "(x \<le> -x) = (x \<le> (Numeral0::real))"
+lemma "(x \<le> -x) = (x \<le> (0::real))"
by arith
-lemma "(-x = (Numeral0::real)) = (x = Numeral0)"
+lemma "(-x = (0::real)) = (x = 0)"
by arith
lemma "-(x - y) = y - (x::real)"
by arith
-lemma "((Numeral0::real) < x - y) = (y < x)"
+lemma "((0::real) < x - y) = (y < x)"
by arith
-lemma "((Numeral0::real) \<le> x - y) = (y \<le> x)"
+lemma "((0::real) \<le> x - y) = (y \<le> x)"
by arith
lemma "(x + y) - x = (y::real)"
@@ -219,16 +219,16 @@
lemma "x < (y::real) ==> \<not>(x = y)"
by arith
-lemma "(x \<le> x + y) = ((Numeral0::real) \<le> y)"
+lemma "(x \<le> x + y) = ((0::real) \<le> y)"
by arith
-lemma "(y \<le> x + y) = ((Numeral0::real) \<le> x)"
+lemma "(y \<le> x + y) = ((0::real) \<le> x)"
by arith
-lemma "(x < x + y) = ((Numeral0::real) < y)"
+lemma "(x < x + y) = ((0::real) < y)"
by arith
-lemma "(y < x + y) = ((Numeral0::real) < x)"
+lemma "(y < x + y) = ((0::real) < x)"
by arith
lemma "(x - y) - x = (-y::real)"
@@ -258,10 +258,10 @@
lemma "(a + b) - (c + d) = (a - c) + (b - (d::real))"
by arith
-lemma "(Numeral0::real) - x = -x"
+lemma "(0::real) - x = -x"
by arith
-lemma "x - (Numeral0::real) = x"
+lemma "x - (0::real) = x"
by arith
lemma "w \<le> x \<and> y < z ==> w + y < x + (z::real)"
@@ -270,10 +270,10 @@
lemma "w < x \<and> y \<le> z ==> w + y < x + (z::real)"
by arith
-lemma "(Numeral0::real) \<le> x \<and> Numeral0 < y ==> Numeral0 < x + (y::real)"
+lemma "(0::real) \<le> x \<and> 0 < y ==> 0 < x + (y::real)"
by arith
-lemma "(Numeral0::real) < x \<and> Numeral0 \<le> y ==> Numeral0 < x + y"
+lemma "(0::real) < x \<and> 0 \<le> y ==> 0 < x + y"
by arith
lemma "-x - y = -(x + (y::real))"
@@ -303,7 +303,7 @@
lemma "x = y ==> x \<le> (y::real)"
by arith
-lemma "(Numeral0::real) < x ==> \<not>(x = Numeral0)"
+lemma "(0::real) < x ==> \<not>(x = 0)"
by arith
lemma "(x + y) * (x - y) = (x * x) - (y * y)"
--- a/src/HOL/Real/real_arith0.ML Thu Nov 01 21:12:13 2001 +0100
+++ b/src/HOL/Real/real_arith0.ML Fri Nov 02 17:55:24 2001 +0100
@@ -9,22 +9,22 @@
local
(* reduce contradictory <= to False *)
-val simps = [order_less_irrefl, zero_eq_numeral_0, one_eq_numeral_1,
- add_real_number_of, minus_real_number_of, diff_real_number_of,
- mult_real_number_of, eq_real_number_of, less_real_number_of,
- le_real_number_of_eq_not_less, real_diff_def,
- real_minus_add_distrib, real_minus_minus, real_mult_assoc];
-
-val add_rules =
- map rename_numerals
- [real_add_zero_left, real_add_zero_right,
- real_add_minus, real_add_minus_left,
- real_mult_0, real_mult_0_right,
- real_mult_1, real_mult_1_right,
- real_mult_minus_1, real_mult_minus_1_right];
+val add_rules =
+ [order_less_irrefl, real_numeral_0_eq_0, real_numeral_1_eq_1,
+ add_real_number_of, minus_real_number_of, diff_real_number_of,
+ mult_real_number_of, eq_real_number_of, less_real_number_of,
+ le_real_number_of_eq_not_less, real_diff_def,
+ real_minus_add_distrib, real_minus_minus, real_mult_assoc,
+ real_minus_zero,
+ real_add_zero_left, real_add_zero_right,
+ real_add_minus, real_add_minus_left,
+ real_mult_0, real_mult_0_right,
+ real_mult_1, real_mult_1_right,
+ real_mult_minus_1, real_mult_minus_1_right];
val simprocs = [Real_Times_Assoc.conv, Real_Numeral_Simprocs.combine_numerals]@
- Real_Numeral_Simprocs.cancel_numerals;
+ Real_Numeral_Simprocs.cancel_numerals @
+ Real_Numeral_Simprocs.eval_numerals;
val mono_ss = simpset() addsimps
[real_add_le_mono,real_add_less_mono,
@@ -66,14 +66,14 @@
mult_mono_thms = mult_mono_thms @ real_mult_mono_thms,
inj_thms = inj_thms, (*FIXME: add real*)
lessD = lessD, (*We don't change LA_Data_Ref.lessD because the real ordering is dense!*)
- simpset = simpset addsimps (add_rules @ simps)
+ simpset = simpset addsimps add_rules
addsimprocs simprocs}),
arith_discrete ("RealDef.real",false),
Simplifier.change_simpset_of (op addsimprocs) [fast_real_arith_simproc]];
(* some thms for injection nat => real:
real_of_nat_zero
-zero_eq_numeral_0
+?zero_eq_numeral_0
real_of_nat_add
*)