src/HOL/Real/RealDef.ML
author paulson
Tue, 23 Nov 1999 11:19:39 +0100
changeset 8027 8a27d0579e37
parent 7499 23e090051cb8
child 9043 ca761fe227d8
permissions -rw-r--r--
distributive laws for * over -
paulson@5588
     1
(*  Title       : Real/RealDef.ML
paulson@7219
     2
    ID          : $Id$
paulson@5588
     3
    Author      : Jacques D. Fleuriot
paulson@5588
     4
    Copyright   : 1998  University of Cambridge
paulson@5588
     5
    Description : The reals
paulson@5588
     6
*)
paulson@5588
     7
paulson@5588
     8
(*** Proving that realrel is an equivalence relation ***)
paulson@5588
     9
paulson@5588
    10
Goal "[| (x1::preal) + y2 = x2 + y1; x2 + y3 = x3 + y2 |] \
paulson@5588
    11
\            ==> x1 + y3 = x3 + y1";        
paulson@5588
    12
by (res_inst_tac [("C","y2")] preal_add_right_cancel 1);
paulson@5588
    13
by (rotate_tac 1 1 THEN dtac sym 1);
paulson@5588
    14
by (asm_full_simp_tac (simpset() addsimps preal_add_ac) 1);
paulson@5588
    15
by (rtac (preal_add_left_commute RS subst) 1);
paulson@5588
    16
by (res_inst_tac [("x1","x1")] (preal_add_assoc RS subst) 1);
paulson@5588
    17
by (asm_full_simp_tac (simpset() addsimps preal_add_ac) 1);
paulson@5588
    18
qed "preal_trans_lemma";
paulson@5588
    19
paulson@5588
    20
(** Natural deduction for realrel **)
paulson@5588
    21
paulson@5588
    22
Goalw [realrel_def]
paulson@5588
    23
    "(((x1,y1),(x2,y2)): realrel) = (x1 + y2 = x2 + y1)";
paulson@5588
    24
by (Blast_tac 1);
paulson@5588
    25
qed "realrel_iff";
paulson@5588
    26
paulson@5588
    27
Goalw [realrel_def]
paulson@5588
    28
    "[| x1 + y2 = x2 + y1 |] ==> ((x1,y1),(x2,y2)): realrel";
paulson@5588
    29
by (Blast_tac  1);
paulson@5588
    30
qed "realrelI";
paulson@5588
    31
paulson@5588
    32
Goalw [realrel_def]
paulson@5588
    33
  "p: realrel --> (EX x1 y1 x2 y2. \
paulson@5588
    34
\                  p = ((x1,y1),(x2,y2)) & x1 + y2 = x2 + y1)";
paulson@5588
    35
by (Blast_tac 1);
paulson@5588
    36
qed "realrelE_lemma";
paulson@5588
    37
paulson@5588
    38
val [major,minor] = goal thy
paulson@5588
    39
  "[| p: realrel;  \
paulson@5588
    40
\     !!x1 y1 x2 y2. [| p = ((x1,y1),(x2,y2));  x1+y2 = x2+y1 \
paulson@5588
    41
\                    |] ==> Q |] ==> Q";
paulson@5588
    42
by (cut_facts_tac [major RS (realrelE_lemma RS mp)] 1);
paulson@5588
    43
by (REPEAT (eresolve_tac [asm_rl,exE,conjE,minor] 1));
paulson@5588
    44
qed "realrelE";
paulson@5588
    45
paulson@5588
    46
AddSIs [realrelI];
paulson@5588
    47
AddSEs [realrelE];
paulson@5588
    48
paulson@5588
    49
Goal "(x,x): realrel";
paulson@5588
    50
by (stac surjective_pairing 1 THEN rtac (refl RS realrelI) 1);
paulson@5588
    51
qed "realrel_refl";
paulson@5588
    52
paulson@5588
    53
Goalw [equiv_def, refl_def, sym_def, trans_def]
paulson@5588
    54
    "equiv {x::(preal*preal).True} realrel";
paulson@5588
    55
by (fast_tac (claset() addSIs [realrel_refl] 
paulson@5588
    56
                      addSEs [sym,preal_trans_lemma]) 1);
paulson@5588
    57
qed "equiv_realrel";
paulson@5588
    58
paulson@5588
    59
val equiv_realrel_iff =
paulson@5588
    60
    [TrueI, TrueI] MRS 
paulson@5588
    61
    ([CollectI, CollectI] MRS 
paulson@5588
    62
    (equiv_realrel RS eq_equiv_class_iff));
paulson@5588
    63
paulson@5588
    64
Goalw  [real_def,realrel_def,quotient_def] "realrel^^{(x,y)}:real";
paulson@5588
    65
by (Blast_tac 1);
paulson@5588
    66
qed "realrel_in_real";
paulson@5588
    67
paulson@5588
    68
Goal "inj_on Abs_real real";
paulson@5588
    69
by (rtac inj_on_inverseI 1);
paulson@5588
    70
by (etac Abs_real_inverse 1);
paulson@5588
    71
qed "inj_on_Abs_real";
paulson@5588
    72
paulson@5588
    73
Addsimps [equiv_realrel_iff,inj_on_Abs_real RS inj_on_iff,
paulson@5588
    74
          realrel_iff, realrel_in_real, Abs_real_inverse];
paulson@5588
    75
paulson@5588
    76
Addsimps [equiv_realrel RS eq_equiv_class_iff];
paulson@5588
    77
val eq_realrelD = equiv_realrel RSN (2,eq_equiv_class);
paulson@5588
    78
paulson@5588
    79
Goal "inj(Rep_real)";
paulson@5588
    80
by (rtac inj_inverseI 1);
paulson@5588
    81
by (rtac Rep_real_inverse 1);
paulson@5588
    82
qed "inj_Rep_real";
paulson@5588
    83
paulson@7077
    84
(** real_of_preal: the injection from preal to real **)
paulson@7077
    85
Goal "inj(real_of_preal)";
paulson@5588
    86
by (rtac injI 1);
paulson@7077
    87
by (rewtac real_of_preal_def);
paulson@5588
    88
by (dtac (inj_on_Abs_real RS inj_onD) 1);
paulson@5588
    89
by (REPEAT (rtac realrel_in_real 1));
paulson@5588
    90
by (dtac eq_equiv_class 1);
paulson@5588
    91
by (rtac equiv_realrel 1);
paulson@5588
    92
by (Blast_tac 1);
paulson@5588
    93
by Safe_tac;
paulson@5588
    94
by (Asm_full_simp_tac 1);
paulson@7077
    95
qed "inj_real_of_preal";
paulson@5588
    96
paulson@5588
    97
val [prem] = goal thy
paulson@5588
    98
    "(!!x y. z = Abs_real(realrel^^{(x,y)}) ==> P) ==> P";
paulson@5588
    99
by (res_inst_tac [("x1","z")] 
paulson@5588
   100
    (rewrite_rule [real_def] Rep_real RS quotientE) 1);
paulson@5588
   101
by (dres_inst_tac [("f","Abs_real")] arg_cong 1);
paulson@5588
   102
by (res_inst_tac [("p","x")] PairE 1);
paulson@5588
   103
by (rtac prem 1);
paulson@5588
   104
by (asm_full_simp_tac (simpset() addsimps [Rep_real_inverse]) 1);
paulson@5588
   105
qed "eq_Abs_real";
paulson@5588
   106
paulson@5588
   107
(**** real_minus: additive inverse on real ****)
paulson@5588
   108
paulson@5588
   109
Goalw [congruent_def]
paulson@5588
   110
  "congruent realrel (%p. split (%x y. realrel^^{(y,x)}) p)";
paulson@5588
   111
by Safe_tac;
paulson@5588
   112
by (asm_full_simp_tac (simpset() addsimps [preal_add_commute]) 1);
paulson@5588
   113
qed "real_minus_congruent";
paulson@5588
   114
paulson@5588
   115
(*Resolve th against the corresponding facts for real_minus*)
paulson@5588
   116
val real_minus_ize = RSLIST [equiv_realrel, real_minus_congruent];
paulson@5588
   117
paulson@5588
   118
Goalw [real_minus_def]
paulson@5588
   119
      "- (Abs_real(realrel^^{(x,y)})) = Abs_real(realrel ^^ {(y,x)})";
paulson@5588
   120
by (res_inst_tac [("f","Abs_real")] arg_cong 1);
paulson@5588
   121
by (simp_tac (simpset() addsimps 
paulson@5588
   122
   [realrel_in_real RS Abs_real_inverse,real_minus_ize UN_equiv_class]) 1);
paulson@5588
   123
qed "real_minus";
paulson@5588
   124
paulson@5588
   125
Goal "- (- z) = (z::real)";
paulson@5588
   126
by (res_inst_tac [("z","z")] eq_Abs_real 1);
paulson@5588
   127
by (asm_simp_tac (simpset() addsimps [real_minus]) 1);
paulson@5588
   128
qed "real_minus_minus";
paulson@5588
   129
paulson@5588
   130
Addsimps [real_minus_minus];
paulson@5588
   131
paulson@5588
   132
Goal "inj(%r::real. -r)";
paulson@5588
   133
by (rtac injI 1);
paulson@5588
   134
by (dres_inst_tac [("f","uminus")] arg_cong 1);
paulson@5588
   135
by (asm_full_simp_tac (simpset() addsimps [real_minus_minus]) 1);
paulson@5588
   136
qed "inj_real_minus";
paulson@5588
   137
paulson@5588
   138
Goalw [real_zero_def] "-0r = 0r";
paulson@5588
   139
by (simp_tac (simpset() addsimps [real_minus]) 1);
paulson@5588
   140
qed "real_minus_zero";
paulson@5588
   141
paulson@5588
   142
Addsimps [real_minus_zero];
paulson@5588
   143
paulson@5588
   144
Goal "(-x = 0r) = (x = 0r)"; 
paulson@5588
   145
by (res_inst_tac [("z","x")] eq_Abs_real 1);
paulson@5588
   146
by (auto_tac (claset(),
paulson@5588
   147
	      simpset() addsimps [real_zero_def, real_minus] @ preal_add_ac));
paulson@5588
   148
