src/HOL/Real/RealOrd.thy
author paulson
Fri, 28 Nov 2003 12:09:37 +0100
changeset 14269 502a7c95de73
parent 14268 5cf13e80be0e
child 14270 342451d763f9
permissions -rw-r--r--
conversion of some Real theories to Isar scripts
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
9013
9dd0274f76af Updated files to remove 0r and 1r from theorems in descendant theories
fleuriot
parents: 8838
diff changeset
     1
(*  Title:	 Real/RealOrd.thy
9dd0274f76af Updated files to remove 0r and 1r from theorems in descendant theories
fleuriot
parents: 8838
diff changeset
     2
    ID: 	 $Id$
9dd0274f76af Updated files to remove 0r and 1r from theorems in descendant theories
fleuriot
parents: 8838
diff changeset
     3
    Author:      Jacques D. Fleuriot and Lawrence C. Paulson
7334
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
     4
    Copyright:   1998  University of Cambridge
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
     5
*)
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
     6
14265
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
     7
header{*The Reals Form an Ordered Field, etc.*}
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
     8
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
     9
theory RealOrd = RealDef:
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
    10
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
    11
defs (overloaded)
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
    12
  real_abs_def:  "abs (r::real) == (if 0 \<le> r then r else -r)"
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
    13
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
    14
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
    15
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
    16
subsection{* The Simproc @{text abel_cancel}*}
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
    17
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
    18
(*Deletion of other terms in the formula, seeking the -x at the front of z*)
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
    19
lemma real_add_cancel_21: "((x::real) + (y + z) = y + u) = ((x + z) = u)"
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
    20
apply (subst real_add_left_commute)
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
    21
apply (rule real_add_left_cancel)
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
    22
done
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
    23
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
    24
(*A further rule to deal with the case that
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
    25
  everything gets cancelled on the right.*)
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
    26
lemma real_add_cancel_end: "((x::real) + (y + z) = y) = (x = -z)"
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
    27
apply (subst real_add_left_commute)
14266
08b34c902618 conversion of integers to use Ring_and_Field;
paulson
parents: 14265
diff changeset
    28
apply (rule_tac t = y in real_add_zero_right [THEN subst], subst real_add_left_cancel)
08b34c902618 conversion of integers to use Ring_and_Field;
paulson
parents: 14265
diff changeset
    29
apply (simp add: real_eq_diff_eq [symmetric])
14265
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
    30
done
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
    31
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
    32
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
    33
ML
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
    34
{*
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
    35
val real_add_cancel_21 = thm "real_add_cancel_21";
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
    36
val real_add_cancel_end = thm "real_add_cancel_end";
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
    37
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
    38
structure Real_Cancel_Data =
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
    39
struct
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
    40
  val ss		= HOL_ss
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
    41
  val eq_reflection	= eq_reflection
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
    42
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
    43
  val sg_ref		= Sign.self_ref (Theory.sign_of (the_context ()))
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
    44
  val T			= HOLogic.realT
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
    45
  val zero		= Const ("0", T)
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
    46
  val restrict_to_left  = restrict_to_left
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
    47
  val add_cancel_21	= real_add_cancel_21
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
    48
  val add_cancel_end	= real_add_cancel_end
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
    49
  val add_left_cancel	= real_add_left_cancel
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
    50
  val add_assoc		= real_add_assoc
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
    51
  val add_commute	= real_add_commute
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
    52
  val add_left_commute	= real_add_left_commute
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
    53
  val add_0		= real_add_zero_left
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
    54
  val add_0_right	= real_add_zero_right
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
    55
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
    56
  val eq_diff_eq	= real_eq_diff_eq
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
    57
  val eqI_rules		= [real_less_eqI, real_eq_eqI, real_le_eqI]
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
    58
  fun dest_eqI th = 
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
    59
      #1 (HOLogic.dest_bin "op =" HOLogic.boolT 
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
    60
	      (HOLogic.dest_Trueprop (concl_of th)))
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
    61
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
    62
  val diff_def		= real_diff_def
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
    63
  val minus_add_distrib	= real_minus_add_distrib
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
    64
  val minus_minus	= real_minus_minus
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
    65
  val minus_0		= real_minus_zero
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
    66
  val add_inverses	= [real_add_minus, real_add_minus_left]
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
    67
  val cancel_simps	= [real_add_minus_cancel, real_minus_add_cancel]
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
    68
end;
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
    69
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
    70
structure Real_Cancel = Abel_Cancel (Real_Cancel_Data);
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
    71
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
    72
Addsimprocs [Real_Cancel.sum_conv, Real_Cancel.rel_conv];
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
    73
*}
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
    74
14266
08b34c902618 conversion of integers to use Ring_and_Field;
paulson
parents: 14265
diff changeset
    75
lemma real_minus_diff_eq [simp]: "- (z - y) = y - (z::real)"
08b34c902618 conversion of integers to use Ring_and_Field;
paulson
parents: 14265
diff changeset
    76
by simp
14265
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
    77
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
    78
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
    79
subsection{*Theorems About the Ordering*}
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
    80
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
    81
lemma real_gt_zero_preal_Ex: "(0 < x) = (\<exists>y. x = real_of_preal y)"
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
    82
apply (auto simp add: real_of_preal_zero_less)
14266
08b34c902618 conversion of integers to use Ring_and_Field;
paulson
parents: 14265
diff changeset
    83
apply (cut_tac x = x in real_of_preal_trichotomy)
14265
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
    84
apply (blast elim!: real_less_irrefl real_of_preal_not_minus_gt_zero [THEN notE])
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
    85
done
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
    86
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
    87
lemma real_gt_preal_preal_Ex: "real_of_preal z < x ==> \<exists>y. x = real_of_preal y"
14266
08b34c902618 conversion of integers to use Ring_and_Field;
paulson
parents: 14265
diff changeset
    88
by (blast dest!: real_of_preal_zero_less [THEN real_less_trans]
14265
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
    89
             intro: real_gt_zero_preal_Ex [THEN iffD1])
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
    90
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
    91
lemma real_ge_preal_preal_Ex: "real_of_preal z \<le> x ==> \<exists>y. x = real_of_preal y"
14266
08b34c902618 conversion of integers to use Ring_and_Field;
paulson
parents: 14265
diff changeset
    92
by (blast dest: order_le_imp_less_or_eq real_gt_preal_preal_Ex)
14265
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
    93
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
    94
lemma real_less_all_preal: "y \<le> 0 ==> \<forall>x. y < real_of_preal x"
14266
08b34c902618 conversion of integers to use Ring_and_Field;
paulson
parents: 14265
diff changeset
    95
by (auto elim: order_le_imp_less_or_eq [THEN disjE] 
14265
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
    96
            intro: real_of_preal_zero_less [THEN [2] real_less_trans] 
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
    97
            simp add: real_of_preal_zero_less)
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
    98
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
    99
lemma real_less_all_real2: "~ 0 < y ==> \<forall>x. y < real_of_preal x"
14266
08b34c902618 conversion of integers to use Ring_and_Field;
paulson
parents: 14265
diff changeset
   100
by (blast intro!: real_less_all_preal real_leI)
14265
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
   101
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
   102
lemma real_of_preal_le_iff: "(real_of_preal m1 \<le> real_of_preal m2) = (m1 \<le> m2)"
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
   103
apply (auto intro!: preal_leI simp add: real_less_le_iff [symmetric])
14266
08b34c902618 conversion of integers to use Ring_and_Field;
paulson
parents: 14265
diff changeset
   104
apply (drule order_le_less_trans, assumption)
14265
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
   105
