src/HOL/Real/RealArith.thy
author paulson
Wed, 10 Dec 2003 15:59:34 +0100
changeset 14288 d149e3cbdb39
parent 14275 031a5a051bb4
child 14290 84fda1b39947
permissions -rw-r--r--
Moving some theorems from Real/RealArith0.ML
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
10722
55c8367bab05 rational linear arithmetic
nipkow
parents: 10574
diff changeset
     1
theory RealArith = RealArith0
14275
031a5a051bb4 Converting more of the "real" development to Isar scripts
paulson
parents: 14270
diff changeset
     2
files ("real_arith.ML"):
031a5a051bb4 Converting more of the "real" development to Isar scripts
paulson
parents: 14270
diff changeset
     3
14288
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14275
diff changeset
     4
lemma real_divide_1: "(x::real)/1 = x"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14275
diff changeset
     5
by (simp add: real_divide_def)
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14275
diff changeset
     6
14275
031a5a051bb4 Converting more of the "real" development to Isar scripts
paulson
parents: 14270
diff changeset
     7
use "real_arith.ML"
9436
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
     8
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
     9
setup real_arith_setup
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
    10
14288
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14275
diff changeset
    11
subsection{* Simprules combining x+y and 0: ARE THEY NEEDED?*}
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14275
diff changeset
    12
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14275
diff changeset
    13
text{*Needed in this non-standard form by Hyperreal/Transcendental*}
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14275
diff changeset
    14
lemma real_0_le_divide_iff:
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14275
diff changeset
    15
     "((0::real) <= x/y) = ((x <= 0 | 0 <= y) & (0 <= x | y <= 0))"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14275
diff changeset
    16
by (simp add: real_divide_def zero_le_mult_iff, auto)
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14275
diff changeset
    17
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14275
diff changeset
    18
lemma real_add_minus_iff [simp]: "(x + - a = (0::real)) = (x=a)" 
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14275
diff changeset
    19
by arith
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14275
diff changeset
    20
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14275
diff changeset
    21
lemma real_add_eq_0_iff [iff]: "(x+y = (0::real)) = (y = -x)"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14275
diff changeset
    22
by auto
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14275
diff changeset
    23
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14275
diff changeset
    24
lemma real_add_less_0_iff [iff]: "(x+y < (0::real)) = (y < -x)"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14275
diff changeset
    25
by auto
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14275
diff changeset
    26
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14275
diff changeset
    27
lemma real_0_less_add_iff [iff]: "((0::real) < x+y) = (-x < y)"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14275
diff changeset
    28
by auto
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14275
diff changeset
    29
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14275
diff changeset
    30
lemma real_add_le_0_iff [iff]: "(x+y \<le> (0::real)) = (y \<le> -x)"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14275
diff changeset
    31
by auto
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14275
diff changeset
    32
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14275
diff changeset
    33
lemma real_0_le_add_iff [iff]: "((0::real) \<le> x+y) = (-x \<le> y)"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14275
diff changeset
    34
by auto
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14275
diff changeset
    35
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14275
diff changeset
    36
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14275
diff changeset
    37
(** Simprules combining x-y and 0; see also real_less_iff_diff_less_0, etc.,
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14275
diff changeset
    38
    in RealBin
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14275
diff changeset
    39
**)
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14275
diff changeset
    40
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14275
diff changeset
    41
lemma real_0_less_diff_iff [iff]: "((0::real) < x-y) = (y < x)"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14275
diff changeset
    42
by auto
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14275
diff changeset
    43
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14275
diff changeset
    44
lemma real_0_le_diff_iff [iff]: "((0::real) \<le> x-y) = (y \<le> x)"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14275
diff changeset
    45
by auto
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14275
diff changeset
    46
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14275
diff changeset
    47
(*
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14275
diff changeset
    48
FIXME: we should have this, as for type int, but many proofs would break.
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14275
diff changeset
    49
It replaces x+-y by x-y.
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14275
diff changeset
    50
Addsimps [symmetric real_diff_def]
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14275
diff changeset
    51
*)
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14275
diff changeset
    52
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14275
diff changeset
    53
(** Division by 1, -1 **)
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14275
diff changeset
    54
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14275
diff changeset
    55
lemma real_divide_minus1 [simp]: "x/-1 = -(x::real)"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14275
diff changeset
    56
by simp
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14275
diff changeset
    57
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14275
diff changeset
    58