qed "real_minus_zero_iff";
paulson@5588
   149
paulson@5588
   150
Addsimps [real_minus_zero_iff];
paulson@5588
   151
paulson@5588
   152
Goal "(-x ~= 0r) = (x ~= 0r)"; 
paulson@5588
   153
by Auto_tac;
paulson@5588
   154
qed "real_minus_not_zero_iff";
paulson@5588
   155
paulson@5588
   156
(*** Congruence property for addition ***)
paulson@5588
   157
Goalw [congruent2_def]
paulson@5588
   158
    "congruent2 realrel (%p1 p2.                  \
paulson@5588
   159
\         split (%x1 y1. split (%x2 y2. realrel^^{(x1+x2, y1+y2)}) p2) p1)";
paulson@5588
   160
by Safe_tac;
paulson@5588
   161
by (asm_simp_tac (simpset() addsimps [preal_add_assoc]) 1);
paulson@5588
   162
by (res_inst_tac [("z1.1","x1a")] (preal_add_left_commute RS ssubst) 1);
paulson@5588
   163
by (asm_simp_tac (simpset() addsimps [preal_add_assoc RS sym]) 1);
paulson@5588
   164
by (asm_simp_tac (simpset() addsimps preal_add_ac) 1);
paulson@5588
   165
qed "real_add_congruent2";
paulson@5588
   166
paulson@5588
   167
(*Resolve th against the corresponding facts for real_add*)
paulson@5588
   168
val real_add_ize = RSLIST [equiv_realrel, real_add_congruent2];
paulson@5588
   169
paulson@5588
   170
Goalw [real_add_def]
paulson@5588
   171
  "Abs_real(realrel^^{(x1,y1)}) + Abs_real(realrel^^{(x2,y2)}) = \
paulson@5588
   172
\  Abs_real(realrel^^{(x1+x2, y1+y2)})";
paulson@5588
   173
by (asm_simp_tac (simpset() addsimps [real_add_ize UN_equiv_class2]) 1);
paulson@5588
   174
qed "real_add";
paulson@5588
   175
paulson@5588
   176
Goal "(z::real) + w = w + z";
paulson@5588
   177
by (res_inst_tac [("z","z")] eq_Abs_real 1);
paulson@5588
   178
by (res_inst_tac [("z","w")] eq_Abs_real 1);
paulson@5588
   179
by (asm_simp_tac (simpset() addsimps preal_add_ac @ [real_add]) 1);
paulson@5588
   180
qed "real_add_commute";
paulson@5588
   181
paulson@5588
   182
Goal "((z1::real) + z2) + z3 = z1 + (z2 + z3)";
paulson@5588
   183
by (res_inst_tac [("z","z1")] eq_Abs_real 1);
paulson@5588
   184
by (res_inst_tac [("z","z2")] eq_Abs_real 1);
paulson@5588
   185
by (res_inst_tac [("z","z3")] eq_Abs_real 1);
paulson@5588
   186
by (asm_simp_tac (simpset() addsimps [real_add, preal_add_assoc]) 1);
paulson@5588
   187
qed "real_add_assoc";
paulson@5588
   188
paulson@5588
   189
(*For AC rewriting*)
paulson@5588
   190
Goal "(x::real)+(y+z)=y+(x+z)";
paulson@5588
   191
by (rtac (real_add_commute RS trans) 1);
paulson@5588
   192
by (rtac (real_add_assoc RS trans) 1);
paulson@5588
   193
by (rtac (real_add_commute RS arg_cong) 1);
paulson@5588
   194
qed "real_add_left_commute";
paulson@5588
   195
paulson@5588
   196
(* real addition is an AC operator *)
wenzelm@7428
   197
bind_thms ("real_add_ac", [real_add_assoc,real_add_commute,real_add_left_commute]);
paulson@5588
   198
paulson@7077
   199
Goalw [real_of_preal_def,real_zero_def] "0r + z = z";
paulson@5588
   200
by (res_inst_tac [("z","z")] eq_Abs_real 1);
paulson@5588
   201
by (asm_full_simp_tac (simpset() addsimps [real_add] @ preal_add_ac) 1);
paulson@5588
   202
qed "real_add_zero_left";
paulson@5588
   203
Addsimps [real_add_zero_left];
paulson@5588
   204
paulson@5588
   205
Goal "z + 0r = z";
paulson@5588
   206
by (simp_tac (simpset() addsimps [real_add_commute]) 1);
paulson@5588
   207
qed "real_add_zero_right";
paulson@5588
   208
Addsimps [real_add_zero_right];
paulson@5588
   209
paulson@7127
   210
Goalw [real_zero_def] "z + (-z) = 0r";
paulson@5588
   211
by (res_inst_tac [("z","z")] eq_Abs_real 1);
paulson@5588
   212
by (asm_full_simp_tac (simpset() addsimps [real_minus,
paulson@5588
   213
        real_add, preal_add_commute]) 1);
paulson@5588
   214
qed "real_add_minus";
paulson@5588
   215
Addsimps [real_add_minus];
paulson@5588
   216
paulson@7127
   217
Goal "(-z) + z = 0r";
paulson@5588
   218
by (simp_tac (simpset() addsimps [real_add_commute]) 1);
paulson@5588
   219
qed "real_add_minus_left";
paulson@5588
   220
Addsimps [real_add_minus_left];
paulson@5588
   221
paulson@5588
   222
paulson@7127
   223
Goal "z + ((- z) + w) = (w::real)";
paulson@5588
   224
by (simp_tac (simpset() addsimps [real_add_assoc RS sym]) 1);
paulson@5588
   225
qed "real_add_minus_cancel";
paulson@5588
   226
paulson@5588
   227
Goal "(-z) + (z + w) = (w::real)";
paulson@5588
   228
by (simp_tac (simpset() addsimps [real_add_assoc RS sym]) 1);
paulson@5588
   229
qed "real_minus_add_cancel";
paulson@5588
   230
paulson@5588
   231
Addsimps [real_add_minus_cancel, real_minus_add_cancel];
paulson@5588
   232
paulson@5588
   233
Goal "? y. (x::real) + y = 0r";
paulson@5588
   234
by (blast_tac (claset() addIs [real_add_minus]) 1);
paulson@5588
   235
qed "real_minus_ex";
paulson@5588
   236
paulson@5588
   237
Goal "?! y. (x::real) + y = 0r";
paulson@5588
   238
by (auto_tac (claset() addIs [real_add_minus],simpset()));
paulson@5588
   239
by (dres_inst_tac [("f","%x. ya+x")] arg_cong 1);
paulson@5588
   240
by (asm_full_simp_tac (simpset() addsimps [real_add_assoc RS sym]) 1);
paulson@5588
   241
by (asm_full_simp_tac (simpset() addsimps [real_add_commute]) 1);
paulson@5588
   242
qed "real_minus_ex1";
paulson@5588
   243
paulson@5588
   244
Goal "?! y. y + (x::real) = 0r";
paulson@5588
   245
by (auto_tac (claset() addIs [real_add_minus_left],simpset()));
paulson@5588
   246
by (dres_inst_tac [("f","%x. x+ya")] arg_cong 1);
paulson@5588
   247
by (asm_full_simp_tac (simpset() addsimps [real_add_assoc]) 1);
paulson@5588
   248
by (asm_full_simp_tac (simpset() addsimps [real_add_commute]) 1);
paulson@5588
   249
qed "real_minus_left_ex1";
paulson@5588
   250
paulson@5588
   251
Goal "x + y = 0r ==> x = -y";
paulson@5588
   252
by (cut_inst_tac [("z","y")] real_add_minus_left 1);
paulson@5588
   253
by (res_inst_tac [("x1","y")] (real_minus_left_ex1 RS ex1E) 1);
paulson@5588
   254
by (Blast_tac 1);
paulson@5588
   255
qed "real_add_minus_eq_minus";
paulson@5588
   256
paulson@7077
   257
Goal "? (y::real). x = -y";
paulson@7077
   258
by (cut_inst_tac [("x","x")] real_minus_ex 1);
paulson@7077
   259
by (etac exE 1 THEN dtac real_add_minus_eq_minus 1);
paulson@7077
   260
by (Fast_tac 1);
paulson@7077
   261
qed "real_as_add_inverse_ex";
paulson@7077
   262
paulson@7127
   263
Goal "-(x + y) = (-x) + (- y :: real)";
paulson@5588
   264
by (res_inst_tac [("z","x")] eq_Abs_real 1);
paulson@5588
   265
by (res_inst_tac [("z","y")] eq_Abs_real 1);
paulson@5588
   266
by (auto_tac (claset(),simpset() addsimps [real_minus,real_add]));
paulson@5588
   267
qed "real_minus_add_distrib";
paulson@5588
   268
paulson@5588
   269
Addsimps [real_minus_add_distrib];
paulson@5588
   270
paulson@5588
   271
Goal "((x::real) + y = x + z) = (y = z)";
paulson@5588
   272
by (Step_tac 1);
paulson@7127
   273
by (dres_inst_tac [("f","%t. (-x) + t")] arg_cong 1);
paulson@5588
   274
by (asm_full_simp_tac (simpset() addsimps [real_add_assoc RS sym]) 1);
paulson@5588
   275
qed "real_add_left_cancel";
paulson@5588
   276
paulson@5588
   277
Goal "(y + (x::real)= z + x) = (y = z)";
paulson@5588
   278
by (simp_tac (simpset() addsimps [real_add_commute,real_add_left_cancel]) 1);
paulson@5588
   279
qed "real_add_right_cancel";
paulson@5588
   280
paulson@7127
   281
Goal "((x::real) = y) = (0r = x + (- y))";
paulson@7077
   282
by (Step_tac 1);
paulson@7077
   283
by (res_inst_tac [("x1","-y")] 
paulson@7077
   284
      (real_add_right_cancel RS iffD1) 2);
paulson@7127
   285
by Auto_tac;
paulson@7077
   286
qed "real_eq_minus_iff"; 
paulson@7077
   287
paulson@7127
   288
Goal "((x::real) = y) = (x + (- y) = 0r)";
paulson@7077
   289
by (Step_tac 1);
paulson@7077
   290