apply (erule preal_less_irrefl)
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
   106
done
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
   107
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
   108
subsection{*Monotonicity of Addition*}
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
   109
14266
08b34c902618 conversion of integers to use Ring_and_Field;
paulson
parents: 14265
diff changeset
   110
lemma real_add_right_cancel_less [simp]: "(v+z < w+z) = (v < (w::real))"
08b34c902618 conversion of integers to use Ring_and_Field;
paulson
parents: 14265
diff changeset
   111
by simp
14265
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
   112
14266
08b34c902618 conversion of integers to use Ring_and_Field;
paulson
parents: 14265
diff changeset
   113
lemma real_add_left_cancel_less [simp]: "(z+v < z+w) = (v < (w::real))"
08b34c902618 conversion of integers to use Ring_and_Field;
paulson
parents: 14265
diff changeset
   114
by simp
14265
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
   115
14266
08b34c902618 conversion of integers to use Ring_and_Field;
paulson
parents: 14265
diff changeset
   116
lemma real_add_right_cancel_le [simp]: "(v+z \<le> w+z) = (v \<le> (w::real))"
08b34c902618 conversion of integers to use Ring_and_Field;
paulson
parents: 14265
diff changeset
   117
by simp
14265
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
   118
14266
08b34c902618 conversion of integers to use Ring_and_Field;
paulson
parents: 14265
diff changeset
   119
lemma real_add_left_cancel_le [simp]: "(z+v \<le> z+w) = (v \<le> (w::real))"
08b34c902618 conversion of integers to use Ring_and_Field;
paulson
parents: 14265
diff changeset
   120
by simp
14265
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
   121
14268
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   122
lemma real_mult_order: "[| 0 < x; 0 < y |] ==> (0::real) < x * y"
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   123
apply (auto simp add: real_gt_zero_preal_Ex)
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   124
apply (rule_tac x = "y*ya" in exI)
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   125
apply (simp (no_asm_use) add: real_of_preal_mult)
14265
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
   126
done
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
   127
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
   128
lemma real_mult_less_mono2: "[| (0::real) < z; x < y |] ==> z * x < z * y"
14268
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   129
apply (rule real_sum_gt_zero_less)
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   130
apply (drule real_less_sum_gt_zero [of x y])
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   131
apply (drule real_mult_order, assumption)
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   132
apply (simp add: real_add_mult_distrib2)
14265
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
   133
done
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
   134
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
   135
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
   136
subsection{*The Reals Form an Ordered Field*}
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
   137
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
   138
instance real :: inverse ..
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
   139
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
   140
instance real :: ordered_field
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
   141
proof
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
   142
  fix x y z :: real
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
   143
  show "(x + y) + z = x + (y + z)" by (rule real_add_assoc)
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
   144
  show "x + y = y + x" by (rule real_add_commute)
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
   145
  show "0 + x = x" by simp
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
   146
  show "- x + x = 0" by simp
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
   147
  show "x - y = x + (-y)" by (simp add: real_diff_def)
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
   148
  show "(x * y) * z = x * (y * z)" by (rule real_mult_assoc)
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
   149
  show "x * y = y * x" by (rule real_mult_commute)
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
   150
  show "1 * x = x" by simp
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
   151
  show "(x + y) * z = x * z + y * z" by (simp add: real_add_mult_distrib)
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
   152
  show "0 \<noteq> (1::real)" by (rule real_zero_not_eq_one)
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
   153
  show "x \<le> y ==> z + x \<le> z + y" by simp
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
   154
  show "x < y ==> 0 < z ==> z * x < z * y" by (simp add: real_mult_less_mono2)
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
   155
  show "\<bar>x\<bar> = (if x < 0 then -x else x)"
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
   156
    by (auto dest: order_le_less_trans simp add: real_abs_def linorder_not_le)
14266
08b34c902618 conversion of integers to use Ring_and_Field;
paulson
parents: 14265
diff changeset
   157
  show "x \<noteq> 0 ==> inverse x * x = 1" by simp
14265
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
   158
  show "y \<noteq> 0 ==> x / y = x * inverse y" by (simp add: real_divide_def)
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
   159
qed
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
   160
14268
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   161
(*"v\<le>w ==> v+z \<le> w+z"*)
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   162
lemmas real_add_less_mono1 = real_add_right_cancel_less [THEN iffD2, standard]
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   163
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   164
(*"v\<le>w ==> v+z \<le> w+z"*)
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   165
lemmas real_add_le_mono1 = real_add_right_cancel_le [THEN iffD2, standard]
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   166
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   167
lemma real_add_less_le_mono: "!!z z'::real. [| w'<w; z'\<le>z |] ==> w' + z' < w + z"
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   168
by (erule real_add_less_mono1 [THEN order_less_le_trans], simp)
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   169
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   170
lemma real_add_le_less_mono: "!!z z'::real. [| w'\<le>w; z'<z |] ==> w' + z' < w + z"
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   171
by (erule real_add_le_mono1 [THEN order_le_less_trans], simp)
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   172
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   173
lemma real_add_less_mono2: "!!(A::real). A < B ==> C + A < C + B"
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   174
by simp
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   175
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   176
lemma real_less_add_right_cancel: "!!(A::real). A + C < B + C ==> A < B"
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   177
apply simp
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   178
done
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   179
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   180
lemma real_less_add_left_cancel: "!!(A::real). C + A < C + B ==> A < B"
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   181
apply simp
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   182
done
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   183
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   184
lemma real_le_add_right_cancel: "!!(A::real). A + C \<le> B + C ==> A \<le> B"
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   185
apply simp
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   186
done
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   187
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   188
lemma real_le_add_left_cancel: "!!(A::real). C + A \<le> C + B ==> A \<le> B"
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   189
apply simp
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   190
done
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   191
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   192
lemma real_zero_less_one: "0 < (1::real)"
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   193
  by (rule Ring_and_Field.zero_less_one)
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   194
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   195
lemma real_add_less_mono: "[| R1 < S1; R2 < S2 |] ==> R1+R2 < S1+(S2::real)"
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   196
 by (rule Ring_and_Field.add_strict_mono)
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   197
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   198
lemma real_add_left_le_mono1: "!!(q1::real). q1 \<le> q2  ==> x + q1 \<le> x + q2"
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   199
by simp
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   200
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   201
lemma real_add_le_mono: "[|i\<le>j;  k\<le>l |] ==> i + k \<le> j + (l::real)"
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   202
 by (rule Ring_and_Field.add_mono)
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   203
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   204
lemma real_le_minus_iff: "(-s \<le> -r) = ((r::real) \<le> s)"
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   205
 by (rule Ring_and_Field.neg_le_iff_le)
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   206
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   207
lemma real_le_square [simp]: "(0::real) \<le> x*x"
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   208
 by (rule Ring_and_Field.zero_le_square)
14265
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
   209
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
   210
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 14268
diff changeset
   211
subsection{*Division Lemmas*}
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 14268
diff changeset
   212
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 14268
diff changeset
   213
(** Inverse of zero!  Useful to simplify certain equations **)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 14268
diff changeset
   214
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 14268
diff changeset
   215
lemma INVERSE_ZERO [simp]: "inverse 0 = (0::real)"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 14268
diff changeset
   216
apply (unfold real_inverse_def)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 14268
diff changeset
   217
apply (rule someI2)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 14268
diff changeset
   218
apply (auto simp add: real_zero_not_eq_one)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 14268
diff changeset
   219
done
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 14268
diff changeset
   220
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 14268
diff changeset
   221