lemma real_minus1_divide [simp]: "-1/(x::real) = - (1/x)"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14275
diff changeset
    59
by (simp add: real_divide_def real_minus_inverse)
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14275
diff changeset
    60
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14275
diff changeset
    61
lemma real_lbound_gt_zero:
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14275
diff changeset
    62
     "[| (0::real) < d1; 0 < d2 |] ==> \<exists>e. 0 < e & e < d1 & e < d2"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14275
diff changeset
    63
apply (rule_tac x = " (min d1 d2) /2" in exI)
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14275
diff changeset
    64
apply (simp add: min_def)
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14275
diff changeset
    65
done
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14275
diff changeset
    66
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14275
diff changeset
    67
(*** Density of the Reals ***)
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14275
diff changeset
    68
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14275
diff changeset
    69
lemma real_less_half_sum: "x < y ==> x < (x+y) / (2::real)"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14275
diff changeset
    70
by auto
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14275
diff changeset
    71
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14275
diff changeset
    72
lemma real_gt_half_sum: "x < y ==> (x+y)/(2::real) < y"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14275
diff changeset
    73
by auto
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14275
diff changeset
    74
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14275
diff changeset
    75
lemma real_dense: "x < y ==> \<exists>r::real. x < r & r < y"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14275
diff changeset
    76
by (blast intro!: real_less_half_sum real_gt_half_sum)
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14275
diff changeset
    77
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
    78
subsection{*Absolute Value Function for the Reals*}
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
    79
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
    80
(** abs (absolute value) **)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
    81
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
    82
lemma abs_nat_number_of: 
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
    83
     "abs (number_of v :: real) =  
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
    84
        (if neg (number_of v) then number_of (bin_minus v)  
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
    85
         else number_of v)"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
    86
apply (simp add: real_abs_def bin_arith_simps minus_real_number_of le_real_number_of_eq_not_less less_real_number_of real_of_int_le_iff)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
    87
done
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
    88
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
    89
declare abs_nat_number_of [simp]
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
    90
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
    91
lemma abs_split [arith_split]: 
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
    92
  "P(abs (x::real)) = ((0 <= x --> P x) & (x < 0 --> P(-x)))"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
    93
apply (unfold real_abs_def, auto)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
    94
done
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
    95
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
    96
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
    97
