src/HOL/Real/Hyperreal/HRealAbs.ML
changeset 10715 c838477b5c18
parent 10699 f0c3da8477e9
--- a/src/HOL/Real/Hyperreal/HRealAbs.ML	Thu Dec 21 10:16:07 2000 +0100
+++ b/src/HOL/Real/Hyperreal/HRealAbs.ML	Thu Dec 21 10:16:33 2000 +0100
@@ -5,17 +5,16 @@
                   Similar to RealAbs.thy
 *) 
 
-
 (*------------------------------------------------------------
   absolute value on hyperreals as pointwise operation on 
   equivalence class representative
  ------------------------------------------------------------*)
 
 Goalw [hrabs_def]
-"abs (Abs_hypreal (hyprel ^^ {X})) = \
-\            Abs_hypreal(hyprel ^^ {%n. abs (X n)})";
-by (auto_tac (claset(),simpset() addsimps [hypreal_zero_def,
-    hypreal_le,hypreal_minus]));
+     "abs (Abs_hypreal (hyprel ^^ {X})) = \
+\     Abs_hypreal(hyprel ^^ {%n. abs (X n)})";
+by (auto_tac (claset(),
+           simpset() addsimps [hypreal_zero_def, hypreal_le,hypreal_minus]));
 by (ALLGOALS(Ultra_tac THEN' arith_tac ));
 qed "hypreal_hrabs";
 
@@ -23,49 +22,38 @@
    Properties of the absolute value function over the reals
    (adapted version of previously proved theorems about abs)
  ------------------------------------------------------------*)
-Goalw [hrabs_def] "abs r = (if (0::hypreal)<=r then r else -r)";
-by (Step_tac 1);
-qed "hrabs_iff";
 
-Goalw [hrabs_def] "abs (0::hypreal) = (0::hypreal)";
-by (rtac (hypreal_le_refl RS if_P) 1);
+Goal "abs (0::hypreal) = 0";
+by (simp_tac (simpset() addsimps [hrabs_def]) 1); 
 qed "hrabs_zero";
-
 Addsimps [hrabs_zero];
 
-Goalw [hrabs_def] "abs (0::hypreal) = -(0::hypreal)";
-by (rtac (hypreal_minus_zero RS ssubst) 1);
-by (rtac if_cancel 1);
-qed "hrabs_minus_zero";
-
-val [prem] = goalw thy [hrabs_def] "(0::hypreal)<=x ==> abs x = x";
-by (rtac (prem RS if_P) 1);
+Goal "(0::hypreal)<=x ==> abs x = x";
+by (asm_simp_tac (simpset() addsimps [hrabs_def]) 1); 
 qed "hrabs_eqI1";
 
-val [prem] = goalw thy [hrabs_def] "(0::hypreal)<x ==> abs x = x";
-by (simp_tac (simpset() addsimps [(prem 
-    RS hypreal_less_imp_le),hrabs_eqI1]) 1);
+Goal "(0::hypreal)<x ==> abs x = x";
+by (asm_simp_tac (simpset() addsimps [hypreal_less_imp_le, hrabs_eqI1]) 1);
 qed "hrabs_eqI2";
 
-val [prem] = goalw thy [hrabs_def,hypreal_le_def] 
-    "x<(0::hypreal) ==> abs x = -x";
-by (simp_tac (simpset() addsimps [prem,if_not_P]) 1);
+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. x<=(0::hypreal) ==> abs x = -x";
-by (dtac hypreal_le_imp_less_or_eq 1);
-by (fast_tac (HOL_cs addIs [hrabs_minus_zero,
-    hrabs_minus_eqI2]) 1);
+Goal "x<=(0::hypreal) ==> abs x = -x";
+by (auto_tac (claset() addDs [order_antisym], simpset() addsimps [hrabs_def])); 
 qed "hrabs_minus_eqI1";
 
-Goalw [hrabs_def,hypreal_le_def] "(0::hypreal)<= abs x";
-by (auto_tac (claset() addDs [hypreal_minus_zero_less_iff RS iffD2,
-    hypreal_less_asym],simpset()));
+Goal "(0::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]));
 qed "hrabs_ge_zero";
 
 Goal "abs(abs x) = abs (x::hypreal)";
-by (res_inst_tac [("r1","abs x")] (hrabs_iff RS ssubst) 1);
-by (blast_tac (claset() addIs [if_P,hrabs_ge_zero]) 1);
+by (auto_tac (claset() addDs [hypreal_minus_zero_less_iff RS iffD2, 
+                              hypreal_less_asym], 
+              simpset() addsimps [hypreal_le_def, hrabs_def]));
 qed "hrabs_idempotent";
 Addsimps [hrabs_idempotent];
 
@@ -87,8 +75,8 @@
 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_mult,abs_mult]));
+by (auto_tac (claset(), 
+              simpset() addsimps [hypreal_hrabs, hypreal_mult,abs_mult]));
 qed "hrabs_mult";
 
 Goal "abs(inverse(x)) = inverse(abs(x::hypreal))";