lemma DIVISION_BY_ZERO [simp]: "a / (0::real) = 0"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 14268
diff changeset
   222
by (simp add: real_divide_def INVERSE_ZERO)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 14268
diff changeset
   223
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 14268
diff changeset
   224
lemma real_mult_left_cancel: "(c::real) ~= 0 ==> (c*a=c*b) = (a=b)"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 14268
diff changeset
   225
apply auto
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 14268
diff changeset
   226
done
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 14268
diff changeset
   227
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 14268
diff changeset
   228
lemma real_mult_right_cancel: "(c::real) ~= 0 ==> (a*c=b*c) = (a=b)"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 14268
diff changeset
   229
apply (auto ); 
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 14268
diff changeset
   230
done
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 14268
diff changeset
   231
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 14268
diff changeset
   232
lemma real_mult_left_cancel_ccontr: "c*a ~= c*b ==> a ~= b"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 14268
diff changeset
   233
by auto
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 14268
diff changeset
   234
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 14268
diff changeset
   235
lemma real_mult_right_cancel_ccontr: "a*c ~= b*c ==> a ~= b"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 14268
diff changeset
   236
by auto
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 14268
diff changeset
   237
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 14268
diff changeset
   238
lemma real_inverse_not_zero: "x ~= 0 ==> inverse(x::real) ~= 0"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 14268
diff changeset
   239
  by (rule Ring_and_Field.nonzero_imp_inverse_nonzero)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 14268
diff changeset
   240
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 14268
diff changeset
   241
lemma real_mult_not_zero: "[| x ~= 0; y ~= 0 |] ==> x * y ~= (0::real)"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 14268
diff changeset
   242
apply (simp add: ); 
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 14268
diff changeset
   243
done
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 14268
diff changeset
   244
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 14268
diff changeset
   245
lemma real_inverse_inverse: "inverse(inverse (x::real)) = x"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 14268
diff changeset
   246
apply (case_tac "x=0", simp)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 14268
diff changeset
   247
apply (rule_tac c1 = "inverse x" in real_mult_right_cancel [THEN iffD1])
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 14268
diff changeset
   248
apply (erule real_inverse_not_zero)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 14268
diff changeset
   249
apply (auto dest: real_inverse_not_zero)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 14268
diff changeset
   250
done
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 14268
diff changeset
   251
declare real_inverse_inverse [simp]
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 14268
diff changeset
   252
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 14268
diff changeset
   253
lemma real_inverse_1: "inverse((1::real)) = (1::real)"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 14268
diff changeset
   254
apply (unfold real_inverse_def)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 14268
diff changeset
   255
apply (cut_tac real_zero_not_eq_one [THEN not_sym, THEN real_mult_inv_left_ex])
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 14268
diff changeset
   256
apply (auto simp add: real_zero_not_eq_one [THEN not_sym])
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 14268
diff changeset
   257
done
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 14268
diff changeset
   258
declare real_inverse_1 [simp]
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 14268
diff changeset
   259
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 14268
diff changeset
   260
lemma real_minus_inverse: "inverse(-x) = -inverse(x::real)"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 14268
diff changeset
   261
apply (case_tac "x=0", simp)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 14268
diff changeset
   262
apply (rule_tac c1 = "-x" in real_mult_right_cancel [THEN iffD1])
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 14268
diff changeset
   263
 prefer 2 apply (subst real_mult_inv_left, auto)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 14268
diff changeset
   264
done
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 14268
diff changeset
   265
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 14268
diff changeset
   266
lemma real_inverse_distrib: "inverse(x*y) = inverse(x)*inverse(y::real)"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 14268
diff changeset
   267
apply (case_tac "x=0", simp)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 14268
diff changeset
   268
apply (case_tac "y=0", simp)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 14268
diff changeset
   269
apply (frule_tac y = y in real_mult_not_zero, assumption)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 14268
diff changeset
   270
apply (rule_tac c1 = x in real_mult_left_cancel [THEN iffD1])
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 14268
diff changeset
   271
apply (auto simp add: real_mult_assoc [symmetric])
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 14268
diff changeset
   272
apply (rule_tac c1 = y in real_mult_left_cancel [THEN iffD1])
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 14268
diff changeset
   273
apply (auto simp add: real_mult_left_commute)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 14268
diff changeset
   274
apply (simp add: real_mult_assoc [symmetric])
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 14268
diff changeset
   275
done
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 14268
diff changeset
   276
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 14268
diff changeset
   277
lemma real_times_divide1_eq: "(x::real) * (y/z) = (x*y)/z"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 14268
diff changeset
   278
by (simp add: real_divide_def real_mult_assoc)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 14268
diff changeset
   279
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 14268
diff changeset
   280
lemma real_times_divide2_eq: "(y/z) * (x::real) = (y*x)/z"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 14268
diff changeset
   281
by (simp add: real_divide_def real_mult_ac)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 14268
diff changeset
   282
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 14268
diff changeset
   283
declare real_times_divide1_eq [simp] real_times_divide2_eq [simp]
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 14268
diff changeset
   284
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 14268
diff changeset
   285
lemma real_divide_divide1_eq: "(x::real) / (y/z) = (x*z)/y"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 14268
diff changeset
   286
by (simp add: real_divide_def real_inverse_distrib real_mult_ac)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 14268
diff changeset
   287
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 14268
diff changeset
   288
lemma real_divide_divide2_eq: "((x::real) / y) / z = x/(y*z)"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 14268
diff changeset
   289
by (simp add: real_divide_def real_inverse_distrib real_mult_assoc)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 14268
diff changeset
   290
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 14268
diff changeset
   291
declare real_divide_divide1_eq [simp] real_divide_divide2_eq [simp]
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 14268
diff changeset
   292
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 14268
diff changeset
   293
(** As with multiplication, pull minus signs OUT of the / operator **)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 14268
diff changeset
   294
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 14268
diff changeset
   295
lemma real_minus_divide_eq: "(-x) / (y::real) = - (x/y)"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 14268
diff changeset
   296
by (simp add: real_divide_def)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 14268
diff changeset
   297
declare real_minus_divide_eq [simp]
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 14268
diff changeset
   298
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 14268
diff changeset
   299
lemma real_divide_minus_eq: "(x / -(y::real)) = - (x/y)"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 14268
diff changeset
   300
by (simp add: real_divide_def real_minus_inverse)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 14268
diff changeset
   301
declare real_divide_minus_eq [simp]
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 14268
diff changeset
   302
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 14268
diff changeset
   303
lemma real_add_divide_distrib: "(x+y)/(z::real) = x/z + y/z"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 14268
diff changeset
   304
by (simp add: real_divide_def real_add_mult_distrib)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 14268
diff changeset
   305
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 14268
diff changeset
   306
(*The following would e.g. convert (x+y)/2 to x/2 + y/2.  Sometimes that
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 14268
diff changeset
   307
  leads to cancellations in x or y, but is also prevents "multiplying out"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 14268
diff changeset
   308
  the 2 in e.g. (x+y)/2 = 5.
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 14268
diff changeset
   309
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 14268
diff changeset
   310
Addsimps [inst "z" "number_of ?w" real_add_divide_distrib]
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 14268
diff changeset
   311
**)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 14268
diff changeset
   312
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 14268
diff changeset
   313
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 14268
diff changeset
   314
14268
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   315
subsection{*An Embedding of the Naturals in the Reals*}
14265
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
   316
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
   317
lemma real_of_posnat_one: "real_of_posnat 0 = (1::real)"
14266
08b34c902618 conversion of integers to use Ring_and_Field;
paulson
parents: 14265
diff changeset
   318
