src/HOL/Real/RealInt.ML
author wenzelm
Mon, 08 Oct 2001 15:23:20 +0200
changeset 11713 883d559b0b8c
parent 11597 cd6d2eddf75f
child 12613 279facb4253a
permissions -rw-r--r--
sane numerals (stage 3): provide generic "1" on all number types;
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]
10834
a7897aebbffc *** empty log message ***
nipkow
parents: 10797
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
Goalw [real_of_int_def]
10919
144ede948e58 renamings: real_of_nat, real_of_int -> (overloaded) real
paulson
parents: 10834
diff changeset
    19
   "real (Abs_Integ (intrel `` {(i, j)})) = \
144ede948e58 renamings: real_of_nat, real_of_int -> (overloaded) real
paulson
parents: 10834
diff changeset
    20
\     Abs_REAL(realrel `` \
7292
dff3470c5c62 real literals using binary arithmetic
paulson
parents:
diff changeset
    21
\       {(preal_of_prat (prat_of_pnat (pnat_of_nat i)), \
dff3470c5c62 real literals using binary arithmetic
paulson
parents:
diff changeset
    22
\         preal_of_prat (prat_of_pnat (pnat_of_nat j)))})";
10919
144ede948e58 renamings: real_of_nat, real_of_int -> (overloaded) real
paulson
parents: 10834
diff changeset
    23
by (res_inst_tac [("f","Abs_REAL")] arg_cong 1);
144ede948e58 renamings: real_of_nat, real_of_int -> (overloaded) real
paulson
parents: 10834
diff changeset
    24
by (simp_tac (simpset() addsimps [realrel_in_real RS Abs_REAL_inverse,
9391
a6ab3a442da6 changed / to // for quotienting; general tidying
paulson
parents: 9365
diff changeset
    25
		[equiv_intrel, real_of_int_congruent] MRS UN_equiv_class]) 1);
7292
dff3470c5c62 real literals using binary arithmetic
paulson
parents:
diff changeset
    26
qed "real_of_int";
dff3470c5c62 real literals using binary arithmetic
paulson
parents:
diff changeset
    27
10919
144ede948e58 renamings: real_of_nat, real_of_int -> (overloaded) real
paulson
parents: 10834
diff changeset
    28
Goal "inj(real :: int => real)";
7292
dff3470c5c62 real literals using binary arithmetic
paulson
parents:
diff changeset
    29
by (rtac injI 1);
dff3470c5c62 real literals using binary arithmetic
paulson
parents:
diff changeset
    30
by (res_inst_tac [("z","x")] eq_Abs_Integ 1);
dff3470c5c62 real literals using binary arithmetic
paulson
parents:
diff changeset
    31
by (res_inst_tac [("z","y")] eq_Abs_Integ 1);
dff3470c5c62 real literals using binary arithmetic
paulson
parents:
diff changeset
    32
by (auto_tac (claset() addSDs [inj_preal_of_prat RS injD,
9069
e8d530582061 a big tidy-up
paulson
parents: 9043
diff changeset
    33
			   inj_prat_of_pnat RS injD, inj_pnat_of_nat RS injD],
e8d530582061 a big tidy-up
paulson
parents: 9043
diff changeset
    34
	      simpset() addsimps [real_of_int,preal_of_prat_add RS sym,
e8d530582061 a big tidy-up
paulson
parents: 9043
diff changeset
    35
				  prat_of_pnat_add RS sym,pnat_of_nat_add]));
7292
dff3470c5c62 real literals using binary arithmetic
paulson
parents:
diff changeset
    36
qed "inj_real_of_int";
dff3470c5c62 real literals using binary arithmetic
paulson
parents:
diff changeset
    37
10919
144ede948e58 renamings: real_of_nat, real_of_int -> (overloaded) real
paulson
parents: 10834
diff changeset
    38
Goalw [int_def,real_zero_def] "real (int 0) = 0";
7292
dff3470c5c62 real literals using binary arithmetic
paulson
parents:
diff changeset
    39
by (simp_tac (simpset() addsimps [real_of_int, preal_add_commute]) 1);
dff3470c5c62 real literals using binary arithmetic
paulson
parents:
diff changeset
    40
qed "real_of_int_zero";
dff3470c5c62 real literals using binary arithmetic
paulson
parents:
diff changeset
    41
11713
883d559b0b8c sane numerals (stage 3): provide generic "1" on all number types;
wenzelm
parents: 11597
diff changeset
    42
Goalw [int_def,real_one_def] "real (int 1) = (1::real)";
7292
dff3470c5c62 real literals using binary arithmetic
paulson
parents:
diff changeset
    43