by (res_inst_tac [("x1","-y")] 
paulson@7077
   291
      (real_add_right_cancel RS iffD1) 2);
paulson@7127
   292
by Auto_tac;
paulson@7077
   293
qed "real_eq_minus_iff2"; 
paulson@7077
   294
paulson@5588
   295
Goal "0r - x = -x";
paulson@5588
   296
by (simp_tac (simpset() addsimps [real_diff_def]) 1);
paulson@5588
   297
qed "real_diff_0";
paulson@5588
   298
paulson@5588
   299
Goal "x - 0r = x";
paulson@5588
   300
by (simp_tac (simpset() addsimps [real_diff_def]) 1);
paulson@5588
   301
qed "real_diff_0_right";
paulson@5588
   302
paulson@5588
   303
Goal "x - x = 0r";
paulson@5588
   304
by (simp_tac (simpset() addsimps [real_diff_def]) 1);
paulson@5588
   305
qed "real_diff_self";
paulson@5588
   306
paulson@5588
   307
Addsimps [real_diff_0, real_diff_0_right, real_diff_self];
paulson@5588
   308
paulson@5588
   309
paulson@5588
   310
(*** Congruence property for multiplication ***)
paulson@5588
   311
paulson@5588
   312
Goal "!!(x1::preal). [| x1 + y2 = x2 + y1 |] ==> \
paulson@5588
   313
\         x * x1 + y * y1 + (x * y2 + x2 * y) = \
paulson@5588
   314
\         x * x2 + y * y2 + (x * y1 + x1 * y)";
paulson@5588
   315
by (asm_full_simp_tac (simpset() addsimps [preal_add_left_commute,
paulson@5588
   316
    preal_add_assoc RS sym,preal_add_mult_distrib2 RS sym]) 1);
paulson@5588
   317
by (rtac (preal_mult_commute RS subst) 1);
paulson@5588
   318
by (res_inst_tac [("y1","x2")] (preal_mult_commute RS subst) 1);
paulson@5588
   319
by (asm_full_simp_tac (simpset() addsimps [preal_add_assoc,
paulson@5588
   320
    preal_add_mult_distrib2 RS sym]) 1);
paulson@5588
   321
by (asm_full_simp_tac (simpset() addsimps [preal_add_commute]) 1);
paulson@5588
   322
qed "real_mult_congruent2_lemma";
paulson@5588
   323
paulson@5588
   324
Goal 
paulson@5588
   325
    "congruent2 realrel (%p1 p2.                  \
paulson@5588
   326
\         split (%x1 y1. split (%x2 y2. realrel^^{(x1*x2 + y1*y2, x1*y2+x2*y1)}) p2) p1)";
paulson@5588
   327
by (rtac (equiv_realrel RS congruent2_commuteI) 1);
paulson@5588
   328
by Safe_tac;
paulson@5588
   329
by (rewtac split_def);
paulson@5588
   330
by (asm_simp_tac (simpset() addsimps [preal_mult_commute,preal_add_commute]) 1);
paulson@5588
   331
by (auto_tac (claset(),simpset() addsimps [real_mult_congruent2_lemma]));
paulson@5588
   332
qed "real_mult_congruent2";
paulson@5588
   333
paulson@5588
   334
(*Resolve th against the corresponding facts for real_mult*)
paulson@5588
   335
val real_mult_ize = RSLIST [equiv_realrel, real_mult_congruent2];
paulson@5588
   336
paulson@5588
   337
Goalw [real_mult_def]
paulson@5588
   338
   "Abs_real((realrel^^{(x1,y1)})) * Abs_real((realrel^^{(x2,y2)})) =   \
paulson@5588
   339
\   Abs_real(realrel ^^ {(x1*x2+y1*y2,x1*y2+x2*y1)})";
paulson@5588
   340
by (simp_tac (simpset() addsimps [real_mult_ize UN_equiv_class2]) 1);
paulson@5588
   341
qed "real_mult";
paulson@5588
   342
paulson@5588
   343
Goal "(z::real) * w = w * z";
paulson@5588
   344
by (res_inst_tac [("z","z")] eq_Abs_real 1);
paulson@5588
   345
by (res_inst_tac [("z","w")] eq_Abs_real 1);
paulson@5588
   346
by (asm_simp_tac
paulson@5588
   347
    (simpset() addsimps [real_mult] @ preal_add_ac @ preal_mult_ac) 1);
paulson@5588
   348
qed "real_mult_commute";
paulson@5588
   349
paulson@5588
   350
Goal "((z1::real) * z2) * z3 = z1 * (z2 * z3)";
paulson@5588
   351
by (res_inst_tac [("z","z1")] eq_Abs_real 1);
paulson@5588
   352
by (res_inst_tac [("z","z2")] eq_Abs_real 1);
paulson@5588
   353
by (res_inst_tac [("z","z3")] eq_Abs_real 1);
paulson@5588
   354
by (asm_simp_tac (simpset() addsimps [preal_add_mult_distrib2,real_mult] @ 
paulson@5588
   355
                                     preal_add_ac @ preal_mult_ac) 1);
paulson@5588
   356
qed "real_mult_assoc";
paulson@5588
   357
paulson@5588
   358
qed_goal "real_mult_left_commute" thy
paulson@5588
   359
    "(z1::real) * (z2 * z3) = z2 * (z1 * z3)"
paulson@5588
   360
 (fn _ => [rtac (real_mult_commute RS trans) 1, rtac (real_mult_assoc RS trans) 1,
paulson@5588
   361
           rtac (real_mult_commute RS arg_cong) 1]);
paulson@5588
   362
paulson@5588
   363
(* real multiplication is an AC operator *)
wenzelm@7428
   364
bind_thms ("real_mult_ac", [real_mult_assoc, real_mult_commute, real_mult_left_commute]);
paulson@5588
   365
paulson@5588
   366
Goalw [real_one_def,pnat_one_def] "1r * z = z";
paulson@5588
   367
by (res_inst_tac [("z","z")] eq_Abs_real 1);
paulson@5588
   368
by (asm_full_simp_tac
paulson@5588
   369
    (simpset() addsimps [real_mult,
paulson@5588
   370
			 preal_add_mult_distrib2,preal_mult_1_right] 
paulson@5588
   371
                        @ preal_mult_ac @ preal_add_ac) 1);
paulson@5588
   372
qed "real_mult_1";
paulson@5588
   373
paulson@5588
   374
Addsimps [real_mult_1];
paulson@5588
   375
paulson@5588
   376
Goal "z * 1r = z";
paulson@5588
   377
by (simp_tac (simpset() addsimps [real_mult_commute]) 1);
paulson@5588
   378
qed "real_mult_1_right";
paulson@5588
   379
paulson@5588
   380
Addsimps [real_mult_1_right];
paulson@5588
   381
paulson@5588
   382
Goalw [real_zero_def,pnat_one_def] "0r * z = 0r";
paulson@5588
   383
by (res_inst_tac [("z","z")] eq_Abs_real 1);
paulson@5588
   384
by (asm_full_simp_tac (simpset() addsimps [real_mult,
paulson@5588
   385
    preal_add_mult_distrib2,preal_mult_1_right] 
paulson@5588
   386
    @ preal_mult_ac @ preal_add_ac) 1);
paulson@5588
   387
qed "real_mult_0";
paulson@5588
   388
paulson@5588
   389
Goal "z * 0r = 0r";
paulson@5588
   390
by (simp_tac (simpset() addsimps [real_mult_commute, real_mult_0]) 1);
paulson@5588
   391
qed "real_mult_0_right";
paulson@5588
   392
paulson@5588
   393
Addsimps [real_mult_0_right, real_mult_0];
paulson@5588
   394
paulson@7127
   395
Goal "-(x * y) = (-x) * (y::real)";
paulson@5588
   396
by (res_inst_tac [("z","x")] eq_Abs_real 1);
paulson@5588
   397
by (res_inst_tac [("z","y")] eq_Abs_real 1);
paulson@5588
   398
by (auto_tac (claset(),
paulson@5588
   399
	      simpset() addsimps [real_minus,real_mult] 
paulson@5588
   400
    @ preal_mult_ac @ preal_add_ac));
paulson@5588
   401
qed "real_minus_mult_eq1";
paulson@5588
   402
paulson@7127
   403
Goal "-(x * y) = x * (- y :: real)";
paulson@5588
   404
by (res_inst_tac [("z","x")] eq_Abs_real 1);
paulson@5588
   405
by (res_inst_tac [("z","y")] eq_Abs_real 1);
paulson@5588
   406
by (auto_tac (claset(),
paulson@5588
   407
	      simpset() addsimps [real_minus,real_mult] 
paulson@5588
   408
    @ preal_mult_ac @ preal_add_ac));
paulson@5588
   409
qed "real_minus_mult_eq2";
paulson@5588
   410
paulson@7127
   411
Goal "(- 1r) * z = -z";
paulson@5588
   412
by (simp_tac (simpset() addsimps [real_minus_mult_eq1 RS sym]) 1);
paulson@5588
   413
qed "real_mult_minus_1";
paulson@5588
   414
paulson@5588
   415
Addsimps [real_mult_minus_1];
paulson@5588
   416
paulson@7127
   417
Goal "z * (- 1r) = -z";
paulson@5588
   418
by (stac real_mult_commute 1);
paulson@5588
   419
by (Simp_tac 1);
paulson@5588
   420
qed "real_mult_minus_1_right";
paulson@5588
   421
paulson@5588
   422
Addsimps [real_mult_minus_1_right];
paulson@5588
   423
paulson@7127
   424
Goal "(-x) * (-y) = x * (y::real)";
paulson@5588
   425
by (full_simp_tac (simpset() addsimps [real_minus_mult_eq2 RS sym,
paulson@7127
   426
				       real_minus_mult_eq1 RS sym]) 1);
paulson@5588
   427
qed "real_minus_mult_cancel";
paulson@5588
   428
paulson@5588
   429
Addsimps [real_minus_mult_cancel];
paulson@5588
   430
paulson@7127
   431
Goal "(-x) * y = x * (- y :: real)";
paulson@5588
   432