by (simp add: real_of_posnat_def pnat_one_iff [symmetric]
08b34c902618 conversion of integers to use Ring_and_Field;
paulson
parents: 14265
diff changeset
   319
              real_of_preal_def symmetric real_one_def)
14265
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
   320
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
   321
lemma real_of_posnat_two: "real_of_posnat (Suc 0) = (1::real) + (1::real)"
14266
08b34c902618 conversion of integers to use Ring_and_Field;
paulson
parents: 14265
diff changeset
   322
by (simp add: real_of_posnat_def real_of_preal_def real_one_def pnat_two_eq
14265
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
   323
                 real_add
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
   324
            prat_of_pnat_add [symmetric] preal_of_prat_add [symmetric]
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
   325
            pnat_add_ac)
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
   326
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
   327
lemma real_of_posnat_add: 
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
   328
     "real_of_posnat n1 + real_of_posnat n2 = real_of_posnat (n1 + n2) + (1::real)"
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
   329
apply (unfold real_of_posnat_def)
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
   330
apply (simp (no_asm_use) add: real_of_posnat_one [symmetric] real_of_posnat_def real_of_preal_add [symmetric] preal_of_prat_add [symmetric] prat_of_pnat_add [symmetric] pnat_of_nat_add)
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
   331
done
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
   332
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
   333
lemma real_of_posnat_add_one: "real_of_posnat (n + 1) = real_of_posnat n + (1::real)"
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
   334
apply (rule_tac x1 = " (1::real) " in real_add_right_cancel [THEN iffD1])
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
   335
apply (rule real_of_posnat_add [THEN subst])
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
   336
apply (simp (no_asm_use) add: real_of_posnat_two real_add_assoc)
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
   337
done
7334
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   338
14265
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
   339
lemma real_of_posnat_Suc: "real_of_posnat (Suc n) = real_of_posnat n + (1::real)"
14266
08b34c902618 conversion of integers to use Ring_and_Field;
paulson
parents: 14265
diff changeset
   340
by (subst real_of_posnat_add_one [symmetric], simp)
14265
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
   341
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
   342
lemma inj_real_of_posnat: "inj(real_of_posnat)"
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
   343
apply (rule inj_onI)
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
   344
apply (unfold real_of_posnat_def)
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
   345
apply (drule inj_real_of_preal [THEN injD])
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
   346
apply (drule inj_preal_of_prat [THEN injD])
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
   347
apply (drule inj_prat_of_pnat [THEN injD])
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
   348
apply (erule inj_pnat_of_nat [THEN injD])
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
   349
done
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
   350
14266
08b34c902618 conversion of integers to use Ring_and_Field;
paulson
parents: 14265
diff changeset
   351
lemma real_of_nat_zero [simp]: "real (0::nat) = 0"
08b34c902618 conversion of integers to use Ring_and_Field;
paulson
parents: 14265
diff changeset
   352
by (simp add: real_of_nat_def real_of_posnat_one)
14265
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
   353
14266
08b34c902618 conversion of integers to use Ring_and_Field;
paulson
parents: 14265
diff changeset
   354
lemma real_of_nat_one [simp]: "real (Suc 0) = (1::real)"
08b34c902618 conversion of integers to use Ring_and_Field;
paulson
parents: 14265
diff changeset
   355
by (simp add: real_of_nat_def real_of_posnat_two real_add_assoc)
14265
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
   356
14266
08b34c902618 conversion of integers to use Ring_and_Field;
paulson
parents: 14265
diff changeset
   357
lemma real_of_nat_add [simp]: 
14265
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
   358
     "real (m + n) = real (m::nat) + real n"
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
   359
apply (simp add: real_of_nat_def real_add_assoc)
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
   360
apply (simp add: real_of_posnat_add real_add_assoc [symmetric])
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
   361
done
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
   362
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
   363
(*Not for addsimps: often the LHS is used to represent a positive natural*)
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
   364
lemma real_of_nat_Suc: "real (Suc n) = real n + (1::real)"
14266
08b34c902618 conversion of integers to use Ring_and_Field;
paulson
parents: 14265
diff changeset
   365
by (simp add: real_of_nat_def real_of_posnat_Suc real_add_ac)
14265
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
   366
14266
08b34c902618 conversion of integers to use Ring_and_Field;
paulson
parents: 14265
diff changeset
   367
lemma real_of_nat_less_iff [iff]: 
14265
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
   368
     "(real (n::nat) < real m) = (n < m)"
14266
08b34c902618 conversion of integers to use Ring_and_Field;
paulson
parents: 14265
diff changeset
   369
by (auto simp add: real_of_nat_def real_of_posnat_def)
14265
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
   370
14266
08b34c902618 conversion of integers to use Ring_and_Field;
paulson
parents: 14265
diff changeset
   371
lemma real_of_nat_le_iff [iff]: "(real (n::nat) \<le> real m) = (n \<le> m)"
08b34c902618 conversion of integers to use Ring_and_Field;
paulson
parents: 14265
diff changeset
   372
by (simp add: linorder_not_less [symmetric])
14265
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
   373
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
   374
lemma inj_real_of_nat: "inj (real :: nat => real)"
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
   375
apply (rule inj_onI)
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
   376
apply (auto intro!: inj_real_of_posnat [THEN injD]
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
   377
            simp add: real_of_nat_def real_add_right_cancel)
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
   378
done
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
   379
14266
08b34c902618 conversion of integers to use Ring_and_Field;
paulson
parents: 14265
diff changeset
   380
lemma real_of_nat_ge_zero [iff]: "0 \<le> real (n::nat)"
14265
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
   381
apply (induct_tac "n")
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
   382
apply (auto simp add: real_of_nat_Suc)
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
   383
apply (drule real_add_le_less_mono)
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
   384
apply (rule real_zero_less_one)
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
   385
apply (simp add: order_less_imp_le)
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
   386
done
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
   387
14266
08b34c902618 conversion of integers to use Ring_and_Field;
paulson
parents: 14265
diff changeset
   388
lemma real_of_nat_mult [simp]: "real (m * n) = real (m::nat) * real n"
14265
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
   389
apply (induct_tac "m")
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
   390
apply (auto simp add: real_of_nat_Suc real_add_mult_distrib real_add_commute)
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
   391
done
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
   392
14266
08b34c902618 conversion of integers to use Ring_and_Field;
paulson
parents: 14265
diff changeset
   393
lemma real_of_nat_inject [iff]: "(real (n::nat) = real m) = (n = m)"
08b34c902618 conversion of integers to use Ring_and_Field;
paulson
parents: 14265
diff changeset
   394
by (auto dest: inj_real_of_nat [THEN injD])
14265
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
   395
14266
08b34c902618 conversion of integers to use Ring_and_Field;
paulson
parents: 14265
diff changeset
   396
lemma real_of_nat_diff [rule_format]:
08b34c902618 conversion of integers to use Ring_and_Field;
paulson
parents: 14265
diff changeset
   397
     "n \<le> m --> real (m - n) = real (m::nat) - real n"
14265
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
   398
apply (induct_tac "m")
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
   399
apply (auto simp add: real_diff_def Suc_diff_le le_Suc_eq real_of_nat_Suc real_of_nat_zero real_add_ac)
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
   400
done
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
   401
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
   402
lemma real_of_nat_zero_iff: "(real (n::nat) = 0) = (n = 0)"
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
   403
  proof 
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
   404
    assume "real n = 0"
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
   405
    have "real n = real (0::nat)" by simp
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
   406
    then show "n = 0" by (simp only: real_of_nat_inject)
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
   407
  next
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
   408
    show "n = 0 \<Longrightarrow> real n = 0" by simp
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
   409
  qed
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
   410