by (auto_tac (claset(),
dff3470c5c62 real literals using binary arithmetic
paulson
parents:
diff changeset
    44
	      simpset() addsimps [real_of_int,
dff3470c5c62 real literals using binary arithmetic
paulson
parents:
diff changeset
    45
				  preal_of_prat_add RS sym,pnat_two_eq,
dff3470c5c62 real literals using binary arithmetic
paulson
parents:
diff changeset
    46
			       prat_of_pnat_add RS sym,pnat_one_iff RS sym]));
dff3470c5c62 real literals using binary arithmetic
paulson
parents:
diff changeset
    47
qed "real_of_int_one";
dff3470c5c62 real literals using binary arithmetic
paulson
parents:
diff changeset
    48
10919
144ede948e58 renamings: real_of_nat, real_of_int -> (overloaded) real
paulson
parents: 10834
diff changeset
    49
Goal "real (x::int) + real y = real (x + y)";
7292
dff3470c5c62 real literals using binary arithmetic
paulson
parents:
diff changeset
    50
by (res_inst_tac [("z","x")] eq_Abs_Integ 1);
dff3470c5c62 real literals using binary arithmetic
paulson
parents:
diff changeset
    51
by (res_inst_tac [("z","y")] eq_Abs_Integ 1);
10919
144ede948e58 renamings: real_of_nat, real_of_int -> (overloaded) real
paulson
parents: 10834
diff changeset
    52
by (auto_tac (claset(),
144ede948e58 renamings: real_of_nat, real_of_int -> (overloaded) real
paulson
parents: 10834
diff changeset
    53
     simpset() addsimps [real_of_int, preal_of_prat_add RS sym,
144ede948e58 renamings: real_of_nat, real_of_int -> (overloaded) real
paulson
parents: 10834
diff changeset
    54
                         prat_of_pnat_add RS sym, zadd,real_add,
144ede948e58 renamings: real_of_nat, real_of_int -> (overloaded) real
paulson
parents: 10834
diff changeset
    55
                         pnat_of_nat_add] @ pnat_add_ac));
7292
dff3470c5c62 real literals using binary arithmetic
paulson
parents:
diff changeset
    56
qed "real_of_int_add";
10784
27e4d90b35b5 more removal of obsolete rules
paulson
parents: 9391
diff changeset
    57
Addsimps [real_of_int_add RS sym];
7292
dff3470c5c62 real literals using binary arithmetic
paulson
parents:
diff changeset
    58
10919
144ede948e58 renamings: real_of_nat, real_of_int -> (overloaded) real
paulson
parents: 10834
diff changeset
    59
Goal "-real (x::int) = real (-x)";
7292
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";
10784
27e4d90b35b5 more removal of obsolete rules
paulson
parents: 9391
diff changeset
    63
Addsimps [real_of_int_minus RS sym];
7292
dff3470c5c62 real literals using binary arithmetic
paulson
parents:
diff changeset
    64
dff3470c5c62 real literals using binary arithmetic
paulson
parents:
diff changeset
    65
Goalw [zdiff_def,real_diff_def]
10919
144ede948e58 renamings: real_of_nat, real_of_int -> (overloaded) real
paulson
parents: 10834
diff changeset
    66
  "real (x::int) - real y = real (x - y)";
10784
27e4d90b35b5 more removal of obsolete rules
paulson
parents: 9391
diff changeset
    67
by (simp_tac (HOL_ss addsimps [real_of_int_add, real_of_int_minus]) 1);
7292
dff3470c5c62 real literals using binary arithmetic
paulson
parents:
diff changeset
    68
qed "real_of_int_diff";
10784
27e4d90b35b5 more removal of obsolete rules
paulson
parents: 9391
diff changeset
    69
Addsimps [real_of_int_diff RS sym];
7292
dff3470c5c62 real literals using binary arithmetic
paulson
parents:
diff changeset
    70
10919
144ede948e58 renamings: real_of_nat, real_of_int -> (overloaded) real
paulson
parents: 10834
diff changeset
    71
Goal "real (x::int) * real y = real (x * y)";
7292
dff3470c5c62 real literals using binary arithmetic
paulson
parents:
diff changeset
    72
by (res_inst_tac [("z","x")] eq_Abs_Integ 1);
dff3470c5c62 real literals using binary arithmetic
paulson
parents:
diff changeset
    73
by (res_inst_tac [("z","y")] eq_Abs_Integ 1);
10784
27e4d90b35b5 more removal of obsolete rules
paulson
parents: 9391
diff changeset
    74
