src/HOL/Real/RealArith.thy
author paulson
Thu, 01 Jan 2004 10:06:32 +0100
changeset 14334 6137d24eef79
parent 14329 ff3210fe968f
child 14352 a8b1a44d8264
permissions -rw-r--r--
tweaking of lemmas in RealDef, RealOrd
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
031a5a051bb4 Converting more of the "real" development to Isar scripts
paulson
parents: 14270
diff changeset
     4
use "real_arith.ML"
9436
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
     5
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
     6
setup real_arith_setup
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
     7
14288
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14275
diff changeset
     8
subsection{* Simprules combining x+y and 0: ARE THEY NEEDED?*}
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14275
diff changeset
     9
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14275
diff changeset
    10
text{*Needed in this non-standard form by Hyperreal/Transcendental*}
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14275
diff changeset
    11
lemma real_0_le_divide_iff:
14294
f4d806fd72ce absolute value theorems moved to HOL/Ring_and_Field
paulson
parents: 14293
diff changeset
    12
     "((0::real) \<le> x/y) = ((x \<le> 0 | 0 \<le> y) & (0 \<le> x | y \<le> 0))"
14288
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14275
diff changeset
    13
by (simp add: real_divide_def zero_le_mult_iff, auto)
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14275
diff changeset
    14
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14275
diff changeset
    15
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
    16
by arith
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_eq_0_iff [iff]: "(x+y = (0::real)) = (y = -x)"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14275
diff changeset
    19
by auto
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_less_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_0_less_add_iff [iff]: "((0::real) < x+y) = (-x < y)"
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_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
    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_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
    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
14329
ff3210fe968f re-organized some hyperreal and real lemmas
paulson
parents: 14321
diff changeset
    34
(** Simprules combining x-y and 0 (needed??) **)
14288
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
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
    37
by auto
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14275
diff changeset
    38
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14275
diff changeset
    39
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
    40
by auto
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14275
diff changeset
    41
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14275
diff changeset
    42