14266
08b34c902618 conversion of integers to use Ring_and_Field;
paulson
parents: 14265
diff changeset
   411
lemma real_of_nat_neg_int [simp]: "neg z ==> real (nat z) = 0"
14268
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   412
apply (simp add: neg_nat real_of_nat_zero)
14265
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
   413
done
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
   414
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
   415
14268
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   416
subsection{*Inverse and Division*}
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   417
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   418
instance real :: division_by_zero
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   419
proof
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   420
  fix x :: real
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   421
  show "inverse 0 = (0::real)" by (rule INVERSE_ZERO)
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   422
  show "x/0 = 0" by (rule DIVISION_BY_ZERO) 
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   423
qed
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   424
14265
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
   425
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
   426
lemma real_inverse_gt_0: "0 < x ==> 0 < inverse (x::real)"
14268
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   427
  by (rule Ring_and_Field.inverse_gt_0)
14265
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
   428
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
   429
lemma real_inverse_less_0: "x < 0 ==> inverse (x::real) < 0"
14268
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   430
  by (rule Ring_and_Field.inverse_less_0)
14265
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
   431
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
   432
lemma real_mult_less_cancel1: "[| (0::real) < z; x * z < y * z |] ==> x < y"
14266
08b34c902618 conversion of integers to use Ring_and_Field;
paulson
parents: 14265
diff changeset
   433
by (force simp add: Ring_and_Field.mult_less_cancel_right 
08b34c902618 conversion of integers to use Ring_and_Field;
paulson
parents: 14265
diff changeset
   434
          elim: order_less_asym) 
14265
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
   435
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
   436
lemma real_mult_less_cancel2: "[| (0::real) < z; z*x < z*y |] ==> x < y"
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
   437
apply (erule real_mult_less_cancel1)
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
   438
apply (simp add: real_mult_commute)
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
   439
done
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
   440
14268
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   441
lemma real_mult_less_mono1: "[| (0::real) < z; x < y |] ==> x*z < y*z"
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   442
 by (rule Ring_and_Field.mult_strict_right_mono)
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   443
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   444
lemma real_mult_less_mono:
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   445
     "[| u<v;  x<y;  (0::real) < v;  0 < x |] ==> u*x < v* y"
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   446
 by (simp add: Ring_and_Field.mult_strict_mono order_less_imp_le)
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   447
14266
08b34c902618 conversion of integers to use Ring_and_Field;
paulson
parents: 14265
diff changeset
   448
lemma real_mult_less_iff1 [simp]: "(0::real) < z ==> (x*z < y*z) = (x < y)"
08b34c902618 conversion of integers to use Ring_and_Field;
paulson
parents: 14265
diff changeset
   449
by (blast intro: real_mult_less_mono1 real_mult_less_cancel1)
14265
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
   450
14266
08b34c902618 conversion of integers to use Ring_and_Field;
paulson
parents: 14265
diff changeset
   451
lemma real_mult_less_iff2 [simp]: "(0::real) < z ==> (z*x < z*y) = (x < y)"
08b34c902618 conversion of integers to use Ring_and_Field;
paulson
parents: 14265
diff changeset
   452
by (blast intro: real_mult_less_mono2 real_mult_less_cancel2)
14265
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
   453
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
   454
(* 05/00 *)
14268
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   455
lemma real_mult_le_cancel_iff1 [simp]: "(0::real) < z ==> (x*z \<le> y*z) = (x\<le>y)"
14266
08b34c902618 conversion of integers to use Ring_and_Field;
paulson
parents: 14265
diff changeset
   456
by (auto simp add: real_le_def)
14265
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
   457
14268
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   458
lemma real_mult_le_cancel_iff2 [simp]: "(0::real) < z ==> (z*x \<le> z*y) = (x\<le>y)"
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   459
by (auto simp add: real_le_def)
14265
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
   460
14266
08b34c902618 conversion of integers to use Ring_and_Field;
paulson
parents: 14265
diff changeset
   461
lemma real_mult_less_mono':
08b34c902618 conversion of integers to use Ring_and_Field;
paulson
parents: 14265
diff changeset
   462
     "[| x < y;  r1 < r2;  (0::real) \<le> r1;  0 \<le> x|] ==> r1 * x < r2 * y"
14265
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
   463
apply (subgoal_tac "0<r2")
14266
08b34c902618 conversion of integers to use Ring_and_Field;
paulson
parents: 14265
diff changeset
   464
 prefer 2 apply (blast intro: order_le_less_trans)
14265
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
   465
apply (case_tac "x=0")
14266
08b34c902618 conversion of integers to use Ring_and_Field;
paulson
parents: 14265
diff changeset
   466
apply (auto dest!: order_le_imp_less_or_eq 
08b34c902618 conversion of integers to use Ring_and_Field;
paulson
parents: 14265
diff changeset
   467
            intro: real_mult_less_mono real_mult_order)
14265
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
   468
done
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
   469
14268
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   470
lemma real_inverse_less_swap:
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   471
     "[| 0 < r; r < x |] ==> inverse x < inverse (r::real)"
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   472
  by (rule Ring_and_Field.less_imp_inverse_less)
14265
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
   473
14266
08b34c902618 conversion of integers to use Ring_and_Field;
paulson
parents: 14265
diff changeset
   474
lemma real_mult_is_0 [iff]: "(x*y = 0) = (x = 0 | y = (0::real))"
08b34c902618 conversion of integers to use Ring_and_Field;
paulson
parents: 14265
diff changeset
   475
by auto
14265
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
   476
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
   477
lemma real_inverse_add: "[| x \<noteq> 0; y \<noteq> 0 |]  
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
   478
      ==> inverse x + inverse y = (x + y) * inverse (x * (y::real))"
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
   479
apply (simp add: real_inverse_distrib real_add_mult_distrib real_mult_assoc [symmetric])
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
   480
apply (subst real_mult_assoc)
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
   481
apply (rule real_mult_left_commute [THEN subst])
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
   482
apply (simp add: real_add_commute)
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
   483
done
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
   484
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
   485
lemma real_sum_squares_cancel: "x * x + y * y = 0 ==> x = (0::real)"
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
   486
apply (drule real_add_minus_eq_minus)
14266
08b34c902618 conversion of integers to use Ring_and_Field;
paulson
parents: 14265
diff changeset
   487
apply (cut_tac x = x in real_le_square)
08b34c902618 conversion of integers to use Ring_and_Field;
paulson
parents: 14265
diff changeset
   488
apply (auto, drule real_le_anti_sym, auto)
14265
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
   489
done
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
   490
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
   491
lemma real_sum_squares_cancel2: "x * x + y * y = 0 ==> y = (0::real)"
14266
08b34c902618 conversion of integers to use Ring_and_Field;
paulson
parents: 14265
diff changeset
   492
apply (rule_tac y = x in real_sum_squares_cancel)
14265
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
   493
apply (simp add: real_add_commute)
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
   494
done
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
   495
14268
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   496
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   497
subsection{*Convenient Biconditionals for Products of Signs*}
14265
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
   498
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
   499
lemma real_0_less_mult_iff: "((0::real) < x*y) = (0<x & 0<y | x<0 & y<0)"
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
   500
  by (rule Ring_and_Field.zero_less_mult_iff) 
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
   501
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
   502