(*----------------------------------------------------------------------------
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
    98
       Properties of the absolute value function over the reals
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
    99
       (adapted version of previously proved theorems about abs)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   100
 ----------------------------------------------------------------------------*)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   101
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   102
lemma abs_iff: "abs (r::real) = (if 0<=r then r else -r)"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   103
apply (unfold real_abs_def, auto)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   104
done
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   105
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   106
lemma abs_zero: "abs 0 = (0::real)"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   107
by (unfold real_abs_def, auto)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   108
declare abs_zero [simp]
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   109
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   110
lemma abs_one: "abs (1::real) = 1"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   111
by (unfold real_abs_def, simp)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   112
declare abs_one [simp]
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   113
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   114
lemma abs_eqI1: "(0::real)<=x ==> abs x = x"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   115
by (unfold real_abs_def, simp)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   116
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   117
lemma abs_eqI2: "(0::real) < x ==> abs x = x"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   118
by (unfold real_abs_def, simp)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   119
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   120
lemma abs_minus_eqI2: "x < (0::real) ==> abs x = -x"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   121
by (unfold real_abs_def real_le_def, simp)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   122
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   123
lemma abs_minus_eqI1: "x<=(0::real) ==> abs x = -x"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   124
by (unfold real_abs_def, simp)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   125
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   126
lemma abs_ge_zero: "(0::real)<= abs x"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   127
by (unfold real_abs_def, simp)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   128
declare abs_ge_zero [simp]
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   129
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   130
lemma abs_idempotent: "abs(abs x)=abs (x::real)"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   131
by (unfold real_abs_def, simp)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   132
declare abs_idempotent [simp]
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   133
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   134
lemma abs_zero_iff: "(abs x = 0) = (x=(0::real))"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   135
by (unfold real_abs_def, simp)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   136
declare abs_zero_iff [iff]
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   137
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   138
lemma abs_ge_self: "x<=abs (x::real)"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   139
by (unfold real_abs_def, simp)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   140
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   141
lemma abs_ge_minus_self: "-x<=abs (x::real)"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   142
by (unfold real_abs_def, simp)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   143
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   144
lemma abs_mult: "abs (x * y) = abs x * abs (y::real)"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   145
apply (unfold real_abs_def)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   146
apply (auto dest!: order_antisym simp add: real_0_le_mult_iff)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   147
done
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   148
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   149
lemma abs_inverse: "abs(inverse(x::real)) = inverse(abs(x))"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   150
apply (unfold real_abs_def)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   151
apply (case_tac "x=0")
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   152
apply (auto simp add: real_minus_inverse real_le_less INVERSE_ZERO real_inverse_gt_0)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   153
done
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   154
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   155
lemma abs_mult_inverse: "abs (x * inverse y) = (abs x) * inverse (abs (y::real))"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   156
by (simp add: abs_mult abs_inverse)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   157
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   158
lemma abs_triangle_ineq: "abs(x+y) <= abs x + abs (y::real)"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   159
by (unfold real_abs_def, simp)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   160
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   161
(*Unused, but perhaps interesting as an example*)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   162
lemma abs_triangle_ineq_four: "abs(w + x + y + z) <= abs(w) + abs(x) + abs(y) + abs(z::real)"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   163
by (simp add: abs_triangle_ineq [THEN order_trans])
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   164
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   165
lemma abs_minus_cancel: "abs(-x)=abs(x::real)"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   166
by (unfold real_abs_def, simp)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   167
declare abs_minus_cancel [simp]
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   168
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   169
lemma abs_minus_add_cancel: "abs(x + (-y)) = abs (y + (-(x::real)))"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   170
by (unfold real_abs_def, simp)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   171
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   172
lemma abs_triangle_minus_ineq: "abs(x + (-y)) <= abs x + abs (y::real)"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   173
by (unfold real_abs_def, simp)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   174
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   175
lemma abs_add_less [rule_format (no_asm)]: "abs x < r --> abs y < s --> abs(x+y) < r+(s::real)"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   176
by (unfold real_abs_def, simp)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   177
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   178
lemma abs_add_minus_less: "abs x < r --> abs y < s --> abs(x+ (-y)) < r+(s::real)"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   179
by (unfold real_abs_def, simp)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   180
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   181
(* lemmas manipulating terms *)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   182
lemma real_mult_0_less: "((0::real)*x < r)=(0 < r)"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   183
by simp
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   184
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   185
lemma real_mult_less_trans: "[| (0::real) < y; x < r; y*r < t*s |] ==> y*x < t*s"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   186
by (blast intro!: real_mult_less_mono2 intro: order_less_trans)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   187
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   188
(*USED ONLY IN NEXT THM*)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   189
lemma real_mult_le_less_trans:
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   190
     "[| (0::real)<=y; x < r; y*r < t*s; 0 < t*s|] ==> y*x < t*s"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   191