(*
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14275
diff changeset
    43
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
    44
It replaces x+-y by x-y.
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14275
diff changeset
    45
Addsimps [symmetric real_diff_def]
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
14321
55c688d2eefa new theorems
paulson
parents: 14308
diff changeset
    48
subsubsection{*Division By @{term "-1"}*}
14288
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14275
diff changeset
    49
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14275
diff changeset
    50
lemma real_divide_minus1 [simp]: "x/-1 = -(x::real)"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14275
diff changeset
    51
by simp
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
lemma real_minus1_divide [simp]: "-1/(x::real) = - (1/x)"
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
    54
by (simp add: real_divide_def inverse_minus_eq)
14288
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14275
diff changeset
    55
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14275
diff changeset
    56
lemma real_lbound_gt_zero:
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14275
diff changeset
    57
     "[| (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
    58
apply (rule_tac x = " (min d1 d2) /2" in exI)
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14275
diff changeset
    59
apply (simp add: min_def)
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14275
diff changeset
    60
done
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14275
diff changeset
    61
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14275
diff changeset
    62
(*** Density of the Reals ***)
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14275
diff changeset
    63
14293
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14290
diff changeset
    64
text{*Similar results are proved in @{text Ring_and_Field}*}
14288
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14275
diff changeset
    65
lemma real_less_half_sum: "x < y ==> x < (x+y) / (2::real)"
14293
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14290
diff changeset
    66
  by auto
14288
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14275
diff changeset
    67
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14275
diff changeset
    68
lemma real_gt_half_sum: "x < y ==> (x+y)/(2::real) < y"
14293
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14290
diff changeset
    69
  by auto
14288
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14275
diff changeset
    70
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14275
diff changeset
    71
lemma real_dense: "x < y ==> \<exists>r::real. x < r & r < y"
14293
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14290
diff changeset
    72
  by (rule Ring_and_Field.dense)
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14290
diff changeset
    73
14288
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14275
diff changeset
    74
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
    75
subsection{*Absolute Value Function for the Reals*}
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
    76
14295
7f115e5c5de4 more general lemmas for Ring_and_Field
paulson
parents: 14294
diff changeset
    77
lemma abs_nat_number_of [simp]: 
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
    78
     "abs (number_of v :: real) =  
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
    79
        (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
    80
         else number_of v)"
14295
7f115e5c5de4 more general lemmas for Ring_and_Field
paulson
parents: 14294
diff changeset
    81
by (simp add: real_abs_def bin_arith_simps minus_real_number_of
7f115e5c5de4 more general lemmas for Ring_and_Field
paulson
parents: 14294
diff changeset
    82
       le_real_number_of_eq_not_less less_real_number_of real_of_int_le_iff)
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
    83
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
    84
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
    85
(*----------------------------------------------------------------------------
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
    86
       Properties of the absolute value function over the reals
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
    87
       (adapted version of previously proved theorems about abs)
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
14294
f4d806fd72ce absolute value theorems moved to HOL/Ring_and_Field
paulson
parents: 14293
diff changeset
    90
text{*FIXME: these should go!*}
f4d806fd72ce absolute value theorems moved to HOL/Ring_and_Field
paulson
parents: 14293
diff changeset
    91
lemma abs_eqI1: "(0::real)\<le>x ==> abs x = x"
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
    92
by (unfold real_abs_def, simp)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
    93
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
    94
lemma abs_eqI2: "(0::real) < x ==> abs x = x"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
    95
by (unfold real_abs_def, simp)
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
lemma abs_minus_eqI2: "x < (0::real) ==> abs x = -x"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
    98
by (unfold real_abs_def real_le_def, simp)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
    99
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   100
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
   101
by (unfold real_abs_def, simp)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   102
14294
f4d806fd72ce absolute value theorems moved to HOL/Ring_and_Field
paulson
parents: 14293
diff changeset
   103
lemma abs_minus_one [simp]: "abs (-1) = (1::real)"
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   104
by (unfold real_abs_def, simp)
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_interval_iff: "(abs x < r) = (-r < x & x < (r::real))"
14295
7f115e5c5de4 more general lemmas for Ring_and_Field
paulson
parents: 14294
diff changeset
   107
by (force simp add: Ring_and_Field.abs_less_iff)
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   108
14294
f4d806fd72ce absolute value theorems moved to HOL/Ring_and_Field
paulson
parents: 14293
diff changeset
   109
lemma abs_le_interval_iff: "(abs x \<le> r) = (-r\<le>x & x\<le>(r::real))"
14295
7f115e5c5de4 more general lemmas for Ring_and_Field
paulson
parents: 14294
diff changeset
   110
by (force simp add: Ring_and_Field.abs_le_iff)
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   111
14295
7f115e5c5de4 more general lemmas for Ring_and_Field
paulson
parents: 14294
diff changeset
   112
lemma abs_add_one_gt_zero [simp]: "(0::real) < 1 + abs(x)"
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   113
by (unfold real_abs_def, auto)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   114
14295
7f115e5c5de4 more general lemmas for Ring_and_Field
paulson
parents: 14294
diff changeset
   115
lemma abs_real_of_nat_cancel [simp]: "abs (real x) = real (x::nat)"
7f115e5c5de4 more general lemmas for Ring_and_Field
paulson
parents: 14294
diff changeset
   116
by (auto intro: abs_eqI1 simp add: real_of_nat_ge_zero)
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   117
14295
7f115e5c5de4 more general lemmas for Ring_and_Field
paulson
parents: 14294
diff changeset
   118
lemma abs_add_one_not_less_self [simp]: "~ abs(x) + (1::real) < x"
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   119
apply (simp add: linorder_not_less)
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   120
apply (auto intro: abs_ge_self [THEN order_trans])
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   121
done
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   122
 
14295
7f115e5c5de4 more general lemmas for Ring_and_Field
paulson
parents: 14294
diff changeset
   123
text{*Used only in Hyperreal/Lim.ML*}
14294
f4d806fd72ce absolute value theorems moved to HOL/Ring_and_Field
paulson
parents: 14293
diff changeset
   124
lemma abs_sum_triangle_ineq: "abs ((x::real) + y + (-l + -m)) \<le> abs(x + -l) + abs(y + -m)"
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   125
apply (simp add: real_add_assoc)
14334
6137d24eef79 tweaking of lemmas in RealDef, RealOrd
paulson
parents: 14329
diff changeset
   126
apply (rule_tac a1 = y in add_left_commute [THEN ssubst])
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   127
apply (rule real_add_assoc [THEN subst])
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   128
apply (rule abs_triangle_ineq)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   129
done
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   130
14288
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14275
diff changeset
   131
14290
84fda1b39947 removal of abel_cancel from Real
paulson
parents: 14288
diff changeset
   132
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   133
ML
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   134
{*
14288
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14275
diff changeset
   135
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
   136
val real_add_minus_iff = thm"real_add_minus_iff";
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14275
diff changeset
   137
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
   138
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
   139
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
   140
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
   141
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
   142
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
   143
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
   144
val real_divide_minus1 = thm"real_divide_minus1";
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14275
diff changeset
   145
val real_minus1_divide = thm"real_minus1_divide";
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14275
diff changeset
   146
val real_lbound_gt_zero = thm"real_lbound_gt_zero";
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14275
diff changeset
   147
val real_less_half_sum = thm"real_less_half_sum";
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14275
diff changeset
   148
val real_gt_half_sum = thm"real_gt_half_sum";
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14275
diff changeset
   149
val real_dense = thm"real_dense";
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14275
diff changeset
   150
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   151
val abs_nat_number_of = thm"abs_nat_number_of";
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   152
val abs_split = thm"abs_split";
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   153
val abs_zero = thm"abs_zero";
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   154
val abs_eqI1 = thm"abs_eqI1";
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   155
val abs_eqI2 = thm"abs_eqI2";
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   156
val abs_minus_eqI2 = thm"abs_minus_eqI2";
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   157
val abs_ge_zero = thm"abs_ge_zero";
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   158
val abs_idempotent = thm"abs_idempotent";
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   159
val abs_zero_iff = thm"abs_zero_iff";
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   160
val abs_ge_self = thm"abs_ge_self";
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   161
val abs_ge_minus_self = thm"abs_ge_minus_self";
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   162
val abs_mult = thm"abs_mult";
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   163
val abs_inverse = thm"abs_inverse";
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   164
val abs_triangle_ineq = thm"abs_triangle_ineq";
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   165
val abs_minus_cancel = thm"abs_minus_cancel";
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   166
val abs_minus_add_cancel = thm"abs_minus_add_cancel";
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   167
val abs_minus_one = thm"abs_minus_one";
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   168
val abs_interval_iff = thm"abs_interval_iff";
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   169
val abs_le_interval_iff = thm"abs_le_interval_iff";
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   170
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
   171
val abs_le_zero_iff = thm"abs_le_zero_iff";
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 10722
diff changeset
   172
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
   173
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
   174
val abs_sum_triangle_ineq = thm"abs_sum_triangle_ineq";
14294
f4d806fd72ce absolute value theorems moved to HOL/Ring_and_Field
paulson
parents: 14293
diff changeset
   175
f4d806fd72ce absolute value theorems moved to HOL/Ring_and_Field
paulson
parents: 14293
diff changeset
   176
val abs_mult_less = thm"abs_mult_less";
14269
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
9436
62bb04ab4b01 rearranged setup of arithmetic procedures, avoiding global reference values;
wenzelm
parents:
diff changeset
   179
end