lemma real_0_le_mult_iff: "((0::real)\<le>x*y) = (0\<le>x & 0\<le>y | x\<le>0 & y\<le>0)"
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
   503
  by (rule zero_le_mult_iff) 
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
   504
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
   505
lemma real_mult_less_0_iff: "(x*y < (0::real)) = (0<x & y<0 | x<0 & 0<y)"
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
   506
  by (rule mult_less_0_iff) 
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
   507
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
   508
lemma real_mult_le_0_iff: "(x*y \<le> (0::real)) = (0\<le>x & y\<le>0 | x\<le>0 & 0\<le>y)"
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
   509
  by (rule mult_le_0_iff) 
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
   510
14268
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   511
subsection{*Hardly Used Theorems to be Deleted*}
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   512
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   513
lemma real_add_order: "[| 0 < x; 0 < y |] ==> (0::real) < x + y"
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   514
apply (erule order_less_trans)
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   515
apply (drule real_add_less_mono2)
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   516
apply simp
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   517
done
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   518
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   519
lemma real_le_add_order: "[| 0 \<le> x; 0 \<le> y |] ==> (0::real) \<le> x + y"
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   520
apply (drule order_le_imp_less_or_eq)+
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   521
apply (auto intro: real_add_order order_less_imp_le)
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   522
done
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   523
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   524
(*One use in HahnBanach/Aux.thy*)
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   525
lemma real_mult_le_less_mono1: "[| (0::real) \<le> z; x < y |] ==> x*z \<le> y*z"
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   526
apply (rule real_less_or_eq_imp_le)
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   527
apply (drule order_le_imp_less_or_eq)
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   528
apply (auto intro: real_mult_less_mono1)
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   529
done
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   530
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   531
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   532
lemma real_of_posnat_gt_zero: "0 < real_of_posnat n"
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   533
apply (unfold real_of_posnat_def)
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   534
apply (rule real_gt_zero_preal_Ex [THEN iffD2])
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   535
apply blast
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   536
done
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   537
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   538
declare real_of_posnat_gt_zero [simp]
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   539
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   540
lemmas real_inv_real_of_posnat_gt_zero =  real_of_posnat_gt_zero [THEN real_inverse_gt_0, standard]
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   541
declare real_inv_real_of_posnat_gt_zero [simp]
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   542
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   543
lemmas real_of_posnat_ge_zero = real_of_posnat_gt_zero [THEN order_less_imp_le, standard]
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   544
declare real_of_posnat_ge_zero [simp]
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   545
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   546
lemma real_of_posnat_not_eq_zero: "real_of_posnat n ~= 0"
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   547
apply (rule real_of_posnat_gt_zero [THEN real_not_refl2, THEN not_sym])
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   548
done
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   549
declare real_of_posnat_not_eq_zero [simp]
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   550
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   551
declare real_of_posnat_not_eq_zero [THEN real_mult_inv_left, simp]
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   552
declare real_of_posnat_not_eq_zero [THEN real_mult_inv_right, simp]
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   553
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   554
lemma real_of_posnat_ge_one: "1 <= real_of_posnat n"
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   555
apply (simp (no_asm) add: real_of_posnat_one [symmetric])
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   556
apply (induct_tac "n")
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   557
apply (auto simp add: real_of_posnat_Suc real_of_posnat_one order_less_imp_le)
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   558
done
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   559
declare real_of_posnat_ge_one [simp]
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   560
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   561
lemma real_of_posnat_real_inv_not_zero: "inverse(real_of_posnat n) ~= 0"
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   562
apply (rule real_inverse_not_zero) 
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   563
apply (rule real_of_posnat_gt_zero [THEN real_not_refl2, THEN not_sym])
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   564
done
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   565
declare real_of_posnat_real_inv_not_zero [simp]
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   566
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   567
lemma real_of_posnat_real_inv_inj: "inverse(real_of_posnat x) = inverse(real_of_posnat y) ==> x = y"
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   568
apply (rule inj_real_of_posnat [THEN injD])
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   569
apply (rule real_of_posnat_real_inv_not_zero
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   570
              [THEN real_mult_left_cancel, THEN iffD1, of x])
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   571
apply (simp add: real_mult_inv_left
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   572
             real_of_posnat_gt_zero [THEN real_not_refl2, THEN not_sym])
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   573
done
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   574
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   575
lemma real_mult_less_self: "0 < r ==> r*(1 + -inverse(real_of_posnat n)) < r"
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   576
apply (simp (no_asm) add: real_add_mult_distrib2)
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   577
apply (rule_tac C = "-r" in real_less_add_left_cancel)
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   578
apply (auto intro: real_mult_order simp add: real_add_assoc [symmetric] real_minus_zero_less_iff2)
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   579
done
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   580
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   581
lemma real_of_posnat_inv_Ex_iff: "(EX n. inverse(real_of_posnat n) < r) = (EX n. 1 < r * real_of_posnat n)"
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   582
apply safe
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   583
apply (drule_tac n1 = "n" in real_of_posnat_gt_zero [THEN real_mult_less_mono1])
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   584
apply (drule_tac [2] n2=n in real_of_posnat_gt_zero [THEN real_inverse_gt_0, THEN real_mult_less_mono1])
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   585
apply (auto simp add: real_of_posnat_gt_zero [THEN real_not_refl2, THEN not_sym] real_mult_assoc)
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   586
done
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   587
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   588
lemma real_of_posnat_inv_iff: "(inverse(real_of_posnat n) < r) = (1 < r * real_of_posnat n)"
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   589
apply safe
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   590
apply (drule_tac n1 = "n" in real_of_posnat_gt_zero [THEN real_mult_less_mono1])
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   591
apply (drule_tac [2] n2=n in real_of_posnat_gt_zero [THEN real_inverse_gt_0, THEN real_mult_less_mono1]) 
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   592
apply (auto simp add: real_mult_assoc)
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   593
done
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   594
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   595
lemma real_mult_le_le_mono1: "[| (0::real) <=z; x<=y |] ==> z*x<=z*y"
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   596
  by (rule Ring_and_Field.mult_left_mono)
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   597
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   598
lemma real_mult_le_le_mono2: "[| (0::real)<=z; x<=y |] ==> x*z<=y*z"
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   599
  by (rule Ring_and_Field.mult_right_mono)
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   600
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   601
lemma real_of_posnat_inv_le_iff: "(inverse(real_of_posnat n) <= r) = (1 <= r * real_of_posnat n)"
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   602
apply safe
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   603
apply (drule_tac n2=n in real_of_posnat_gt_zero [THEN order_less_imp_le, THEN real_mult_le_le_mono1])
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   604
apply (drule_tac [2] n3=n in real_of_posnat_gt_zero [THEN real_inverse_gt_0, THEN order_less_imp_le, THEN real_mult_le_le_mono1])
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   605
apply (auto simp add: real_mult_ac)
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   606
done
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   607
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   608
lemma real_of_posnat_less_iff: 
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   609
      "(real_of_posnat n < real_of_posnat m) = (n < m)"
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   610
apply (unfold real_of_posnat_def)
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   611
apply auto
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   612
done
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   613
declare real_of_posnat_less_iff [simp]
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   614
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   615
lemma real_of_posnat_le_iff: "(real_of_posnat n <= real_of_posnat m) = (n <= m)"
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   616
apply (auto dest: inj_real_of_posnat [THEN injD] simp add: real_le_less le_eq_less_or_eq)
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   617
done
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   618
declare real_of_posnat_le_iff [simp]
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   619
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   620
lemma real_mult_less_cancel3: "[| (0::real)<z; x*z<y*z |] ==> x<y"
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   621
apply auto
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   622
done
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   623
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   624
lemma real_mult_less_cancel4: "[| (0::real)<z; z*x<z*y |] ==> x<y"
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   625
apply auto
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   626
done
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   627
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   628
lemma real_of_posnat_less_inv_iff: "0 < u  ==> (u < inverse (real_of_posnat n)) = (real_of_posnat n < inverse(u))"
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   629
apply safe
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   630
apply (rule_tac n2=n in real_of_posnat_gt_zero [THEN real_inverse_gt_0, THEN real_mult_less_cancel3])
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   631
apply (rule_tac [2] x1 = "u" in real_inverse_gt_0 [THEN real_mult_less_cancel3])
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   632
apply (auto simp add: real_not_refl2 [THEN not_sym])
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   633
apply (rule_tac z = "u" in real_mult_less_cancel4)
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   634
apply (rule_tac [3] n1 = "n" in real_of_posnat_gt_zero [THEN real_mult_less_cancel4])
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   635
apply (auto simp add: real_not_refl2 [THEN not_sym] real_mult_assoc [symmetric])
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   636
done
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   637
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   638
lemma real_of_posnat_inv_eq_iff: "0 < u ==> (u = inverse(real_of_posnat n)) = (real_of_posnat n = inverse u)"
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   639
apply auto
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   640
done
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   641
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   642
lemma real_add_one_minus_inv_ge_zero: "0 <= 1 + -inverse(real_of_posnat n)"
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   643
apply (rule_tac C = "inverse (real_of_posnat n) " in real_le_add_right_cancel)
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   644
apply (simp (no_asm) add: real_add_assoc real_of_posnat_inv_le_iff)
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   645
done
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   646
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   647
lemma real_mult_add_one_minus_ge_zero: "0 < r ==> 0 <= r*(1 + -inverse(real_of_posnat n))"
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   648
apply (drule real_add_one_minus_inv_ge_zero [THEN real_mult_le_less_mono1])
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   649
apply auto
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   650
done
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   651
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   652
lemma real_inverse_unique: "x*y = (1::real) ==> y = inverse x"
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   653
apply (case_tac "x ~= 0")
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   654
apply (rule_tac c1 = "x" in real_mult_left_cancel [THEN iffD1])
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   655
apply auto
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   656
done
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   657
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   658
lemma real_inverse_gt_one: "[| (0::real) < x; x < 1 |] ==> 1 < inverse x"
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   659
apply (auto dest: real_inverse_less_swap)
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   660
done
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   661
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   662
lemma real_of_nat_gt_zero_cancel_iff: "(0 < real (n::nat)) = (0 < n)"
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   663
apply (rule real_of_nat_less_iff [THEN subst])
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   664
apply auto
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   665
done
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   666
declare real_of_nat_gt_zero_cancel_iff [simp]
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   667
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   668
lemma real_of_nat_le_zero_cancel_iff: "(real (n::nat) <= 0) = (n = 0)"
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   669
apply (rule real_of_nat_zero [THEN subst])
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   670
apply (subst real_of_nat_le_iff)
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   671
apply auto
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   672
done
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   673
declare real_of_nat_le_zero_cancel_iff [simp]
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   674
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   675
lemma not_real_of_nat_less_zero: "~ real (n::nat) < 0"
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   676
apply (simp (no_asm) add: real_le_def [symmetric] real_of_nat_ge_zero)
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   677
done
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   678
declare not_real_of_nat_less_zero [simp]
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   679
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   680
lemma real_of_nat_ge_zero_cancel_iff: 
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   681
      "(0 <= real (n::nat)) = (0 <= n)"
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   682
apply (unfold real_le_def le_def)
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   683
apply (simp (no_asm))
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   684
done
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   685
declare real_of_nat_ge_zero_cancel_iff [simp]
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   686
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   687
lemma real_of_nat_num_if: "real n = (if n=0 then 0 else 1 + real ((n::nat) - 1))"
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   688
apply (case_tac "n")
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   689
apply (auto simp add: real_of_nat_Suc)
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   690
done
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   691
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   692
(*RING AND FIELD*)
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   693
		lemma mult_less_imp_less_left:
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   694
		    "[|c*a < c*b; 0 < c|] ==> a < (b::'a::ordered_ring)"
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   695
		  by (force elim: order_less_asym simp add: mult_less_cancel_left)
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   696
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   697
		lemma mult_less_imp_less_right:
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   698
		    "[|a*c < b*c; 0 < c|] ==> a < (b::'a::ordered_ring)"
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   699
		  by (force elim: order_less_asym simp add: mult_less_cancel_right)