@@ -103,31 +91,23 @@
 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]));
+by (auto_tac (claset(), 
+      simpset() addsimps [hypreal_hrabs, hypreal_add,hypreal_le,
+                        abs_triangle_ineq]));
 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 hypreal_le_trans),hypreal_add_left_le_mono1],
+by (auto_tac (claset() addSIs [hrabs_triangle_ineq RS hypreal_le_trans,
+                               hypreal_add_left_le_mono1],
     simpset() addsimps [hypreal_add_assoc]));
 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]));
+by (auto_tac (claset() addSDs [not_hypreal_leE, hypreal_less_asym] 
+                       addIs [hypreal_le_anti_sym],
+              simpset() addsimps [hypreal_ge_zero_iff]));
 qed "hrabs_minus_cancel";
-
-Goal "abs((x::hypreal) + -y) = abs (y + -x)";
-by (rtac (hrabs_minus_cancel RS subst) 1);
-by (simp_tac (simpset() addsimps [hypreal_add_commute]) 1);
-qed "hrabs_minus_add_cancel";
-
-Goal "abs((x::hypreal) + -y) <= abs x + abs y";
-by (res_inst_tac [("x1","y")] (hrabs_minus_cancel RS subst) 1);
-by (rtac hrabs_triangle_ineq 1);
-qed "rhabs_triangle_minus_ineq";
+Addsimps [hrabs_minus_cancel];
 
 val prem1::prem2::rest = goal thy 
     "[| abs x < r; abs y < s |] ==> abs(x+y) < r + (s::hypreal)";
@@ -136,13 +116,6 @@
 by (rtac ([prem1,prem2] MRS hypreal_add_less_mono) 1);
 qed "hrabs_add_less";
 
-Goal "[| abs x < r; abs y < s |] \
-\     ==> abs(x+ -y) < r + (s::hypreal)";
-by (rotate_tac 1 1);
-by (dtac (hrabs_minus_cancel RS ssubst) 1);
-by (asm_simp_tac (simpset() addsimps [hrabs_add_less]) 1);
-qed "hrabs_add_minus_less";
-
 val prem1::prem2::rest = 
     goal thy "[| abs x<r; abs y<s |] ==> abs(x*y)<r*(s::hypreal)";
 by (simp_tac (simpset() addsimps [hrabs_mult]) 1);
@@ -153,11 +126,11 @@
 by (rtac (prem2 RS (hrabs_ge_zero RS hypreal_le_less_trans)) 1);
 by (rtac prem1 1);
 by (rtac ([prem1 RS (hrabs_ge_zero RS hypreal_le_less_trans),
-   prem2 RS (hrabs_ge_zero RS hypreal_le_less_trans)] 
-   MRS hypreal_mult_order) 1);
+           prem2 RS (hrabs_ge_zero RS hypreal_le_less_trans)] 
+          MRS hypreal_mult_order) 1);
 qed "hrabs_mult_less";
 
-Goal "!! x y r. 1hr < abs x ==> abs y <= abs(x*y)";
+Goal "1hr < abs x ==> abs y <= abs(x*y)";
 by (cut_inst_tac [("x1","y")] (hrabs_ge_zero RS hypreal_le_imp_less_or_eq) 1);
 by (EVERY1[etac disjE,rtac hypreal_less_imp_le]);
 by (dres_inst_tac [("x1","1hr")]  (hypreal_less_minus_iff RS iffD1) 1);
@@ -170,30 +143,24 @@
 by (asm_full_simp_tac (simpset() addsimps [hypreal_le_refl,hrabs_mult]) 1);
 qed "hrabs_mult_le";
 
-Goal "!!x. [| 1hr < abs x; r < abs y|] ==> r < abs(x*y)";
+Goal "[| 1hr < abs x; r < abs y|] ==> r < abs(x*y)";
 by (fast_tac (HOL_cs addIs [hrabs_mult_le, hypreal_less_le_trans]) 1);
 qed "hrabs_mult_gt";
 
-Goal "!!r. abs x < r ==> (0::hypreal) < r";
+Goal "abs x < r ==> (0::hypreal) < r";
 by (blast_tac (claset() addSIs [hypreal_le_less_trans,
     hrabs_ge_zero]) 1);
 qed "hrabs_less_gt_zero";
 
 Goalw [hrabs_def] "abs 1hr = 1hr";
-by (auto_tac (claset() addSDs [not_hypreal_leE 
-    RS hypreal_less_asym],simpset() addsimps 
-    [hypreal_zero_less_one]));
+by (auto_tac (claset() addSDs [not_hypreal_leE RS hypreal_less_asym], 
+      simpset() addsimps [hypreal_zero_less_one]));
 qed "hrabs_one";
 
