| author | paulson | 
| Thu, 14 Sep 2000 11:34:13 +0200 | |
| changeset 9955 | 6ed42bcba707 | 
| parent 9391 | a6ab3a442da6 | 
| child 10784 | 27e4d90b35b5 | 
| permissions | -rw-r--r-- | 
| 7292 | 1 | (* Title: RealInt.ML | 
| 2 | ID: $Id$ | |
| 3 | Author: Jacques D. Fleuriot | |
| 4 | Copyright: 1999 University of Edinburgh | |
| 5 | Description: Embedding the integers in the reals | |
| 6 | *) | |
| 7 | ||
| 8 | ||
| 9 | Goalw [congruent_def] | |
| 9365 | 10 | "congruent intrel (%p. (%(i,j). realrel ^^ \ | 
| 7292 | 11 | \  {(preal_of_prat (prat_of_pnat (pnat_of_nat i)), \
 | 
| 12 | \ preal_of_prat (prat_of_pnat (pnat_of_nat j)))}) p)"; | |
| 9069 | 13 | by (auto_tac (claset(), | 
| 14 | simpset() addsimps [pnat_of_nat_add, prat_of_pnat_add RS sym, | |
| 15 | preal_of_prat_add RS sym])); | |
| 7292 | 16 | qed "real_of_int_congruent"; | 
| 17 | ||
| 18 | Goalw [real_of_int_def] | |
| 19 |    "real_of_int (Abs_Integ (intrel ^^ {(i, j)})) = \
 | |
| 20 | \ Abs_real(realrel ^^ \ | |
| 21 | \       {(preal_of_prat (prat_of_pnat (pnat_of_nat i)), \
 | |
| 22 | \ preal_of_prat (prat_of_pnat (pnat_of_nat j)))})"; | |
| 23 | by (res_inst_tac [("f","Abs_real")] arg_cong 1);
 | |
| 9069 | 24 | by (simp_tac (simpset() addsimps [realrel_in_real RS Abs_real_inverse, | 
| 9391 | 25 | [equiv_intrel, real_of_int_congruent] MRS UN_equiv_class]) 1); | 
| 7292 | 26 | qed "real_of_int"; | 
| 27 | ||
| 28 | Goal "inj(real_of_int)"; | |
| 29 | by (rtac injI 1); | |
| 30 | by (res_inst_tac [("z","x")] eq_Abs_Integ 1);
 | |
| 31 | by (res_inst_tac [("z","y")] eq_Abs_Integ 1);
 | |
| 32 | by (auto_tac (claset() addSDs [inj_preal_of_prat RS injD, | |
| 9069 | 33 | inj_prat_of_pnat RS injD, inj_pnat_of_nat RS injD], | 
| 34 | simpset() addsimps [real_of_int,preal_of_prat_add RS sym, | |
| 35 | prat_of_pnat_add RS sym,pnat_of_nat_add])); | |
| 7292 | 36 | qed "inj_real_of_int"; | 
| 37 | ||
| 9043 
ca761fe227d8
First round of changes, towards installation of simprocs
 paulson parents: 
8856diff
changeset | 38 | Goalw [int_def,real_zero_def] "real_of_int (int 0) = 0"; | 
| 7292 | 39 | by (simp_tac (simpset() addsimps [real_of_int, preal_add_commute]) 1); | 
| 40 | qed "real_of_int_zero"; | |
| 41 | ||
| 42 | Goalw [int_def,real_one_def] "real_of_int (int 1) = 1r"; | |
| 43 | by (auto_tac (claset(), | |
| 44 | simpset() addsimps [real_of_int, | |
| 45 | preal_of_prat_add RS sym,pnat_two_eq, | |
| 46 | prat_of_pnat_add RS sym,pnat_one_iff RS sym])); | |
| 47 | qed "real_of_int_one"; | |
| 48 | ||
| 49 | Goal "real_of_int x + real_of_int y = real_of_int (x + y)"; | |
| 50 | by (res_inst_tac [("z","x")] eq_Abs_Integ 1);
 | |
| 51 | by (res_inst_tac [("z","y")] eq_Abs_Integ 1);
 | |
| 52 | by (auto_tac (claset(),simpset() addsimps [real_of_int, | |
| 53 | preal_of_prat_add RS sym,prat_of_pnat_add RS sym, | |
| 54 | zadd,real_add,pnat_of_nat_add] @ pnat_add_ac)); | |
| 55 | qed "real_of_int_add"; | |
| 56 | ||
| 57 | Goal "-real_of_int x = real_of_int (-x)"; | |
| 58 | by (res_inst_tac [("z","x")] eq_Abs_Integ 1);
 | |