by (full_simp_tac (simpset() addsimps [real_minus_mult_eq2 RS sym,
paulson@7127
   433
				       real_minus_mult_eq1 RS sym]) 1);
paulson@5588
   434
qed "real_minus_mult_commute";
paulson@5588
   435
paulson@5588
   436
(*-----------------------------------------------------------------------------
paulson@5588
   437
paulson@7127
   438
 ----------------------------------------------------------------------------*)
paulson@5588
   439
paulson@5588
   440
(** Lemmas **)
paulson@5588
   441
paulson@5588
   442
qed_goal "real_add_assoc_cong" thy
paulson@5588
   443
    "!!z. (z::real) + v = z' + v' ==> z + (v + w) = z' + (v' + w)"
paulson@5588
   444
 (fn _ => [(asm_simp_tac (simpset() addsimps [real_add_assoc RS sym]) 1)]);
paulson@5588
   445
paulson@5588
   446
qed_goal "real_add_assoc_swap" thy "(z::real) + (v + w) = v + (z + w)"
paulson@5588
   447
 (fn _ => [(REPEAT (ares_tac [real_add_commute RS real_add_assoc_cong] 1))]);
paulson@5588
   448
paulson@5588
   449
Goal "((z1::real) + z2) * w = (z1 * w) + (z2 * w)";
paulson@5588
   450
by (res_inst_tac [("z","z1")] eq_Abs_real 1);
paulson@5588
   451
by (res_inst_tac [("z","z2")] eq_Abs_real 1);
paulson@5588
   452
by (res_inst_tac [("z","w")] eq_Abs_real 1);
paulson@5588
   453
by (asm_simp_tac 
paulson@5588
   454
    (simpset() addsimps [preal_add_mult_distrib2, real_add, real_mult] @ 
paulson@5588
   455
                        preal_add_ac @ preal_mult_ac) 1);
paulson@5588
   456
qed "real_add_mult_distrib";
paulson@5588
   457
paulson@5588
   458
val real_mult_commute'= read_instantiate [("z","w")] real_mult_commute;
paulson@5588
   459
paulson@5588
   460
Goal "(w::real) * (z1 + z2) = (w * z1) + (w * z2)";
paulson@5588
   461