14265
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
   700
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
   701
ML
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
   702
{*
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
   703
val real_abs_def = thm "real_abs_def";
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
   704
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
   705
val real_minus_diff_eq = thm "real_minus_diff_eq";
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
   706
val real_gt_zero_preal_Ex = thm "real_gt_zero_preal_Ex";
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
   707
val real_gt_preal_preal_Ex = thm "real_gt_preal_preal_Ex";
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
   708
val real_ge_preal_preal_Ex = thm "real_ge_preal_preal_Ex";
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
   709
val real_less_all_preal = thm "real_less_all_preal";
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
   710
val real_less_all_real2 = thm "real_less_all_real2";
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
   711
val real_of_preal_le_iff = thm "real_of_preal_le_iff";
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
   712
val real_mult_order = thm "real_mult_order";
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
   713
val real_zero_less_one = thm "real_zero_less_one";
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
   714
val real_add_right_cancel_less = thm "real_add_right_cancel_less";
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
   715
val real_add_left_cancel_less = thm "real_add_left_cancel_less";
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
   716
val real_add_right_cancel_le = thm "real_add_right_cancel_le";
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
   717
val real_add_left_cancel_le = thm "real_add_left_cancel_le";
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
   718
val real_add_less_mono1 = thm "real_add_less_mono1";
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
   719
val real_add_le_mono1 = thm "real_add_le_mono1";
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
   720
val real_add_less_le_mono = thm "real_add_less_le_mono";
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
   721
val real_add_le_less_mono = thm "real_add_le_less_mono";
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
   722
val real_add_less_mono2 = thm "real_add_less_mono2";
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
   723
val real_less_add_right_cancel = thm "real_less_add_right_cancel";
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
   724
val real_less_add_left_cancel = thm "real_less_add_left_cancel";
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
   725
val real_le_add_right_cancel = thm "real_le_add_right_cancel";
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
   726
val real_le_add_left_cancel = thm "real_le_add_left_cancel";
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
   727
val real_add_order = thm "real_add_order";
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
   728
val real_le_add_order = thm "real_le_add_order";
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
   729
val real_add_less_mono = thm "real_add_less_mono";
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
   730
val real_add_left_le_mono1 = thm "real_add_left_le_mono1";
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
   731
val real_add_le_mono = thm "real_add_le_mono";
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
   732
val real_le_minus_iff = thm "real_le_minus_iff";
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
   733
val real_le_square = thm "real_le_square";
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
   734
val real_mult_less_mono1 = thm "real_mult_less_mono1";
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
   735
val real_mult_less_mono2 = thm "real_mult_less_mono2";
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
   736
val real_of_posnat_one = thm "real_of_posnat_one";
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
   737
val real_of_posnat_two = thm "real_of_posnat_two";
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
   738
val real_of_posnat_add = thm "real_of_posnat_add";
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
   739
val real_of_posnat_add_one = thm "real_of_posnat_add_one";
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
   740
val real_of_posnat_Suc = thm "real_of_posnat_Suc";
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
   741
val inj_real_of_posnat = thm "inj_real_of_posnat";
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
   742
val real_of_nat_zero = thm "real_of_nat_zero";
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
   743
val real_of_nat_one = thm "real_of_nat_one";
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
   744
val real_of_nat_add = thm "real_of_nat_add";
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
   745
val real_of_nat_Suc = thm "real_of_nat_Suc";
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
   746
val real_of_nat_less_iff = thm "real_of_nat_less_iff";
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
   747
val real_of_nat_le_iff = thm "real_of_nat_le_iff";
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
   748
val inj_real_of_nat = thm "inj_real_of_nat";
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
   749
val real_of_nat_ge_zero = thm "real_of_nat_ge_zero";
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
   750
val real_of_nat_mult = thm "real_of_nat_mult";
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
   751
val real_of_nat_inject = thm "real_of_nat_inject";
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
   752
val real_of_nat_diff = thm "real_of_nat_diff";
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
   753
val real_of_nat_zero_iff = thm "real_of_nat_zero_iff";
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
   754
val real_of_nat_neg_int = thm "real_of_nat_neg_int";
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
   755
val real_inverse_gt_0 = thm "real_inverse_gt_0";
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
   756
val real_inverse_less_0 = thm "real_inverse_less_0";
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
   757
val real_mult_less_cancel1 = thm "real_mult_less_cancel1";
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
   758
val real_mult_less_cancel2 = thm "real_mult_less_cancel2";
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
   759
val real_mult_less_iff1 = thm "real_mult_less_iff1";
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
   760
val real_mult_less_iff2 = thm "real_mult_less_iff2";
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
   761
val real_mult_le_cancel_iff1 = thm "real_mult_le_cancel_iff1";
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
   762
val real_mult_le_cancel_iff2 = thm "real_mult_le_cancel_iff2";
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
   763
val real_mult_less_mono = thm "real_mult_less_mono";
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
   764
val real_mult_less_mono' = thm "real_mult_less_mono'";
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
   765
val real_inverse_less_swap = thm "real_inverse_less_swap";
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
   766
val real_mult_is_0 = thm "real_mult_is_0";
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
   767
val real_inverse_add = thm "real_inverse_add";
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
   768
val real_sum_squares_cancel = thm "real_sum_squares_cancel";
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
   769
val real_sum_squares_cancel2 = thm "real_sum_squares_cancel2";
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
   770
val real_0_less_mult_iff = thm "real_0_less_mult_iff";
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
   771
val real_0_le_mult_iff = thm "real_0_le_mult_iff";
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
   772
val real_mult_less_0_iff = thm "real_mult_less_0_iff";
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
   773
val real_mult_le_0_iff = thm "real_mult_le_0_iff";
14268
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   774
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   775
val real_of_posnat_gt_zero = thm "real_of_posnat_gt_zero";
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   776
val real_inv_real_of_posnat_gt_zero = thm "real_inv_real_of_posnat_gt_zero";
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   777
val real_of_posnat_ge_zero = thm "real_of_posnat_ge_zero";
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   778
val real_of_posnat_not_eq_zero = thm "real_of_posnat_not_eq_zero";
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   779
val real_of_posnat_ge_one = thm "real_of_posnat_ge_one";
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   780
val real_of_posnat_real_inv_not_zero = thm "real_of_posnat_real_inv_not_zero";
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   781
val real_of_posnat_real_inv_inj = thm "real_of_posnat_real_inv_inj";
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   782
val real_mult_less_self = thm "real_mult_less_self";
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   783
val real_of_posnat_inv_Ex_iff = thm "real_of_posnat_inv_Ex_iff";
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   784
val real_of_posnat_inv_iff = thm "real_of_posnat_inv_iff";
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   785
val real_mult_le_le_mono1 = thm "real_mult_le_le_mono1";
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   786
val real_mult_le_le_mono2 = thm "real_mult_le_le_mono2";
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   787
val real_of_posnat_inv_le_iff = thm "real_of_posnat_inv_le_iff";
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   788
val real_of_posnat_less_iff = thm "real_of_posnat_less_iff";
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   789
val real_of_posnat_le_iff = thm "real_of_posnat_le_iff";
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   790
val real_mult_less_cancel3 = thm "real_mult_less_cancel3";
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   791
val real_mult_less_cancel4 = thm "real_mult_less_cancel4";
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   792
val real_of_posnat_less_inv_iff = thm "real_of_posnat_less_inv_iff";
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   793
val real_of_posnat_inv_eq_iff = thm "real_of_posnat_inv_eq_iff";
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   794
val real_add_one_minus_inv_ge_zero = thm "real_add_one_minus_inv_ge_zero";
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   795
val real_mult_add_one_minus_ge_zero = thm "real_mult_add_one_minus_ge_zero";
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   796
val real_inverse_unique = thm "real_inverse_unique";
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   797
val real_inverse_gt_one = thm "real_inverse_gt_one";
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   798
val real_of_nat_gt_zero_cancel_iff = thm "real_of_nat_gt_zero_cancel_iff";
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   799
val real_of_nat_le_zero_cancel_iff = thm "real_of_nat_le_zero_cancel_iff";
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   800
val not_real_of_nat_less_zero = thm "not_real_of_nat_less_zero";
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   801
val real_of_nat_ge_zero_cancel_iff = thm "real_of_nat_ge_zero_cancel_iff";
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14266
diff changeset
   802
val real_of_nat_num_if = thm "real_of_nat_num_if";
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 14268
diff changeset
   803
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 14268
diff changeset
   804
val INVERSE_ZERO = thm"INVERSE_ZERO";
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 14268
diff changeset
   805
val DIVISION_BY_ZERO = thm"DIVISION_BY_ZERO";
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 14268
diff changeset
   806
val real_mult_left_cancel = thm"real_mult_left_cancel";
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 14268
diff changeset
   807
val real_mult_right_cancel = thm"real_mult_right_cancel";
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 14268
diff changeset
   808
val real_mult_left_cancel_ccontr = thm"real_mult_left_cancel_ccontr";
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 14268
diff changeset
   809
val real_mult_right_cancel_ccontr = thm"real_mult_right_cancel_ccontr";
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 14268
diff changeset
   810
val real_inverse_not_zero = thm"real_inverse_not_zero";
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 14268
diff changeset
   811
val real_mult_not_zero = thm"real_mult_not_zero";
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 14268
diff changeset
   812
val real_inverse_inverse = thm"real_inverse_inverse";
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 14268
diff changeset
   813
val real_inverse_1 = thm"real_inverse_1";
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 14268
diff changeset
   814
val real_minus_inverse = thm"real_minus_inverse";
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 14268
diff changeset
   815
val real_inverse_distrib = thm"real_inverse_distrib";
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 14268
diff changeset
   816
val real_times_divide1_eq = thm"real_times_divide1_eq";
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 14268
diff changeset
   817
val real_times_divide2_eq = thm"real_times_divide2_eq";
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 14268
diff changeset
   818
val real_divide_divide1_eq = thm"real_divide_divide1_eq";
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 14268
diff changeset
   819
val real_divide_divide2_eq = thm"real_divide_divide2_eq";
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 14268
diff changeset
   820
val real_minus_divide_eq = thm"real_minus_divide_eq";
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 14268
diff changeset
   821
val real_divide_minus_eq = thm"real_divide_minus_eq";
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 14268
diff changeset
   822
val real_add_divide_distrib = thm"real_add_divide_distrib";
14265
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents: 9043
diff changeset
   823
*}
9043
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
   824
7334
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   825
end