| 9069 | 59 | by (auto_tac (claset(), simpset() addsimps [real_of_int, real_minus,zminus])); | 
| 7292 | 60 | qed "real_of_int_minus"; | 
| 61 | ||
| 62 | Goalw [zdiff_def,real_diff_def] | |
| 63 | "real_of_int x - real_of_int y = real_of_int (x - y)"; | |
| 9069 | 64 | by (simp_tac (simpset() addsimps [real_of_int_add, real_of_int_minus]) 1); | 
| 7292 | 65 | qed "real_of_int_diff"; | 
| 66 | ||
| 67 | Goal "real_of_int x * real_of_int y = real_of_int (x * y)"; | |
| 68 | by (res_inst_tac [("z","x")] eq_Abs_Integ 1);
 | |
| 69 | by (res_inst_tac [("z","y")] eq_Abs_Integ 1);
 | |
| 70 | by (auto_tac (claset(),simpset() addsimps [real_of_int, | |
| 71 | real_mult,zmult,preal_of_prat_mult RS sym,pnat_of_nat_mult, | |
| 72 | prat_of_pnat_mult RS sym,preal_of_prat_add RS sym, | |
| 73 | prat_of_pnat_add RS sym,pnat_of_nat_add] @ mult_ac @add_ac | |
| 74 | @ pnat_add_ac)); | |
| 75 | qed "real_of_int_mult"; | |
| 76 | ||
| 77 | Goal "real_of_int (int (Suc n)) = real_of_int (int n) + 1r"; | |
| 78 | by (simp_tac (simpset() addsimps [real_of_int_one RS sym, | |
| 8856 | 79 | real_of_int_add,zadd_int] @ zadd_ac) 1); | 
| 7292 | 80 | qed "real_of_int_Suc"; | 
| 81 | ||
| 82 | (* relating two of the embeddings *) | |
| 83 | Goal "real_of_int (int n) = real_of_nat n"; | |
| 84 | by (induct_tac "n" 1); | |
| 8856 | 85 | by (ALLGOALS (simp_tac (HOL_ss addsimps [real_of_int_zero, | 
| 86 | real_of_nat_zero,real_of_int_Suc,real_of_nat_Suc]))); | |
| 87 | by (Simp_tac 1); | |
| 7292 | 88 | qed "real_of_int_real_of_nat"; | 
| 89 | ||
| 90 | Goal "~neg z ==> real_of_nat (nat z) = real_of_int z"; | |
| 91 | by (asm_simp_tac (simpset() addsimps [not_neg_nat, | |
| 9069 | 92 | real_of_int_real_of_nat RS sym]) 1); | 
| 7292 | 93 | qed "real_of_nat_real_of_int"; | 
| 94 | ||
| 9043 
ca761fe227d8
First round of changes, towards installation of simprocs
 paulson parents: 
8856diff
changeset | 95 | Goal "(real_of_int x = 0) = (x = int 0)"; | 
| 7292 | 96 | by (auto_tac (claset() addIs [inj_real_of_int RS injD], | 
| 9069 | 97 | HOL_ss addsimps [real_of_int_zero])); | 
| 7292 | 98 | qed "real_of_int_zero_cancel"; | 
| 99 | Addsimps [real_of_int_zero_cancel]; | |
| 100 | ||
| 101 | Goal "real_of_int x < real_of_int y ==> x < y"; | |
| 102 | by (rtac ccontr 1 THEN dtac (linorder_not_less RS iffD1) 1); | |
| 103 | by (auto_tac (claset(), | |
| 104 | simpset() addsimps [zle_iff_zadd, real_of_int_add RS sym, | |
| 105 | real_of_int_real_of_nat, | |
| 9069 | 106 | linorder_not_le RS sym])); | 
| 7292 | 107 | qed "real_of_int_less_cancel"; | 
| 108 | ||
| 9069 | 109 | Goal "(real_of_int x = real_of_int y) = (x = y)"; | 
| 110 | by (blast_tac (claset() addSDs [inj_real_of_int RS injD]) 1); | |
| 111 | qed "real_of_int_eq_iff"; | |
| 112 | AddIffs [real_of_int_eq_iff]; | |
| 113 | ||
| 7292 | 114 | Goal "x < y ==> (real_of_int x < real_of_int y)"; | 
| 9069 | 115 | by (full_simp_tac (simpset() addsimps [linorder_not_le RS sym]) 1); | 
| 116 | by (auto_tac (claset() addSDs [real_of_int_less_cancel], | |
| 117 | simpset() addsimps [order_le_less])); | |
| 7292 | 118 | qed "real_of_int_less_mono"; | 
| 119 | ||
| 120 | Goal "(real_of_int x < real_of_int y) = (x < y)"; | |
| 9069 | 121 | by (blast_tac (claset() addIs [real_of_int_less_cancel, | 
| 122 | real_of_int_less_mono]) 1); | |
| 7292 | 123 | qed "real_of_int_less_iff"; | 
| 9069 | 124 | AddIffs [real_of_int_less_iff]; | 
| 7292 | 125 | |
| 126 | Goal "(real_of_int x <= real_of_int y) = (x <= y)"; | |
| 9069 | 127 | by (full_simp_tac (simpset() addsimps [linorder_not_less RS sym]) 1); | 
| 7292 | 128 | qed "real_of_int_le_iff"; | 
| 129 | Addsimps [real_of_int_le_iff]; |