src/HOL/Real/RealInt.ML
changeset 9069 e8d530582061
parent 9043 ca761fe227d8
child 9365 0cced1b20d68
--- a/src/HOL/Real/RealInt.ML	Wed Jun 14 18:19:20 2000 +0200
+++ b/src/HOL/Real/RealInt.ML	Wed Jun 14 18:21:25 2000 +0200
@@ -10,8 +10,9 @@
   "congruent intrel (%p. split (%i j. realrel ^^ \
 \  {(preal_of_prat (prat_of_pnat (pnat_of_nat i)), \
 \    preal_of_prat (prat_of_pnat (pnat_of_nat j)))}) p)";
-by (auto_tac (claset(),simpset() addsimps [pnat_of_nat_add,
-    prat_of_pnat_add RS sym,preal_of_prat_add RS sym]));
+by (auto_tac (claset(),
+	      simpset() addsimps [pnat_of_nat_add, prat_of_pnat_add RS sym,
+				  preal_of_prat_add RS sym]));
 qed "real_of_int_congruent";
 
 val real_of_int_ize = RSLIST [equiv_intrel, real_of_int_congruent];
@@ -22,9 +23,8 @@
 \       {(preal_of_prat (prat_of_pnat (pnat_of_nat i)), \
 \         preal_of_prat (prat_of_pnat (pnat_of_nat j)))})";
 by (res_inst_tac [("f","Abs_real")] arg_cong 1);
-by (simp_tac (simpset() addsimps 
-   [realrel_in_real RS Abs_real_inverse,
-    real_of_int_ize UN_equiv_class]) 1);
+by (simp_tac (simpset() addsimps [realrel_in_real RS Abs_real_inverse,
+				  real_of_int_ize UN_equiv_class]) 1);
 qed "real_of_int";
 
 Goal "inj(real_of_int)";
@@ -32,9 +32,9 @@
 by (res_inst_tac [("z","x")] eq_Abs_Integ 1);
 by (res_inst_tac [("z","y")] eq_Abs_Integ 1);
 by (auto_tac (claset() addSDs [inj_preal_of_prat RS injD,
-    inj_prat_of_pnat RS injD,inj_pnat_of_nat RS injD],
-    simpset() addsimps [real_of_int,preal_of_prat_add RS sym,
-     prat_of_pnat_add RS sym,pnat_of_nat_add]));
+			   inj_prat_of_pnat RS injD, inj_pnat_of_nat RS injD],
+	      simpset() addsimps [real_of_int,preal_of_prat_add RS sym,
+				  prat_of_pnat_add RS sym,pnat_of_nat_add]));
 qed "inj_real_of_int";
 
 Goalw [int_def,real_zero_def] "real_of_int (int 0) = 0";
@@ -58,14 +58,12 @@
 
 Goal "-real_of_int x = real_of_int (-x)";
 by (res_inst_tac [("z","x")] eq_Abs_Integ 1);
-by (auto_tac (claset(),simpset() addsimps [real_of_int,
-    real_minus,zminus]));
+by (auto_tac (claset(), simpset() addsimps [real_of_int, real_minus,zminus]));
 qed "real_of_int_minus";
 
 Goalw [zdiff_def,real_diff_def]
   "real_of_int x - real_of_int y = real_of_int (x - y)";
-by (simp_tac (simpset() addsimps [real_of_int_add,
-    real_of_int_minus]) 1);
+by (simp_tac (simpset() addsimps [real_of_int_add, real_of_int_minus]) 1);
 qed "real_of_int_diff";
 
 Goal "real_of_int x * real_of_int y = real_of_int (x * y)";
@@ -93,19 +91,12 @@
 
 Goal "~neg z ==> real_of_nat (nat z) = real_of_int z";
 by (asm_simp_tac (simpset() addsimps [not_neg_nat,
-    real_of_int_real_of_nat RS sym]) 1);
+				      real_of_int_real_of_nat RS sym]) 1);
 qed "real_of_nat_real_of_int";
 
-(* put with other properties of real_of_nat? *)
-Goal "neg z ==> real_of_nat (nat z) = 0";
-by (asm_simp_tac (simpset() addsimps [neg_nat,
-    real_of_nat_zero]) 1);
-qed "real_of_nat_neg_int";
-Addsimps [real_of_nat_neg_int];
-
 Goal "(real_of_int x = 0) = (x = int 0)";
 by (auto_tac (claset() addIs [inj_real_of_int RS injD],
-    HOL_ss addsimps [real_of_int_zero]));
+	      HOL_ss addsimps [real_of_int_zero]));
 qed "real_of_int_zero_cancel";
 Addsimps [real_of_int_zero_cancel];
 
@@ -114,33 +105,27 @@
 by (auto_tac (claset(),
 	      simpset() addsimps [zle_iff_zadd, real_of_int_add RS sym,
 				  real_of_int_real_of_nat,
-				  real_of_nat_zero RS sym]));
+				  linorder_not_le RS sym]));
 qed "real_of_int_less_cancel";
 
+Goal "(real_of_int x = real_of_int y) = (x = y)";
+by (blast_tac (claset() addSDs [inj_real_of_int RS injD]) 1);
+qed "real_of_int_eq_iff";
+AddIffs [real_of_int_eq_iff];
+
 Goal "x < y ==> (real_of_int x < real_of_int y)";
-by (auto_tac (claset(),
-	      HOL_ss addsimps [zless_iff_Suc_zadd, real_of_int_add RS sym,
-				  real_of_int_real_of_nat,
-				  real_of_nat_Suc]));
-by (simp_tac (simpset() addsimps [real_of_nat_one RS
-    sym,real_of_nat_zero RS sym,real_of_nat_add]) 1); 
+by (full_simp_tac (simpset() addsimps [linorder_not_le RS sym]) 1);
+by (auto_tac (claset() addSDs [real_of_int_less_cancel],
+	      simpset() addsimps [order_le_less]));
 qed "real_of_int_less_mono";
 
 Goal "(real_of_int x < real_of_int y) = (x < y)";
-by (auto_tac (claset() addIs [real_of_int_less_cancel,
-			      real_of_int_less_mono],
-	      simpset()));
+by (blast_tac (claset() addIs [real_of_int_less_cancel,
+			       real_of_int_less_mono]) 1);
 qed "real_of_int_less_iff";
-Addsimps [real_of_int_less_iff];
+AddIffs [real_of_int_less_iff];
 
 Goal "(real_of_int x <= real_of_int y) = (x <= y)";
-by (auto_tac (claset(),
-	      simpset() addsimps [real_le_def, zle_def]));
+by (full_simp_tac (simpset() addsimps [linorder_not_less RS sym]) 1);
 qed "real_of_int_le_iff";
 Addsimps [real_of_int_le_iff];
-
-Goal "(real_of_int x = real_of_int y) = (x = y)";
-by (auto_tac (claset() addSIs [order_antisym],
-	      simpset() addsimps [real_of_int_le_iff RS iffD1]));
-qed "real_of_int_eq_iff";
-Addsimps [real_of_int_eq_iff];