apply (drule order_le_imp_less_or_eq)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   192
apply (fast  elim: real_mult_0_less [THEN iffD2] real_mult_less_trans) 
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   193
done
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   194
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   195
lemma abs_mult_less: "[| abs x < r; abs y < s |] ==> abs(x*y) < r*(s::real)"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   196
apply (simp add: abs_mult)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   197
apply (rule real_mult_le_less_trans)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   198
apply (rule abs_ge_zero, assumption)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   199
apply (rule_tac [2] real_mult_order)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   200
apply (auto intro!: real_mult_less_mono1 abs_ge_zero intro: order_le_less_trans)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   201
done
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   202
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   203
lemma abs_mult_less2: "[| abs x < r; abs y < s |] ==> abs(x)*abs(y) < r*(s::real)"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   204
by (auto intro: abs_mult_less simp add: abs_mult [symmetric])
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   205
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   206
lemma abs_less_gt_zero: "abs(x) < r ==> (0::real) < r"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   207
by (blast intro!: order_le_less_trans abs_ge_zero)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   208
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   209
lemma abs_minus_one: "abs (-1) = (1::real)"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   210
by (unfold real_abs_def, simp)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   211
declare abs_minus_one [simp]
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   212
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   213
lemma abs_disj: "abs x =x | abs x = -(x::real)"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   214
by (unfold real_abs_def, auto)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   215
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   216
lemma abs_interval_iff: "(abs x < r) = (-r < x & x < (r::real))"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   217
by (unfold real_abs_def, auto)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   218
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   219
lemma abs_le_interval_iff: "(abs x <= r) = (-r<=x & x<=(r::real))"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   220
by (unfold real_abs_def, auto)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   221
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   222
lemma abs_add_pos_gt_zero: "(0::real) < k ==> 0 < k + abs(x)"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   223
by (unfold real_abs_def, auto)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   224
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   225
lemma abs_add_one_gt_zero: "(0::real) < 1 + abs(x)"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   226
by (unfold real_abs_def, auto)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   227
declare abs_add_one_gt_zero [simp]
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   228
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   229
lemma abs_not_less_zero: "~ abs x < (0::real)"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   230
by (unfold real_abs_def, auto)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   231
declare abs_not_less_zero [simp]
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   232
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   233
lemma abs_circle: "abs h < abs y - abs x ==> abs (x + h) < abs (y::real)"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   234
by (auto intro: abs_triangle_ineq [THEN order_le_less_trans])
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   235
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   236
lemma abs_le_zero_iff: "(abs x <= (0::real)) = (x = 0)"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   237
by (unfold real_abs_def, auto)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   238
declare abs_le_zero_iff [simp]
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   239
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   240
lemma real_0_less_abs_iff: "((0::real) < abs x) = (x ~= 0)"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   241
by (simp add: real_abs_def, arith)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   242
declare real_0_less_abs_iff [simp]
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   243
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   244
lemma abs_real_of_nat_cancel: "abs (real x) = real (x::nat)"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   245
by (auto intro: abs_eqI1 simp add: real_of_nat_ge_zero)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   246
declare abs_real_of_nat_cancel [simp]
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   247
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   248
lemma abs_add_one_not_less_self: "~ abs(x) + (1::real) < x"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   249
apply (rule real_leD)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   250
apply (auto intro: abs_ge_self [THEN order_trans])
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   251
done
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   252
declare abs_add_one_not_less_self [simp]
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   253
 
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   254
(* used in vector theory *)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   255
lemma abs_triangle_ineq_three: "abs(w + x + (y::real)) <= abs(w) + abs(x) + abs(y)"
14270
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   256
by (auto intro!: abs_triangle_ineq [THEN order_trans] real_add_left_mono
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   257
                 simp add: real_add_assoc)
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   258
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   259
lemma abs_diff_less_imp_gt_zero: "abs(x - y) < y ==> (0::real) < y"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   260
apply (unfold real_abs_def)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   261
apply (case_tac "0 <= x - y", auto)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   262
done
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   263
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   264
lemma abs_diff_less_imp_gt_zero2: "abs(x - y) < x ==> (0::real) < x"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   265
apply (unfold real_abs_def)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   266
apply (case_tac "0 <= x - y", auto)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   267
done
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   268
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   269
lemma abs_diff_less_imp_gt_zero3: "abs(x - y) < y ==> (0::real) < x"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   270
by (auto simp add: abs_interval_iff)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   271
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   272
lemma abs_diff_less_imp_gt_zero4: "abs(x - y) < -y ==> x < (0::real)"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   273
by (auto simp add: abs_interval_iff)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   274
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   275
lemma abs_triangle_ineq_minus_cancel: 
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   276
     "abs(x) <= abs(x + (-y)) + abs((y::real))"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   277