by (simp_tac (simpset() addsimps [real_mult_commute',real_add_mult_distrib]) 1);
paulson@5588
   462
qed "real_add_mult_distrib2";
paulson@5588
   463
paulson@8027
   464
Goalw [real_diff_def] "((z1::real) - z2) * w = (z1 * w) - (z2 * w)";
paulson@8027
   465
by (simp_tac (simpset() addsimps [real_add_mult_distrib, 
paulson@8027
   466
				  real_minus_mult_eq1]) 1);
paulson@8027
   467
qed "real_diff_mult_distrib";
paulson@8027
   468
paulson@8027
   469
Goal "(w::real) * (z1 - z2) = (w * z1) - (w * z2)";
paulson@8027
   470
by (simp_tac (simpset() addsimps [real_mult_commute', 
paulson@8027
   471
				  real_diff_mult_distrib]) 1);
paulson@8027
   472
qed "real_diff_mult_distrib2";
paulson@8027
   473
paulson@5588
   474
(*** one and zero are distinct ***)
paulson@5588
   475
Goalw [real_zero_def,real_one_def] "0r ~= 1r";
paulson@5588
   476
by (auto_tac (claset(),
paulson@5588
   477
         simpset() addsimps [preal_self_less_add_left RS preal_not_refl2]));
paulson@5588
   478
qed "real_zero_not_eq_one";
paulson@5588
   479
paulson@5588
   480
(*** existence of inverse ***)
paulson@5588
   481
(** lemma -- alternative definition for 0r **)
paulson@5588
   482
Goalw [real_zero_def] "0r = Abs_real (realrel ^^ {(x, x)})";
paulson@5588
   483
by (auto_tac (claset(),simpset() addsimps [preal_add_commute]));
paulson@5588
   484
qed "real_zero_iff";
paulson@5588
   485
paulson@5588
   486
Goalw [real_zero_def,real_one_def] 
paulson@5588
   487
          "!!(x::real). x ~= 0r ==> ? y. x*y = 1r";
paulson@5588
   488
by (res_inst_tac [("z","x")] eq_Abs_real 1);
paulson@5588
   489
by (cut_inst_tac [("r1.0","xa"),("r2.0","y")] preal_linear 1);
paulson@5588
   490
by (auto_tac (claset() addSDs [preal_less_add_left_Ex],
paulson@5588
   491
           simpset() addsimps [real_zero_iff RS sym]));
paulson@7077
   492
by (res_inst_tac [("x","Abs_real (realrel ^^ \
paulson@7077
   493
\   {(preal_of_prat(prat_of_pnat 1p),pinv(D)+\
paulson@7077
   494
\    preal_of_prat(prat_of_pnat 1p))})")] exI 1);
paulson@7077
   495
by (res_inst_tac [("x","Abs_real (realrel ^^ \
paulson@7077
   496
\   {(pinv(D)+preal_of_prat(prat_of_pnat 1p),\
paulson@7077
   497
\    preal_of_prat(prat_of_pnat 1p))})")] exI 2);
paulson@5588
   498
by (auto_tac (claset(),
paulson@5588
   499
	      simpset() addsimps [real_mult,
paulson@5588
   500
    pnat_one_def,preal_mult_1_right,preal_add_mult_distrib2,
paulson@5588
   501
    preal_add_mult_distrib,preal_mult_1,preal_mult_inv_right] 
paulson@5588
   502
    @ preal_add_ac @ preal_mult_ac));
paulson@5588
   503
qed "real_mult_inv_right_ex";
paulson@5588
   504
paulson@5588
   505
Goal "!!(x::real). x ~= 0r ==> ? y. y*x = 1r";
paulson@5588
   506
by (asm_simp_tac (simpset() addsimps [real_mult_commute,
paulson@7127
   507
				      real_mult_inv_right_ex]) 1);
paulson@5588
   508
qed "real_mult_inv_left_ex";
paulson@5588
   509
paulson@7127
   510
Goalw [rinv_def] "x ~= 0r ==> rinv(x)*x = 1r";
wenzelm@7499
   511
by (ftac real_mult_inv_left_ex 1);
paulson@5588
   512
by (Step_tac 1);
paulson@5588
   513
by (rtac selectI2 1);
paulson@5588
   514
by Auto_tac;
paulson@5588
   515
qed "real_mult_inv_left";
paulson@5588
   516
paulson@7127
   517
Goal "x ~= 0r ==> x*rinv(x) = 1r";
paulson@5588
   518
by (auto_tac (claset() addIs [real_mult_commute RS subst],
paulson@5588
   519
              simpset() addsimps [real_mult_inv_left]));
paulson@5588
   520
qed "real_mult_inv_right";
paulson@5588
   521
paulson@5588
   522
Goal "(c::real) ~= 0r ==> (c*a=c*b) = (a=b)";
paulson@5588
   523
by Auto_tac;
paulson@5588
   524
by (dres_inst_tac [("f","%x. x*rinv c")] arg_cong 1);
paulson@5588
   525
by (asm_full_simp_tac (simpset() addsimps [real_mult_inv_right] @ real_mult_ac)  1);
paulson@5588
   526
qed "real_mult_left_cancel";
paulson@5588
   527
    
paulson@5588
   528
Goal "(c::real) ~= 0r ==> (a*c=b*c) = (a=b)";
paulson@5588
   529
by (Step_tac 1);
paulson@5588
   530
by (dres_inst_tac [("f","%x. x*rinv c")] arg_cong 1);
paulson@7127
   531
by (asm_full_simp_tac
paulson@7127
   532
    (simpset() addsimps [real_mult_inv_right] @ real_mult_ac)  1);
paulson@5588
   533
qed "real_mult_right_cancel";
paulson@5588
   534
paulson@7127
   535
Goal "c*a ~= c*b ==> a ~= b";
paulson@7127
   536
by Auto_tac;
paulson@7077
   537
qed "real_mult_left_cancel_ccontr";
paulson@7077
   538
paulson@7127
   539
Goal "a*c ~= b*c ==> a ~= b";
paulson@7127
   540
by Auto_tac;
paulson@7077
   541
qed "real_mult_right_cancel_ccontr";
paulson@7077
   542
paulson@5588
   543
Goalw [rinv_def] "x ~= 0r ==> rinv(x) ~= 0r";
wenzelm@7499
   544
by (ftac real_mult_inv_left_ex 1);
paulson@5588
   545
by (etac exE 1);
paulson@5588
   546
by (rtac selectI2 1);
paulson@5588
   547
by (auto_tac (claset(),
paulson@5588
   548
	      simpset() addsimps [real_mult_0,
paulson@5588
   549
    real_zero_not_eq_one]));
paulson@5588
   550
qed "rinv_not_zero";
paulson@5588
   551
paulson@5588
   552
Addsimps [real_mult_inv_left,real_mult_inv_right];
paulson@5588
   553
paulson@7127
   554
Goal "[| x ~= 0r; y ~= 0r |] ==> x * y ~= 0r";
paulson@7077
   555
by (Step_tac 1);
paulson@7077
   556
by (dres_inst_tac [("f","%z. rinv x*z")] arg_cong 1);
paulson@7077
   557
by (asm_full_simp_tac (simpset() addsimps [real_mult_assoc RS sym]) 1);
paulson@7077
   558
qed "real_mult_not_zero";
paulson@7077
   559
paulson@7077
   560
bind_thm ("real_mult_not_zeroE",real_mult_not_zero RS notE);
paulson@7077
   561
paulson@5588
   562
Goal "x ~= 0r ==> rinv(rinv x) = x";
paulson@5588
   563
by (res_inst_tac [("c1","rinv x")] (real_mult_right_cancel RS iffD1) 1);
paulson@5588
   564
by (etac rinv_not_zero 1);
paulson@5588
   565
by (auto_tac (claset() addDs [rinv_not_zero],simpset()));
paulson@5588
   566
qed "real_rinv_rinv";
paulson@5588
   567
paulson@5588
   568
Goalw [rinv_def] "rinv(1r) = 1r";
paulson@5588
   569
by (cut_facts_tac [real_zero_not_eq_one RS 
paulson@5588
   570
       not_sym RS real_mult_inv_left_ex] 1);
paulson@5588
   571
by (etac exE 1);
paulson@5588
   572
by (rtac selectI2 1);
paulson@5588
   573
by (auto_tac (claset(),
paulson@5588
   574
	      simpset() addsimps 
paulson@5588
   575
    [real_zero_not_eq_one RS not_sym]));
paulson@5588
   576
qed "real_rinv_1";
paulson@7077
   577
Addsimps [real_rinv_1];
paulson@5588
   578
paulson@5588
   579
Goal "x ~= 0r ==> rinv(-x) = -rinv(x)";
paulson@5588
   580
by (res_inst_tac [("c1","-x")] (real_mult_right_cancel RS iffD1) 1);
paulson@5588
   581
by Auto_tac;
paulson@5588
   582
qed "real_minus_rinv";
paulson@5588
   583
paulson@7127
   584
Goal "[| x ~= 0r; y ~= 0r |] ==> rinv(x*y) = rinv(x)*rinv(y)";
paulson@7077
   585
by (forw_inst_tac [("y","y")] real_mult_not_zero 1 THEN assume_tac 1);
paulson@7077
   586
by (res_inst_tac [("c1","x")] (real_mult_left_cancel RS iffD1) 1);
paulson@7077
   587
by (auto_tac (claset(),simpset() addsimps [real_mult_assoc RS sym]));
paulson@7077
   588
by (res_inst_tac [("c1","y")] (real_mult_left_cancel RS iffD1) 1);
paulson@7077
   589
by (auto_tac (claset(),simpset() addsimps [real_mult_left_commute]));
paulson@7077
   590
by (asm_simp_tac (simpset() addsimps [real_mult_assoc RS sym]) 1);
paulson@7077
   591
qed "real_rinv_distrib";
paulson@7077
   592
paulson@7077
   593
(*---------------------------------------------------------
paulson@7077
   594
     Theorems for ordering
paulson@7077
   595
 --------------------------------------------------------*)
paulson@5588
   596
(* prove introduction and elimination rules for real_less *)
paulson@5588
   597
paulson@7077
   598
(* real_less is a strong order i.e. nonreflexive and transitive *)
paulson@7077
   599
paulson@5588
   600
(*** lemmas ***)
paulson@5588
   601
Goal "!!(x::preal). [| x = y; x1 = y1 |] ==> x + y1 = x1 + y";
paulson@5588
   602
by (asm_simp_tac (simpset() addsimps [preal_add_commute]) 1);
paulson@5588
   603
qed "preal_lemma_eq_rev_sum";
paulson@5588
   604
paulson@5588
   605
Goal "!!(b::preal). x + (b + y) = x1 + (b + y1) ==> x + y = x1 + y1";
paulson@5588
   606
by (asm_full_simp_tac (simpset() addsimps preal_add_ac) 1);
paulson@5588
   607
qed "preal_add_left_commute_cancel";
paulson@5588
   608
paulson@5588
   609
Goal "!!(x::preal). [| x + y2a = x2a + y; \
paulson@5588
   610
\                      x + y2b = x2b + y |] \
paulson@5588
   611
\                   ==> x2a + y2b = x2b + y2a";
paulson@5588
   612
by (dtac preal_lemma_eq_rev_sum 1);
paulson@5588
   613
by (assume_tac 1);
paulson@5588
   614
by (thin_tac "x + y2b = x2b + y" 1);
paulson@5588
   615
by (asm_full_simp_tac (simpset() addsimps preal_add_ac) 1);
paulson@5588
   616
by (dtac preal_add_left_commute_cancel 1);
paulson@5588
   617
by (asm_full_simp_tac (simpset() addsimps preal_add_ac) 1);
paulson@5588
   618
qed "preal_lemma_for_not_refl";
paulson@5588
   619
paulson@5588
   620
Goal "~ (R::real) < R";
paulson@5588
   621
by (res_inst_tac [("z","R")] eq_Abs_real 1);
paulson@5588
   622
by (auto_tac (claset(),simpset() addsimps [real_less_def]));
paulson@5588
   623
by (dtac preal_lemma_for_not_refl 1);
paulson@5588
   624
by (assume_tac 1 THEN rotate_tac 2 1);
paulson@5588
   625
by (auto_tac (claset(),simpset() addsimps [preal_less_not_refl]));
paulson@5588
   626
qed "real_less_not_refl";
paulson@5588
   627
paulson@5588
   628
(*** y < y ==> P ***)
paulson@5588
   629
bind_thm("real_less_irrefl", real_less_not_refl RS notE);
paulson@5588
   630
AddSEs [real_less_irrefl];
paulson@5588
   631
paulson@5588
   632
Goal "!!(x::real). x < y ==> x ~= y";
paulson@5588
   633
by (auto_tac (claset(),simpset() addsimps [real_less_not_refl]));
paulson@5588
   634
qed "real_not_refl2";
paulson@5588
   635
paulson@5588
   636
(* lemma re-arranging and eliminating terms *)
paulson@5588
   637
Goal "!! (a::preal). [| a + b = c + d; \
paulson@5588
   638
\            x2b + d + (c + y2e) < a + y2b + (x2e + b) |] \
paulson@5588
   639
\         ==> x2b + y2e < x2e + y2b";
paulson@5588
   640
by (asm_full_simp_tac (simpset() addsimps preal_add_ac) 1);
paulson@5588
   641
by (res_inst_tac [("C","c+d")] preal_add_left_less_cancel 1);
paulson@5588
   642
by (asm_full_simp_tac (simpset() addsimps [preal_add_assoc RS sym]) 1);
paulson@5588
   643
qed "preal_lemma_trans";
paulson@5588
   644
paulson@5588
   645
(** heavy re-writing involved*)
paulson@5588
   646
Goal "!!(R1::real). [| R1 < R2; R2 < R3 |] ==> R1 < R3";
paulson@5588
   647
by (res_inst_tac [("z","R1")] eq_Abs_real 1);
paulson@5588
   648
by (res_inst_tac [("z","R2")] eq_Abs_real 1);
paulson@5588
   649
by (res_inst_tac [("z","R3")] eq_Abs_real 1);
paulson@5588
   650
by (auto_tac (claset(),simpset() addsimps [real_less_def]));
paulson@5588
   651
by (REPEAT(rtac exI 1));
paulson@5588
   652
by (EVERY[rtac conjI 1, rtac conjI 2]);
paulson@5588
   653
by (REPEAT(Blast_tac 2));
paulson@5588
   654
by (dtac preal_lemma_for_not_refl 1 THEN assume_tac 1);
paulson@5588
   655
by (blast_tac (claset() addDs [preal_add_less_mono] 
paulson@5588
   656
    addIs [preal_lemma_trans]) 1);
paulson@5588
   657
qed "real_less_trans";
paulson@5588
   658
paulson@5588
   659
Goal "!! (R1::real). [| R1 < R2; R2 < R1 |] ==> P";
paulson@5588
   660
by (dtac real_less_trans 1 THEN assume_tac 1);
paulson@5588
   661
by (asm_full_simp_tac (simpset() addsimps [real_less_not_refl]) 1);
paulson@5588
   662
qed "real_less_asym";
paulson@5588
   663
paulson@5588
   664
(****)(****)(****)(****)(****)(****)(****)(****)(****)(****)
paulson@5588
   665
    (****** Map and more real_less ******)
paulson@5588
   666
(*** mapping from preal into real ***)
paulson@7077
   667
Goalw [real_of_preal_def] 
paulson@7077
   668
     "real_of_preal ((z1::preal) + z2) = \
paulson@7077
   669
\     real_of_preal z1 + real_of_preal z2";
paulson@5588
   670
by (asm_simp_tac (simpset() addsimps [real_add,
paulson@5588
   671
       preal_add_mult_distrib,preal_mult_1] addsimps preal_add_ac) 1);
paulson@7077
   672
qed "real_of_preal_add";
paulson@5588
   673
paulson@7077
   674
Goalw [real_of_preal_def] 
paulson@7077
   675
     "real_of_preal ((z1::preal) * z2) = \
paulson@7077
   676
\     real_of_preal z1* real_of_preal z2";
paulson@5588
   677
by (full_simp_tac (simpset() addsimps [real_mult,
paulson@5588
   678
        preal_add_mult_distrib2,preal_mult_1,
paulson@5588
   679
        preal_mult_1_right,pnat_one_def] 
paulson@5588
   680
        @ preal_add_ac @ preal_mult_ac) 1);
paulson@7077
   681
qed "real_of_preal_mult";
paulson@5588
   682
paulson@7077
   683
Goalw [real_of_preal_def]
paulson@7077
   684
      "!!(x::preal). y < x ==> \
paulson@7077
   685
\      ? m. Abs_real (realrel ^^ {(x,y)}) = real_of_preal m";
paulson@5588
   686
by (auto_tac (claset() addSDs [preal_less_add_left_Ex],
paulson@5588
   687
    simpset() addsimps preal_add_ac));
paulson@7077
   688
qed "real_of_preal_ExI";
paulson@5588
   689
paulson@7077
   690
Goalw [real_of_preal_def]
paulson@7077
   691
      "!!(x::preal). ? m. Abs_real (realrel ^^ {(x,y)}) = \
paulson@7077
   692
\                    real_of_preal m ==> y < x";
paulson@5588
   693
by (auto_tac (claset(),
paulson@5588
   694
	      simpset() addsimps 
paulson@5588
   695
    [preal_add_commute,preal_add_assoc]));
paulson@5588
   696
by (asm_full_simp_tac (simpset() addsimps 
paulson@5588
   697
    [preal_add_assoc RS sym,preal_self_less_add_left]) 1);
paulson@7077
   698
qed "real_of_preal_ExD";
paulson@5588
   699
paulson@7077
   700
Goal "(? m. Abs_real (realrel ^^ {(x,y)}) = real_of_preal m) = (y < x)";
paulson@7077
   701
by (blast_tac (claset() addSIs [real_of_preal_ExI,real_of_preal_ExD]) 1);
paulson@7077
   702
qed "real_of_preal_iff";
paulson@5588
   703
paulson@5588
   704
(*** Gleason prop 9-4.4 p 127 ***)
paulson@7077
   705
Goalw [real_of_preal_def,real_zero_def] 
paulson@7077
   706
      "? m. (x::real) = real_of_preal m | x = 0r | x = -(real_of_preal m)";
paulson@5588
   707
by (res_inst_tac [("z","x")] eq_Abs_real 1);
paulson@5588
   708
by (auto_tac (claset(),simpset() addsimps [real_minus] @ preal_add_ac));
paulson@5588
   709
by (cut_inst_tac [("r1.0","x"),("r2.0","y")] preal_linear 1);
paulson@5588
   710
by (auto_tac (claset() addSDs [preal_less_add_left_Ex],
paulson@5588
   711
    simpset() addsimps [preal_add_assoc RS sym]));
paulson@5588
   712
by (auto_tac (claset(),simpset() addsimps [preal_add_commute]));
paulson@7077
   713
qed "real_of_preal_trichotomy";
paulson@5588
   714
paulson@7077
   715
Goal "!!P. [| !!m. x = real_of_preal m ==> P; \
paulson@5588
   716
\             x = 0r ==> P; \
paulson@7077
   717
\             !!m. x = -(real_of_preal m) ==> P |] ==> P";
paulson@7077
   718
by (cut_inst_tac [("x","x")] real_of_preal_trichotomy 1);
paulson@5588
   719
by Auto_tac;
paulson@7077
   720
qed "real_of_preal_trichotomyE";
paulson@5588
   721
paulson@7077
   722
Goalw [real_of_preal_def] 
paulson@7077
   723
      "real_of_preal m1 < real_of_preal m2 ==> m1 < m2";
paulson@5588
   724
by (auto_tac (claset(),simpset() addsimps [real_less_def] @ preal_add_ac));
paulson@5588
   725
by (auto_tac (claset(),simpset() addsimps [preal_add_assoc RS sym]));
paulson@5588
   726
by (auto_tac (claset(),simpset() addsimps preal_add_ac));
paulson@7077
   727
qed "real_of_preal_lessD";
paulson@5588
   728
paulson@7077
   729
Goal "m1 < m2 ==> real_of_preal m1 < real_of_preal m2";
paulson@5588
   730
by (dtac preal_less_add_left_Ex 1);
paulson@5588
   731
by (auto_tac (claset(),
paulson@7077
   732
	      simpset() addsimps [real_of_preal_add,
paulson@7077
   733
    real_of_preal_def,real_less_def]));
paulson@5588
   734
by (REPEAT(rtac exI 1));
paulson@5588
   735
by (EVERY[rtac conjI 1, rtac conjI 2]);
paulson@5588
   736
by (REPEAT(Blast_tac 2));
paulson@5588
   737
by (simp_tac (simpset() addsimps [preal_self_less_add_left] 
paulson@5588
   738
    delsimps [preal_add_less_iff2]) 1);
paulson@7077
   739
qed "real_of_preal_lessI";
paulson@5588
   740
paulson@7077
   741
Goal "(real_of_preal m1 < real_of_preal m2) = (m1 < m2)";
paulson@7077
   742
by (blast_tac (claset() addIs [real_of_preal_lessI,real_of_preal_lessD]) 1);
paulson@7077
   743
qed "real_of_preal_less_iff1";
paulson@5588
   744
paulson@7077
   745
Addsimps [real_of_preal_less_iff1];
paulson@5588
   746
paulson@7077
   747
Goal "- real_of_preal m < real_of_preal m";
paulson@5588
   748
by (auto_tac (claset(),
paulson@5588
   749
	      simpset() addsimps 
paulson@7077
   750
    [real_of_preal_def,real_less_def,real_minus]));
paulson@5588
   751
by (REPEAT(rtac exI 1));
paulson@5588
   752
by (EVERY[rtac conjI 1, rtac conjI 2]);
paulson@5588
   753
by (REPEAT(Blast_tac 2));
paulson@5588
   754
by (full_simp_tac (simpset() addsimps preal_add_ac) 1);
paulson@5588
   755
by (full_simp_tac (simpset() addsimps [preal_self_less_add_right,
paulson@5588
   756
    preal_add_assoc RS sym]) 1);
paulson@7077
   757
qed "real_of_preal_minus_less_self";
paulson@5588
   758
paulson@7077
   759
Goalw [real_zero_def] "- real_of_preal m < 0r";
paulson@5588
   760
by (auto_tac (claset(),
paulson@7292
   761
	      simpset() addsimps [real_of_preal_def,
paulson@7292
   762
				  real_less_def,real_minus]));
paulson@5588
   763
by (REPEAT(rtac exI 1));
paulson@5588
   764
by (EVERY[rtac conjI 1, rtac conjI 2]);
paulson@5588
   765
by (REPEAT(Blast_tac 2));
paulson@5588
   766
by (full_simp_tac (simpset() addsimps 
paulson@5588
   767
  [preal_self_less_add_right] @ preal_add_ac) 1);
paulson@7077
   768
qed "real_of_preal_minus_less_zero";
paulson@5588
   769
paulson@7077
   770
Goal "~ 0r < - real_of_preal m";
paulson@7077
   771
by (cut_facts_tac [real_of_preal_minus_less_zero] 1);
paulson@5588
   772
by (fast_tac (claset() addDs [real_less_trans] 
paulson@5588
   773
                        addEs [real_less_irrefl]) 1);
paulson@7077
   774
qed "real_of_preal_not_minus_gt_zero";
paulson@5588
   775
paulson@7077
   776
Goalw [real_zero_def] "0r < real_of_preal m";
paulson@7077
   777
by (auto_tac (claset(),simpset() addsimps 
paulson@7077
   778
   [real_of_preal_def,real_less_def,real_minus]));
paulson@5588
   779
by (REPEAT(rtac exI 1));
paulson@5588
   780
by (EVERY[rtac conjI 1, rtac conjI 2]);
paulson@5588
   781
by (REPEAT(Blast_tac 2));
paulson@5588
   782
by (full_simp_tac (simpset() addsimps 
paulson@5588
   783
		   [preal_self_less_add_right] @ preal_add_ac) 1);
paulson@7077
   784
qed "real_of_preal_zero_less";
paulson@5588
   785
paulson@7077
   786
Goal "~ real_of_preal m < 0r";
paulson@7077
   787
by (cut_facts_tac [real_of_preal_zero_less] 1);
paulson@5588
   788
by (blast_tac (claset() addDs [real_less_trans] 
paulson@7292
   789
                        addEs [real_less_irrefl]) 1);
paulson@7077
   790
qed "real_of_preal_not_less_zero";
paulson@5588
   791
paulson@7127
   792
Goal "0r < - (- real_of_preal m)";
paulson@5588
   793
by (simp_tac (simpset() addsimps 
paulson@7077
   794
    [real_of_preal_zero_less]) 1);
paulson@5588
   795
qed "real_minus_minus_zero_less";
paulson@5588
   796
paulson@5588
   797
(* another lemma *)
paulson@7077
   798
Goalw [real_zero_def]
paulson@7077
   799
      "0r < real_of_preal m + real_of_preal m1";
paulson@5588
   800
by (auto_tac (claset(),
paulson@7077
   801
	      simpset() addsimps [real_of_preal_def,
paulson@7292
   802
				  real_less_def,real_add]));
paulson@5588
   803
by (REPEAT(rtac exI 1));
paulson@5588
   804
by (EVERY[rtac conjI 1, rtac conjI 2]);
paulson@5588
   805
by (REPEAT(Blast_tac 2));
paulson@5588
   806
by (full_simp_tac (simpset() addsimps preal_add_ac) 1);
paulson@5588
   807
by (full_simp_tac (simpset() addsimps [preal_self_less_add_right,
paulson@5588
   808
    preal_add_assoc RS sym]) 1);
paulson@7077
   809
qed "real_of_preal_sum_zero_less";
paulson@5588
   810
paulson@7077
   811
Goal "- real_of_preal m < real_of_preal m1";
paulson@5588
   812
by (auto_tac (claset(),
paulson@7077
   813
	      simpset() addsimps [real_of_preal_def,
paulson@7077
   814
              real_less_def,real_minus]));
paulson@5588
   815
by (REPEAT(rtac exI 1));
paulson@5588
   816
by (EVERY[rtac conjI 1, rtac conjI 2]);
paulson@5588
   817
by (REPEAT(Blast_tac 2));
paulson@5588
   818
by (full_simp_tac (simpset() addsimps preal_add_ac) 1);
paulson@5588
   819
by (full_simp_tac (simpset() addsimps [preal_self_less_add_right,
paulson@5588
   820
    preal_add_assoc RS sym]) 1);
paulson@7077
   821
qed "real_of_preal_minus_less_all";
paulson@5588
   822
paulson@7077
   823
Goal "~ real_of_preal m < - real_of_preal m1";
paulson@7077
   824
by (cut_facts_tac [real_of_preal_minus_less_all] 1);
paulson@5588
   825
by (blast_tac (claset() addDs [real_less_trans] 
paulson@5588
   826
               addEs [real_less_irrefl]) 1);
paulson@7077
   827
qed "real_of_preal_not_minus_gt_all";
paulson@5588
   828
paulson@7077
   829
Goal "- real_of_preal m1 < - real_of_preal m2 \
paulson@7077
   830
\     ==> real_of_preal m2 < real_of_preal m1";
paulson@5588
   831
by (auto_tac (claset(),
paulson@7077
   832
	      simpset() addsimps [real_of_preal_def,
paulson@7077
   833
              real_less_def,real_minus]));
paulson@5588
   834
by (REPEAT(rtac exI 1));
paulson@5588
   835
by (EVERY[rtac conjI 1, rtac conjI 2]);
paulson@5588
   836
by (REPEAT(Blast_tac 2));
paulson@5588
   837
by (auto_tac (claset(),simpset() addsimps preal_add_ac));
paulson@5588
   838
by (asm_full_simp_tac (simpset() addsimps [preal_add_assoc RS sym]) 1);
paulson@5588
   839
by (auto_tac (claset(),simpset() addsimps preal_add_ac));
paulson@7077
   840
qed "real_of_preal_minus_less_rev1";
paulson@5588
   841
paulson@7077
   842
Goal "real_of_preal m1 < real_of_preal m2 \
paulson@7077
   843
\     ==> - real_of_preal m2 < - real_of_preal m1";
paulson@5588
   844
by (auto_tac (claset(),
paulson@7077
   845
	      simpset() addsimps [real_of_preal_def,
paulson@7077
   846
              real_less_def,real_minus]));
paulson@5588
   847
by (REPEAT(rtac exI 1));
paulson@5588
   848
by (EVERY[rtac conjI 1, rtac conjI 2]);
paulson@5588
   849
by (REPEAT(Blast_tac 2));
paulson@5588
   850
by (auto_tac (claset(),simpset() addsimps preal_add_ac));
paulson@5588
   851
by (asm_full_simp_tac (simpset() addsimps [preal_add_assoc RS sym]) 1);
paulson@5588
   852
by (auto_tac (claset(),simpset() addsimps preal_add_ac));
paulson@7077
   853
qed "real_of_preal_minus_less_rev2";
paulson@5588
   854
paulson@7077
   855
Goal "(- real_of_preal m1 < - real_of_preal m2) = \
paulson@7077
   856
\     (real_of_preal m2 < real_of_preal m1)";
paulson@7077
   857
by (blast_tac (claset() addSIs [real_of_preal_minus_less_rev1,
paulson@7077
   858
               real_of_preal_minus_less_rev2]) 1);
paulson@7077
   859
qed "real_of_preal_minus_less_rev_iff";
paulson@5588
   860
paulson@7077
   861
Addsimps [real_of_preal_minus_less_rev_iff];
paulson@5588
   862
paulson@5588
   863
(*** linearity ***)
paulson@5588
   864
Goal "(R1::real) < R2 | R1 = R2 | R2 < R1";
paulson@7077
   865
by (res_inst_tac [("x","R1")]  real_of_preal_trichotomyE 1);
paulson@7077
   866
by (ALLGOALS(res_inst_tac [("x","R2")]  real_of_preal_trichotomyE));
paulson@5588
   867
by (auto_tac (claset() addSDs [preal_le_anti_sym],
paulson@7077
   868
              simpset() addsimps [preal_less_le_iff,real_of_preal_minus_less_zero,
paulson@7077
   869
               real_of_preal_zero_less,real_of_preal_minus_less_all]));
paulson@5588
   870
qed "real_linear";
paulson@5588
   871
paulson@5588
   872
Goal "!!w::real. (w ~= z) = (w<z | z<w)";
paulson@5588
   873
by (cut_facts_tac [real_linear] 1);
paulson@5588
   874
by (Blast_tac 1);
paulson@5588
   875
qed "real_neq_iff";
paulson@5588
   876
paulson@5588
   877
Goal "!!(R1::real). [| R1 < R2 ==> P;  R1 = R2 ==> P; \
paulson@5588
   878
\                      R2 < R1 ==> P |] ==> P";
paulson@5588
   879
by (cut_inst_tac [("R1.0","R1"),("R2.0","R2")] real_linear 1);
paulson@5588
   880
by Auto_tac;
paulson@5588
   881
qed "real_linear_less2";
paulson@5588
   882
paulson@5588
   883
(*** Properties of <= ***)
paulson@5588
   884
paulson@5588
   885
Goalw [real_le_def] "~(w < z) ==> z <= (w::real)";
paulson@5588
   886
by (assume_tac 1);
paulson@5588
   887
qed "real_leI";
paulson@5588
   888
paulson@5588
   889
Goalw [real_le_def] "z<=w ==> ~(w<(z::real))";
paulson@5588
   890
by (assume_tac 1);
paulson@5588
   891
qed "real_leD";
paulson@5588
   892
wenzelm@7428
   893
bind_thm ("real_leE", make_elim real_leD);
paulson@5588
   894
paulson@5588
   895
Goal "(~(w < z)) = (z <= (w::real))";
paulson@5588
   896
by (blast_tac (claset() addSIs [real_leI,real_leD]) 1);
paulson@5588
   897
qed "real_less_le_iff";
paulson@5588
   898
paulson@5588
   899
Goalw [real_le_def] "~ z <= w ==> w<(z::real)";
paulson@5588
   900
by (Blast_tac 1);
paulson@5588
   901
qed "not_real_leE";
paulson@5588
   902
paulson@5588
   903
Goalw [real_le_def] "z < w ==> z <= (w::real)";
paulson@5588
   904
by (blast_tac (claset() addEs [real_less_asym]) 1);
paulson@5588
   905
qed "real_less_imp_le";
paulson@5588
   906
paulson@5588
   907
Goalw [real_le_def] "!!(x::real). x <= y ==> x < y | x = y";
paulson@5588
   908
by (cut_facts_tac [real_linear] 1);
paulson@5588
   909
by (blast_tac (claset() addEs [real_less_irrefl,real_less_asym]) 1);
paulson@5588
   910
qed "real_le_imp_less_or_eq";
paulson@5588
   911
paulson@5588
   912
Goalw [real_le_def] "z<w | z=w ==> z <=(w::real)";
paulson@5588
   913
by (cut_facts_tac [real_linear] 1);
paulson@5588
   914
by (fast_tac (claset() addEs [real_less_irrefl,real_less_asym]) 1);
paulson@5588
   915
qed "real_less_or_eq_imp_le";
paulson@5588
   916
paulson@5588
   917
Goal "(x <= (y::real)) = (x < y | x=y)";
paulson@5588
   918
by (REPEAT(ares_tac [iffI, real_less_or_eq_imp_le, real_le_imp_less_or_eq] 1));
paulson@5588
   919
qed "real_le_less";
paulson@5588
   920
paulson@5588
   921
Goal "w <= (w::real)";
paulson@5588
   922
by (simp_tac (simpset() addsimps [real_le_less]) 1);
paulson@5588
   923
qed "real_le_refl";
paulson@5588
   924
paulson@5588
   925
AddIffs [real_le_refl];
paulson@5588
   926
paulson@5588
   927
(* Axiom 'linorder_linear' of class 'linorder': *)
paulson@5588
   928
Goal "(z::real) <= w | w <= z";
paulson@5588
   929
by (simp_tac (simpset() addsimps [real_le_less]) 1);
paulson@5588
   930
by (cut_facts_tac [real_linear] 1);
paulson@5588
   931
by (Blast_tac 1);
paulson@5588
   932
qed "real_le_linear";
paulson@5588
   933
paulson@5588
   934
Goal "[| i <= j; j < k |] ==> i < (k::real)";
paulson@5588
   935
by (dtac real_le_imp_less_or_eq 1);
paulson@5588
   936
by (blast_tac (claset() addIs [real_less_trans]) 1);
paulson@5588
   937
qed "real_le_less_trans";
paulson@5588
   938
paulson@5588
   939
Goal "!! (i::real). [| i < j; j <= k |] ==> i < k";
paulson@5588
   940
by (dtac real_le_imp_less_or_eq 1);
paulson@5588
   941
by (blast_tac (claset() addIs [real_less_trans]) 1);
paulson@5588
   942
qed "real_less_le_trans";
paulson@5588
   943
paulson@5588
   944
Goal "[| i <= j; j <= k |] ==> i <= (k::real)";
paulson@5588
   945
by (EVERY1 [dtac real_le_imp_less_or_eq, dtac real_le_imp_less_or_eq,
paulson@5588
   946
            rtac real_less_or_eq_imp_le, blast_tac (claset() addIs [real_less_trans])]);
paulson@5588
   947
qed "real_le_trans";
paulson@5588
   948
paulson@5588
   949
Goal "[| z <= w; w <= z |] ==> z = (w::real)";
paulson@5588
   950
by (EVERY1 [dtac real_le_imp_less_or_eq, dtac real_le_imp_less_or_eq,
paulson@5588
   951
            fast_tac (claset() addEs [real_less_irrefl,real_less_asym])]);
paulson@5588
   952
qed "real_le_anti_sym";
paulson@5588
   953
paulson@5588
   954
Goal "[| ~ y < x; y ~= x |] ==> x < (y::real)";
paulson@5588
   955
by (rtac not_real_leE 1);
paulson@5588
   956
by (blast_tac (claset() addDs [real_le_imp_less_or_eq]) 1);
paulson@5588
   957
qed "not_less_not_eq_real_less";
paulson@5588
   958
paulson@5588
   959
(* Axiom 'order_less_le' of class 'order': *)
paulson@5588
   960
Goal "(w::real) < z = (w <= z & w ~= z)";
paulson@5588
   961
by (simp_tac (simpset() addsimps [real_le_def, real_neq_iff]) 1);
paulson@5588
   962
by (blast_tac (claset() addSEs [real_less_asym]) 1);
paulson@5588
   963
qed "real_less_le";
paulson@5588
   964
paulson@5588
   965
Goal "(0r < -R) = (R < 0r)";
paulson@7077
   966
by (res_inst_tac [("x","R")]  real_of_preal_trichotomyE 1);
paulson@5588
   967
by (auto_tac (claset(),
paulson@7077
   968
	      simpset() addsimps [real_of_preal_not_minus_gt_zero,
paulson@7077
   969
                        real_of_preal_not_less_zero,real_of_preal_zero_less,
paulson@7077
   970
                        real_of_preal_minus_less_zero]));
paulson@5588
   971
qed "real_minus_zero_less_iff";
paulson@5588
   972
paulson@5588
   973
Addsimps [real_minus_zero_less_iff];
paulson@5588
   974
paulson@5588
   975
Goal "(-R < 0r) = (0r < R)";
paulson@7077
   976
by (res_inst_tac [("x","R")]  real_of_preal_trichotomyE 1);
paulson@5588
   977
by (auto_tac (claset(),
paulson@7077
   978
	      simpset() addsimps [real_of_preal_not_minus_gt_zero,
paulson@7077
   979
                        real_of_preal_not_less_zero,real_of_preal_zero_less,
paulson@7077
   980
                        real_of_preal_minus_less_zero]));
paulson@5588
   981
qed "real_minus_zero_less_iff2";
paulson@5588
   982
paulson@5588
   983
(*Alternative definition for real_less*)
paulson@7127
   984
Goal "R < S ==> ? T. 0r < T & R + T = S";
paulson@7077
   985
by (res_inst_tac [("x","R")]  real_of_preal_trichotomyE 1);
paulson@7077
   986
by (ALLGOALS(res_inst_tac [("x","S")]  real_of_preal_trichotomyE));
paulson@5588
   987
by (auto_tac (claset() addSDs [preal_less_add_left_Ex],
paulson@7077
   988
	      simpset() addsimps [real_of_preal_not_minus_gt_all,
paulson@7077
   989
				  real_of_preal_add, real_of_preal_not_less_zero,
paulson@5588
   990
				  real_less_not_refl,
paulson@7077
   991
				  real_of_preal_not_minus_gt_zero]));
paulson@7077
   992
by (res_inst_tac [("x","real_of_preal D")] exI 1);
paulson@7077
   993
by (res_inst_tac [("x","real_of_preal m+real_of_preal ma")] exI 2);
paulson@7077
   994
by (res_inst_tac [("x","real_of_preal m")] exI 3);
paulson@7077
   995
by (res_inst_tac [("x","real_of_preal D")] exI 4);
paulson@5588
   996
by (auto_tac (claset(),
paulson@7077
   997
	      simpset() addsimps [real_of_preal_zero_less,
paulson@7077
   998
				  real_of_preal_sum_zero_less,real_add_assoc]));
paulson@5588
   999
qed "real_less_add_positive_left_Ex";
paulson@5588
  1000
paulson@5588
  1001
(** change naff name(s)! **)
paulson@7127
  1002
Goal "(W < S) ==> (0r < S + (-W))";
paulson@5588
  1003
by (dtac real_less_add_positive_left_Ex 1);
paulson@5588
  1004
by (auto_tac (claset(),
paulson@5588
  1005
	      simpset() addsimps [real_add_minus,
paulson@5588
  1006
    real_add_zero_right] @ real_add_ac));
paulson@5588
  1007
qed "real_less_sum_gt_zero";
paulson@5588
  1008
paulson@7127
  1009
Goal "!!S::real. T = S + W ==> S = T + (-W)";
paulson@5588
  1010
by (asm_simp_tac (simpset() addsimps real_add_ac) 1);
paulson@5588
  1011
qed "real_lemma_change_eq_subj";
paulson@5588
  1012
paulson@5588
  1013
(* FIXME: long! *)
paulson@7127
  1014
Goal "(0r < S + (-W)) ==> (W < S)";
paulson@5588
  1015
by (rtac ccontr 1);
paulson@5588
  1016
by (dtac (real_leI RS real_le_imp_less_or_eq) 1);
paulson@5588
  1017
by (auto_tac (claset(),
paulson@5588
  1018
	      simpset() addsimps [real_less_not_refl]));
paulson@5588
  1019
by (EVERY1[dtac real_less_add_positive_left_Ex, etac exE, etac conjE]);
paulson@5588
  1020
by (Asm_full_simp_tac 1);
paulson@5588
  1021
by (dtac real_lemma_change_eq_subj 1);
paulson@5588
  1022
by Auto_tac;
paulson@5588
  1023
by (dtac real_less_sum_gt_zero 1);
paulson@5588
  1024
by (asm_full_simp_tac (simpset() addsimps real_add_ac) 1);
paulson@5588
  1025
by (EVERY1[rotate_tac 1, dtac (real_add_left_commute RS ssubst)]);
paulson@5588
  1026
by (auto_tac (claset() addEs [real_less_asym], simpset()));
paulson@5588
  1027
qed "real_sum_gt_zero_less";
paulson@5588
  1028
paulson@7127
  1029
Goal "(0r < S + (-W)) = (W < S)";
paulson@5588
  1030
by (blast_tac (claset() addIs [real_less_sum_gt_zero,
paulson@5588
  1031
			       real_sum_gt_zero_less]) 1);
paulson@5588
  1032
qed "real_less_sum_gt_0_iff";
paulson@5588
  1033
paulson@5588
  1034
paulson@5588
  1035
Goalw [real_diff_def] "(x<y) = (x-y < 0r)";
paulson@5588
  1036
by (stac (real_minus_zero_less_iff RS sym) 1);
paulson@5588
  1037
by (simp_tac (simpset() addsimps [real_add_commute,
paulson@5588
  1038
				  real_less_sum_gt_0_iff]) 1);
paulson@5588
  1039
qed "real_less_eq_diff";
paulson@5588
  1040
paulson@5588
  1041
paulson@5588
  1042
(*** Subtraction laws ***)
paulson@5588
  1043
paulson@5588
  1044
Goal "x + (y - z) = (x + y) - (z::real)";
paulson@5588
  1045
by (simp_tac (simpset() addsimps real_diff_def::real_add_ac) 1);
paulson@5588
  1046
qed "real_add_diff_eq";
paulson@5588
  1047
paulson@5588
  1048
Goal "(x - y) + z = (x + z) - (y::real)";
paulson@5588
  1049
by (simp_tac (simpset() addsimps real_diff_def::real_add_ac) 1);
paulson@5588
  1050
qed "real_diff_add_eq";
paulson@5588
  1051
paulson@5588
  1052
Goal "(x - y) - z = x - (y + (z::real))";
paulson@5588
  1053
by (simp_tac (simpset() addsimps real_diff_def::real_add_ac) 1);
paulson@5588
  1054
qed "real_diff_diff_eq";
paulson@5588
  1055
paulson@5588
  1056
Goal "x - (y - z) = (x + z) - (y::real)";
paulson@5588
  1057
by (simp_tac (simpset() addsimps real_diff_def::real_add_ac) 1);
paulson@5588
  1058
qed "real_diff_diff_eq2";
paulson@5588
  1059
paulson@5588
  1060
Goal "(x-y < z) = (x < z + (y::real))";
paulson@5588
  1061
by (stac real_less_eq_diff 1);
paulson@5588
  1062
by (res_inst_tac [("y1", "z")] (real_less_eq_diff RS ssubst) 1);
paulson@5588
  1063
by (simp_tac (simpset() addsimps real_diff_def::real_add_ac) 1);
paulson@5588
  1064
qed "real_diff_less_eq";
paulson@5588
  1065
paulson@5588
  1066
Goal "(x < z-y) = (x + (y::real) < z)";
paulson@5588
  1067
by (stac real_less_eq_diff 1);
paulson@5588
  1068
by (res_inst_tac [("y1", "z-y")] (real_less_eq_diff RS ssubst) 1);
paulson@5588
  1069
by (simp_tac (simpset() addsimps real_diff_def::real_add_ac) 1);
paulson@5588
  1070
qed "real_less_diff_eq";
paulson@5588
  1071
paulson@5588
  1072
Goalw [real_le_def] "(x-y <= z) = (x <= z + (y::real))";
paulson@5588
  1073
by (simp_tac (simpset() addsimps [real_less_diff_eq]) 1);
paulson@5588
  1074
qed "real_diff_le_eq";
paulson@5588
  1075
paulson@5588
  1076
Goalw [real_le_def] "(x <= z-y) = (x + (y::real) <= z)";
paulson@5588
  1077
by (simp_tac (simpset() addsimps [real_diff_less_eq]) 1);
paulson@5588
  1078
qed "real_le_diff_eq";
paulson@5588
  1079
paulson@5588
  1080
Goalw [real_diff_def] "(x-y = z) = (x = z + (y::real))";
paulson@5588
  1081
by (auto_tac (claset(), simpset() addsimps [real_add_assoc]));
paulson@5588
  1082
qed "real_diff_eq_eq";
paulson@5588
  1083
paulson@5588
  1084
Goalw [real_diff_def] "(x = z-y) = (x + (y::real) = z)";
paulson@5588
  1085
by (auto_tac (claset(), simpset() addsimps [real_add_assoc]));
paulson@5588
  1086
qed "real_eq_diff_eq";
paulson@5588
  1087
paulson@5588
  1088
(*This list of rewrites simplifies (in)equalities by bringing subtractions
paulson@5588
  1089
  to the top and then moving negative terms to the other side.  
paulson@5588
  1090
  Use with real_add_ac*)
paulson@5588
  1091
val real_compare_rls = 
paulson@5588
  1092
  [symmetric real_diff_def,
paulson@5588
  1093
   real_add_diff_eq, real_diff_add_eq, real_diff_diff_eq, real_diff_diff_eq2, 
paulson@5588
  1094
   real_diff_less_eq, real_less_diff_eq, real_diff_le_eq, real_le_diff_eq, 
paulson@5588
  1095
   real_diff_eq_eq, real_eq_diff_eq];
paulson@5588
  1096
paulson@5588
  1097
paulson@5588
  1098
(** For the cancellation simproc.
paulson@5588
  1099
    The idea is to cancel like terms on opposite sides by subtraction **)
paulson@5588
  1100
paulson@5588
  1101
Goal "(x::real) - y = x' - y' ==> (x<y) = (x'<y')";
paulson@5588
  1102
by (stac real_less_eq_diff 1);
paulson@5588
  1103
by (res_inst_tac [("y1", "y")] (real_less_eq_diff RS ssubst) 1);
paulson@5588
  1104
by (Asm_simp_tac 1);
paulson@5588
  1105
qed "real_less_eqI";
paulson@5588
  1106
paulson@5588
  1107
Goal "(x::real) - y = x' - y' ==> (y<=x) = (y'<=x')";
paulson@5588
  1108
by (dtac real_less_eqI 1);
paulson@5588
  1109
by (asm_simp_tac (simpset() addsimps [real_le_def]) 1);
paulson@5588
  1110
qed "real_le_eqI";
paulson@5588
  1111
paulson@5588
  1112
Goal "(x::real) - y = x' - y' ==> (x=y) = (x'=y')";
paulson@5588
  1113
by Safe_tac;
paulson@5588
  1114
by (ALLGOALS
paulson@5588
  1115
    (asm_full_simp_tac
paulson@5588
  1116
     (simpset() addsimps [real_eq_diff_eq, real_diff_eq_eq])));
paulson@5588
  1117
qed "real_eq_eqI";