by (auto_tac (claset(),
27e4d90b35b5 more removal of obsolete rules
paulson
parents: 9391
diff changeset
    75
   simpset() addsimps [real_of_int, real_mult,zmult,
10919
144ede948e58 renamings: real_of_nat, real_of_int -> (overloaded) real
paulson
parents: 10834
diff changeset
    76
		       preal_of_prat_mult RS sym,pnat_of_nat_mult,
144ede948e58 renamings: real_of_nat, real_of_int -> (overloaded) real
paulson
parents: 10834
diff changeset
    77
		       prat_of_pnat_mult RS sym,preal_of_prat_add RS sym,
144ede948e58 renamings: real_of_nat, real_of_int -> (overloaded) real
paulson
parents: 10834
diff changeset
    78
		       prat_of_pnat_add RS sym,pnat_of_nat_add] @ 
144ede948e58 renamings: real_of_nat, real_of_int -> (overloaded) real
paulson
parents: 10834
diff changeset
    79
                      mult_ac @ add_ac @ pnat_add_ac));
7292
dff3470c5c62 real literals using binary arithmetic
paulson
parents:
diff changeset
    80
qed "real_of_int_mult";
10784
27e4d90b35b5 more removal of obsolete rules
paulson
parents: 9391
diff changeset
    81
Addsimps [real_of_int_mult RS sym];
7292
dff3470c5c62 real literals using binary arithmetic
paulson
parents:
diff changeset
    82
11713
883d559b0b8c sane numerals (stage 3): provide generic "1" on all number types;
wenzelm
parents: 11597
diff changeset
    83
Goal "real (int (Suc n)) = real (int n) + (1::real)";
10919
144ede948e58 renamings: real_of_nat, real_of_int -> (overloaded) real
paulson
parents: 10834
diff changeset
    84
by (simp_tac (simpset() addsimps [real_of_int_one RS sym, zadd_int] @ 
144ede948e58 renamings: real_of_nat, real_of_int -> (overloaded) real
paulson
parents: 10834
diff changeset
    85
                                 zadd_ac) 1);
7292
dff3470c5c62 real literals using binary arithmetic
paulson
parents:
diff changeset
    86
qed "real_of_int_Suc";
dff3470c5c62 real literals using binary arithmetic
paulson
parents:
diff changeset
    87
dff3470c5c62 real literals using binary arithmetic
paulson
parents:
diff changeset
    88
(* relating two of the embeddings *)
10919
144ede948e58 renamings: real_of_nat, real_of_int -> (overloaded) real
paulson
parents: 10834
diff changeset
    89
Goal "real (int n) = real n";
7292
dff3470c5c62 real literals using binary arithmetic
paulson
parents:
diff changeset
    90
by (induct_tac "n" 1);
10919
144ede948e58 renamings: real_of_nat, real_of_int -> (overloaded) real
paulson
parents: 10834
diff changeset
    91
by (ALLGOALS (simp_tac (HOL_ss addsimps [real_of_int_zero, real_of_nat_zero,
144ede948e58 renamings: real_of_nat, real_of_int -> (overloaded) real
paulson
parents: 10834
diff changeset
    92
                                         real_of_int_Suc,real_of_nat_Suc])));
8856
435187ffc64e fixed theory deps;
wenzelm
parents: 7292
diff changeset
    93
by (Simp_tac 1);
7292
dff3470c5c62 real literals using binary arithmetic
paulson
parents:
diff changeset
    94
qed "real_of_int_real_of_nat";
dff3470c5c62 real literals using binary arithmetic
paulson
parents:
diff changeset
    95
10919
144ede948e58 renamings: real_of_nat, real_of_int -> (overloaded) real
paulson
parents: 10834
diff changeset
    96
Goal "~neg z ==> real (nat z) = real z";
7292
dff3470c5c62 real literals using binary arithmetic
paulson
parents:
diff changeset
    97
by (asm_simp_tac (simpset() addsimps [not_neg_nat,
9069
e8d530582061 a big tidy-up
paulson
parents: 9043
diff changeset
    98
				      real_of_int_real_of_nat RS sym]) 1);
7292
dff3470c5c62 real literals using binary arithmetic
paulson
parents:
diff changeset
    99
qed "real_of_nat_real_of_int";
dff3470c5c62 real literals using binary arithmetic
paulson
parents:
diff changeset
   100
10919
144ede948e58 renamings: real_of_nat, real_of_int -> (overloaded) real
paulson
parents: 10834
diff changeset
   101
Goal "(real x = 0) = (x = int 0)";
7292
dff3470c5c62 real literals using binary arithmetic
paulson
parents:
diff changeset
   102
by (auto_tac (claset() addIs [inj_real_of_int RS injD],
9069
e8d530582061 a big tidy-up
paulson
parents: 9043
diff changeset
   103
	      HOL_ss addsimps [real_of_int_zero]));
7292
dff3470c5c62 real literals using binary arithmetic
paulson
parents:
diff changeset
   104
qed "real_of_int_zero_cancel";
dff3470c5c62 real literals using binary arithmetic
paulson
parents:
diff changeset
   105
