--- 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];
--- a/src/HOL/Real/RealOrd.ML Wed Jun 14 18:19:20 2000 +0200
+++ b/src/HOL/Real/RealOrd.ML Wed Jun 14 18:21:25 2000 +0200
@@ -825,64 +825,65 @@
Goal "((0::real) < x*y) = (0 < x & 0 < y | x < 0 & y < 0)";
by (auto_tac (claset(),
- simpset() addsimps [order_le_less, real_less_le_iff,
+ simpset() addsimps [order_le_less, linorder_not_less,
real_mult_order, real_mult_less_zero1]));
by (ALLGOALS (rtac ccontr));
-by (auto_tac (claset(), simpset() addsimps [order_le_less, real_less_le_iff]));
+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 (auto_tac (claset() addDs [order_less_not_sym],
simpset() addsimps [real_mult_commute]));
-qed "real_zero_less_times_iff";
+qed "real_zero_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, real_less_le_iff,
- real_zero_less_times_iff]));
-qed "real_zero_le_times_iff";
+ simpset() addsimps [order_le_less, linorder_not_less,
+ real_zero_less_mult_iff]));
+qed "real_zero_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_times_iff,
+ simpset() addsimps [real_zero_le_mult_iff,
linorder_not_le RS sym]));
by (auto_tac (claset() addDs [order_less_not_sym],
simpset() addsimps [linorder_not_le]));
-qed "real_times_less_zero_iff";
+qed "real_mult_less_zero_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_times_iff,
+ simpset() addsimps [real_zero_less_mult_iff,
linorder_not_less RS sym]));
-qed "real_times_le_zero_iff";
+qed "real_mult_le_zero_iff";
(*----------------------------------------------------------------------------
Another embedding of the naturals in the reals (see real_of_posnat)
----------------------------------------------------------------------------*)
Goalw [real_of_nat_def] "real_of_nat 0 = 0";
-by (full_simp_tac (simpset() addsimps [real_of_posnat_one]) 1);
+by (simp_tac (simpset() addsimps [real_of_posnat_one]) 1);
qed "real_of_nat_zero";
Goalw [real_of_nat_def] "real_of_nat 1 = 1r";
-by (full_simp_tac (simpset() addsimps [real_of_posnat_two,
- real_add_assoc]) 1);
+by (simp_tac (simpset() addsimps [real_of_posnat_two, real_add_assoc]) 1);
qed "real_of_nat_one";
+Addsimps [real_of_nat_zero, real_of_nat_one];
Goalw [real_of_nat_def]
- "real_of_nat n1 + real_of_nat n2 = real_of_nat (n1 + n2)";
+ "real_of_nat (m + n) = real_of_nat m + real_of_nat n";
by (simp_tac (simpset() addsimps
- [real_of_posnat_add,real_add_assoc RS sym]) 1);
+ [real_of_posnat_add,real_add_assoc RS sym]) 1);
qed "real_of_nat_add";
Goalw [real_of_nat_def] "real_of_nat (Suc n) = real_of_nat n + 1r";
by (simp_tac (simpset() addsimps [real_of_posnat_Suc] @ real_add_ac) 1);
qed "real_of_nat_Suc";
+Addsimps [real_of_nat_Suc];
-Goalw [real_of_nat_def] "(n < m) = (real_of_nat n < real_of_nat m)";
+Goalw [real_of_nat_def] "(real_of_nat n < real_of_nat m) = (n < m)";
by Auto_tac;
qed "real_of_nat_less_iff";
-Addsimps [real_of_nat_less_iff RS sym];
+AddIffs [real_of_nat_less_iff];
Goal "inj real_of_nat";
by (rtac injI 1);
@@ -894,14 +895,12 @@
by (res_inst_tac [("C","1r")] real_le_add_right_cancel 1);
by (asm_full_simp_tac (simpset() addsimps [real_add_assoc]) 1);
qed "real_of_nat_ge_zero";
-Addsimps [real_of_nat_ge_zero];
+AddIffs [real_of_nat_ge_zero];
-Goal "real_of_nat n1 * real_of_nat n2 = real_of_nat (n1 * n2)";
-by (induct_tac "n1" 1);
-by (dtac sym 2);
+Goal "real_of_nat (m * n) = real_of_nat m * real_of_nat n";
+by (induct_tac "m" 1);
by (auto_tac (claset(),
- simpset() addsimps [real_of_nat_zero,
- real_of_nat_add RS sym,real_of_nat_Suc,
+ simpset() addsimps [real_of_nat_add,
real_add_mult_distrib, real_add_commute]));
qed "real_of_nat_mult";
@@ -910,33 +909,36 @@
simpset()));
qed "real_of_nat_eq_cancel";
-Goal "n2 <= n1 --> real_of_nat (n1 - n2) = real_of_nat n1 + (-real_of_nat n2)";
-by (induct_tac "n1" 1);
+Goal "n <= m --> real_of_nat (m - n) = real_of_nat m + (-real_of_nat n)";
+by (induct_tac "m" 1);
by (auto_tac (claset(),
simpset() addsimps [Suc_diff_le, le_Suc_eq, real_of_nat_Suc,
real_of_nat_zero] @ real_add_ac));
qed_spec_mp "real_of_nat_minus";
(* 05/00 *)
-Goal "n2 < n1 ==> real_of_nat (n1 - n2) = \
-\ real_of_nat n1 + -real_of_nat n2";
+Goal "n < m ==> real_of_nat (m - n) = \
+\ real_of_nat m + -real_of_nat n";
by (auto_tac (claset() addIs [real_of_nat_minus],simpset()));
qed "real_of_nat_minus2";
-Goalw [real_diff_def] "n2 < n1 ==> real_of_nat (n1 - n2) = \
-\ real_of_nat n1 - real_of_nat n2";
+Goalw [real_diff_def]
+ "n < m ==> real_of_nat (m - n) = real_of_nat m - real_of_nat n";
by (etac real_of_nat_minus2 1);
qed "real_of_nat_diff";
-Goalw [real_diff_def] "n2 <= n1 ==> real_of_nat (n1 - n2) = \
-\ real_of_nat n1 - real_of_nat n2";
+Goalw [real_diff_def]
+ "n <= m ==> real_of_nat (m - n) = real_of_nat m - real_of_nat n";
by (etac real_of_nat_minus 1);
qed "real_of_nat_diff2";
-Goal "(real_of_nat n ~= 0) = (n ~= 0)";
-by (auto_tac (claset() addIs [inj_real_of_nat RS injD],
- simpset() addsimps [real_of_nat_zero RS sym]
- delsimps [neq0_conv]));
-qed "real_of_nat_not_zero_iff";
-Addsimps [real_of_nat_not_zero_iff];
+Goal "(real_of_nat n = 0) = (n = 0)";
+by (auto_tac (claset() addIs [inj_real_of_nat RS injD], simpset()));
+qed "real_of_nat_zero_iff";
+AddIffs [real_of_nat_zero_iff];
+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];
+