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