Addsimps [real_of_int_zero_cancel];
dff3470c5c62 real literals using binary arithmetic
paulson
parents:
diff changeset
   106
10919
144ede948e58 renamings: real_of_nat, real_of_int -> (overloaded) real
paulson
parents: 10834
diff changeset
   107
Goal "real (x::int) < real y ==> x < y";
7292
dff3470c5c62 real literals using binary arithmetic
paulson
parents:
diff changeset
   108
by (rtac ccontr 1 THEN dtac (linorder_not_less RS iffD1) 1);
dff3470c5c62 real literals using binary arithmetic
paulson
parents:
diff changeset
   109
by (auto_tac (claset(),
dff3470c5c62 real literals using binary arithmetic
paulson
parents:
diff changeset
   110
	      simpset() addsimps [zle_iff_zadd, real_of_int_add RS sym,
dff3470c5c62 real literals using binary arithmetic
paulson
parents:
diff changeset
   111
				  real_of_int_real_of_nat,
9069
e8d530582061 a big tidy-up
paulson
parents: 9043
diff changeset
   112
				  linorder_not_le RS sym]));
7292
dff3470c5c62 real literals using binary arithmetic
paulson
parents:
diff changeset
   113
qed "real_of_int_less_cancel";
dff3470c5c62 real literals using binary arithmetic
paulson
parents:
diff changeset
   114
10919
144ede948e58 renamings: real_of_nat, real_of_int -> (overloaded) real
paulson
parents: 10834
diff changeset
   115
Goal "(real (x::int) = real y) = (x = y)";
9069
e8d530582061 a big tidy-up
paulson
parents: 9043
diff changeset
   116
by (blast_tac (claset() addSDs [inj_real_of_int RS injD]) 1);
11597
cd6d2eddf75f renamed real_of_int_eq_iff to real_of_int_inject;
wenzelm
parents: 10919
diff changeset
   117
qed "real_of_int_inject";
cd6d2eddf75f renamed real_of_int_eq_iff to real_of_int_inject;
wenzelm
parents: 10919
diff changeset
   118
AddIffs [real_of_int_inject];
9069
e8d530582061 a big tidy-up
paulson
parents: 9043
diff changeset
   119
10919
144ede948e58 renamings: real_of_nat, real_of_int -> (overloaded) real
paulson
parents: 10834
diff changeset
   120
Goal "x < y ==> (real (x::int) < real y)";
9069
e8d530582061 a big tidy-up
paulson
parents: 9043
diff changeset
   121
by (full_simp_tac (simpset() addsimps [linorder_not_le RS sym]) 1);
e8d530582061 a big tidy-up
paulson
parents: 9043
diff changeset
   122
by (auto_tac (claset() addSDs [real_of_int_less_cancel],
e8d530582061 a big tidy-up
paulson
parents: 9043
diff changeset
   123
	      simpset() addsimps [order_le_less]));
7292
dff3470c5c62 real literals using binary arithmetic
paulson
parents:
diff changeset
   124
qed "real_of_int_less_mono";
dff3470c5c62 real literals using binary arithmetic
paulson
parents:
diff changeset
   125
10919
144ede948e58 renamings: real_of_nat, real_of_int -> (overloaded) real
paulson
parents: 10834
diff changeset
   126
Goal "(real (x::int) < real y) = (x < y)";
144ede948e58 renamings: real_of_nat, real_of_int -> (overloaded) real
paulson
parents: 10834
diff changeset
   127
by (blast_tac (claset() addDs [real_of_int_less_cancel]
144ede948e58 renamings: real_of_nat, real_of_int -> (overloaded) real
paulson
parents: 10834
diff changeset
   128
			addIs [real_of_int_less_mono]) 1);
7292
dff3470c5c62 real literals using binary arithmetic
paulson
parents:
diff changeset
   129
qed "real_of_int_less_iff";
9069
e8d530582061 a big tidy-up
paulson
parents: 9043
diff changeset
   130
AddIffs [real_of_int_less_iff];
7292
dff3470c5c62 real literals using binary arithmetic
paulson
parents:
diff changeset
   131
10919
144ede948e58 renamings: real_of_nat, real_of_int -> (overloaded) real
paulson
parents: 10834
diff changeset
   132
Goal "(real (x::int) <= real y) = (x <= y)";
9069
e8d530582061 a big tidy-up
paulson
parents: 9043
diff changeset
   133
by (full_simp_tac (simpset() addsimps [linorder_not_less RS sym]) 1);
7292
dff3470c5c62 real literals using binary arithmetic
paulson
parents:
diff changeset
   134
qed "real_of_int_le_iff";
dff3470c5c62 real literals using binary arithmetic
paulson
parents:
diff changeset
   135
Addsimps [real_of_int_le_iff];