-val prem1::prem2::rest = 
-    goal thy "[| (0::hypreal) < x ; x < r |] ==> abs x < r";
-by (simp_tac (simpset() addsimps [(prem1 RS hrabs_eqI2),prem2]) 1);
-qed "hrabs_lessI";
-
 Goal "abs x = (x::hypreal) | abs x = -x";
 by (cut_inst_tac [("x","0"),("y","x")] hypreal_linear 1);
 by (fast_tac (claset() addIs [hrabs_eqI2,hrabs_minus_eqI2,
-                            hrabs_zero,hrabs_minus_zero]) 1);
+                            hrabs_zero]) 1);
 qed "hrabs_disj";
 
 Goal "abs x = (y::hypreal) ==> x = y | -x = y";
@@ -217,54 +184,28 @@
 qed "hrabs_interval_iff";
 
 Goal "(abs x < (r::hypreal)) = (- x < r & x < r)";
-by (auto_tac (claset(),simpset() addsimps [hrabs_interval_iff]));
+by (auto_tac (claset(),  simpset() addsimps [hrabs_interval_iff]));
 by (dtac (hypreal_less_swap_iff RS iffD1) 1);
 by (dtac (hypreal_less_swap_iff RS iffD1) 2);
 by (Auto_tac);
 qed "hrabs_interval_iff2";
 
-Goal 
-     "(abs (x + -y) < (r::hypreal)) = (y + -r < x & x < y + r)";
-by (auto_tac (claset(),simpset() addsimps 
-     [hrabs_interval_iff]));
-by (ALLGOALS(dtac (hypreal_less_minus_iff RS iffD1)));
-by (ALLGOALS(dtac (hypreal_less_minus_iff RS iffD1)));
-by (ALLGOALS(rtac (hypreal_less_minus_iff RS iffD2)));
-by (auto_tac (claset(),simpset() addsimps 
-    [hypreal_minus_add_distrib] addsimps hypreal_add_ac));
-qed "hrabs_add_minus_interval_iff";
-
-Goal "x < (y::hypreal) ==> abs(y + -x) = y + -x";
-by (dtac (hypreal_less_minus_iff RS iffD1) 1);
-by (etac hrabs_eqI2 1);
-qed "hrabs_less_eqI2";
-
-Goal "x < (y::hypreal) ==> abs(x + -y) = y + -x";
-by (auto_tac (claset() addDs [hrabs_less_eqI2],
-              simpset() addsimps [hrabs_minus_add_cancel]));
-qed "hrabs_less_eqI2a";
-
-Goal "x <= (y::hypreal) ==> abs(y + -x) = y + -x";
-by (auto_tac (claset() addDs  [hypreal_le_imp_less_or_eq,
-              hrabs_less_eqI2],simpset()));
-qed "hrabs_le_eqI2";
-
-Goal "x <= (y::hypreal) ==> abs(x + -y) = y + -x";
-by (auto_tac (claset() addDs [hrabs_le_eqI2],
-              simpset() addsimps [hrabs_minus_add_cancel]));
-qed "hrabs_le_eqI2a";
-
 (* Needed in Geom.ML *)
-Goal "(y::hypreal) + - x + (y + - z) = abs (x + - z) \
-\     ==> y = z | x = y";
+Goal "(y::hypreal) + - x + (y + - z) = abs (x + - z) ==> y = z | x = y";
 by (res_inst_tac [("z","x")] eq_Abs_hypreal 1);
 by (res_inst_tac [("z","y")] eq_Abs_hypreal 1);
 by (res_inst_tac [("z","z")] eq_Abs_hypreal 1);
-by (auto_tac (claset(),simpset() addsimps [hypreal_hrabs,
-    hypreal_minus,hypreal_add]));
+by (auto_tac (claset(), 
+              simpset() addsimps [hypreal_hrabs, hypreal_minus,hypreal_add]));
 by (Ultra_tac 1 THEN arith_tac 1);
 qed "hrabs_add_lemma_disj";
 
+(***SHOULD BE TRIVIAL GIVEN SIMPROCS
+Goal "abs((x::hypreal) + -y) = abs (y + -x)";
+by (rtac (hrabs_minus_cancel RS subst) 1);
+by (simp_tac (simpset() addsimps [hypreal_add_commute]) 1);
+qed "hrabs_minus_add_cancel";
+
 (* Needed in Geom.ML *)
 Goal "(x::hypreal) + - y + (z + - y) = abs (x + - z) \
 \     ==> y = z | x = y";
@@ -274,11 +215,12 @@
 by (asm_full_simp_tac (simpset() addsimps [hrabs_minus_add_cancel] 
          @ hypreal_add_ac) 1);
 qed "hrabs_add_lemma_disj2";
+***)
  
 (*----------------------------------------------------------
     Relating hrabs to abs through embedding of IR into IR*
  ----------------------------------------------------------*)
 Goalw [hypreal_of_real_def] 
     "abs (hypreal_of_real r) = hypreal_of_real (abs r)";
-by (auto_tac (claset(),simpset() addsimps [hypreal_hrabs]));
+by (auto_tac (claset(), simpset() addsimps [hypreal_hrabs]));
 qed "hypreal_of_real_hrabs";