apply (unfold real_abs_def, auto)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   278
done
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   279
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   280
lemma abs_sum_triangle_ineq: "abs ((x::real) + y + (-l + -m)) <= abs(x + -l) + abs(y + -m)"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   281
apply (simp add: real_add_assoc)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   282
apply (rule_tac x1 = y in real_add_left_commute [THEN ssubst])
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   283
apply (rule real_add_assoc [THEN subst])
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   284
apply (rule abs_triangle_ineq)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   285
done
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   286
14288
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14275
diff changeset
   287
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   288
ML
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   289
{*
14288
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14275
diff changeset
   290
val real_0_le_divide_iff = thm"real_0_le_divide_iff";
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14275
diff changeset
   291
val real_add_minus_iff = thm"real_add_minus_iff";
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14275
diff changeset
   292
val real_add_eq_0_iff = thm"real_add_eq_0_iff";
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14275
diff changeset
   293
val real_add_less_0_iff = thm"real_add_less_0_iff";
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14275
diff changeset
   294
val real_0_less_add_iff = thm"real_0_less_add_iff";
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14275
diff changeset
   295
val real_add_le_0_iff = thm"real_add_le_0_iff";
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14275
diff changeset
   296
val real_0_le_add_iff = thm"real_0_le_add_iff";
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14275
diff changeset
   297
val real_0_less_diff_iff = thm"real_0_less_diff_iff";
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14275
diff changeset
   298
val real_0_le_diff_iff = thm"real_0_le_diff_iff";
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14275
diff changeset
   299
val real_divide_minus1 = thm"real_divide_minus1";
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14275
diff changeset
   300
val real_minus1_divide = thm"real_minus1_divide";
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14275
diff changeset
   301
val real_lbound_gt_zero = thm"real_lbound_gt_zero";
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14275
diff changeset
   302
val real_less_half_sum = thm"real_less_half_sum";
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14275
diff changeset
   303
val real_gt_half_sum = thm"real_gt_half_sum";
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14275
diff changeset
   304
val real_dense = thm"real_dense";
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14275
diff changeset
   305
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   306
val abs_nat_number_of = thm"abs_nat_number_of";
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   307
val abs_split = thm"abs_split";
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   308
val abs_iff = thm"abs_iff";
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   309
val abs_zero = thm"abs_zero";
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   310
val abs_one = thm"abs_one";
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   311
val abs_eqI1 = thm"abs_eqI1";
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   312
val abs_eqI2 = thm"abs_eqI2";
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   313
val abs_minus_eqI2 = thm"abs_minus_eqI2";
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   314
val abs_minus_eqI1 = thm"abs_minus_eqI1";
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   315
val abs_ge_zero = thm"abs_ge_zero";
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   316
val abs_idempotent = thm"abs_idempotent";
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   317
val abs_zero_iff = thm"abs_zero_iff";
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   318
val abs_ge_self = thm"abs_ge_self";
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   319
val abs_ge_minus_self = thm"abs_ge_minus_self";
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   320
val abs_mult = thm"abs_mult";
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   321
val abs_inverse = thm"abs_inverse";
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   322
val abs_mult_inverse = thm"abs_mult_inverse";
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   323
val abs_triangle_ineq = thm"abs_triangle_ineq";
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   324
val abs_triangle_ineq_four = thm"abs_triangle_ineq_four";
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   325
val abs_minus_cancel = thm"abs_minus_cancel";
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   326
val abs_minus_add_cancel = thm"abs_minus_add_cancel";
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   327
val abs_triangle_minus_ineq = thm"abs_triangle_minus_ineq";
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   328
val abs_add_less = thm"abs_add_less";
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   329
val abs_add_minus_less = thm"abs_add_minus_less";
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   330
val real_mult_0_less = thm"real_mult_0_less";
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   331
val real_mult_less_trans = thm"real_mult_less_trans";
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   332
val real_mult_le_less_trans = thm"real_mult_le_less_trans";
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   333
val abs_mult_less = thm"abs_mult_less";
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   334
val abs_mult_less2 = thm"abs_mult_less2";
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   335
val abs_less_gt_zero = thm"abs_less_gt_zero";
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   336
val abs_minus_one = thm"abs_minus_one";
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   337
val abs_disj = thm"abs_disj";
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   338
val abs_interval_iff = thm"abs_interval_iff";
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   339
val abs_le_interval_iff = thm"abs_le_interval_iff";
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   340
val abs_add_pos_gt_zero = thm"abs_add_pos_gt_zero";
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   341
val abs_add_one_gt_zero = thm"abs_add_one_gt_zero";
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   342
val abs_not_less_zero = thm"abs_not_less_zero";
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   343
val abs_circle = thm"abs_circle";
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   344
val abs_le_zero_iff = thm"abs_le_zero_iff";
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   345
val real_0_less_abs_iff = thm"real_0_less_abs_iff";
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   346
val abs_real_of_nat_cancel = thm"abs_real_of_nat_cancel";
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   347
val abs_add_one_not_less_self = thm"abs_add_one_not_less_self";
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   348
val abs_triangle_ineq_three = thm"abs_triangle_ineq_three";
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   349
val abs_diff_less_imp_gt_zero = thm"abs_diff_less_imp_gt_zero";
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   350
val abs_diff_less_imp_gt_zero2 = thm"abs_diff_less_imp_gt_zero2";
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   351
val abs_diff_less_imp_gt_zero3 = thm"abs_diff_less_imp_gt_zero3";
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   352
val abs_diff_less_imp_gt_zero4 = thm"abs_diff_less_imp_gt_zero4";
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   353
val abs_triangle_ineq_minus_cancel = thm"abs_triangle_ineq_minus_cancel";
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   354
val abs_sum_triangle_ineq = thm"abs_sum_triangle_ineq";
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   355
*}
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   356
9436
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   357
end