src/HOL/Real/RealOrd.ML
author wenzelm
Thu, 27 Sep 2001 18:46:32 +0200
changeset 11599 12cc28aafb4d
parent 11464 ddea204de5bc
child 11701 3d51fbf81c17
permissions -rw-r--r--
renamed real_of_nat_eq_cancel to real_of_nat_inject, and declared as iff rule;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
7334
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
     1
(*  Title:       HOL/Real/Real.ML
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
     2
    ID:          $Id$
9013
9dd0274f76af Updated files to remove 0r and 1r from theorems in descendant theories
fleuriot
parents: 8867
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
    Description: Type "real" is a linear order
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
     6
*)
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
     7
9043
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
     8
(**** The simproc abel_cancel ****)
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
     9
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
    10
(*** Two lemmas needed for the simprocs ***)
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
    11
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
    12
(*Deletion of other terms in the formula, seeking the -x at the front of z*)
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
    13
Goal "((x::real) + (y + z) = y + u) = ((x + z) = u)";
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
    14
by (stac real_add_left_commute 1);
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
    15
by (rtac real_add_left_cancel 1);
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
    16
qed "real_add_cancel_21";
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
    17
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
    18
(*A further rule to deal with the case that
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
    19
  everything gets cancelled on the right.*)
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
    20
Goal "((x::real) + (y + z) = y) = (x = -z)";
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
    21
by (stac real_add_left_commute 1);
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
    22
by (res_inst_tac [("t", "y")] (real_add_zero_right RS subst) 1
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
    23
    THEN stac real_add_left_cancel 1);
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
    24
by (simp_tac (simpset() addsimps [real_eq_diff_eq RS sym]) 1);
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
    25
qed "real_add_cancel_end";
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
    26
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
    27
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
    28
structure Real_Cancel_Data =
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
    29
struct
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
    30
  val ss		= HOL_ss
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
    31
  val eq_reflection	= eq_reflection
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
    32
9434
d2fa043ab24f do nat pass theory value, but sg_ref;
wenzelm
parents: 9081
diff changeset
    33
  val sg_ref		= Sign.self_ref (Theory.sign_of (the_context ()))
9043
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
    34
  val T			= HOLogic.realT
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
    35
  val zero		= Const ("0", T)
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
    36
  val restrict_to_left  = restrict_to_left
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
    37
  val add_cancel_21	= real_add_cancel_21
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
    38
  val add_cancel_end	= real_add_cancel_end
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
    39
  val add_left_cancel	= real_add_left_cancel
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
    40
  val add_assoc		= real_add_assoc
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
    41
  val add_commute	= real_add_commute
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
    42
  val add_left_commute	= real_add_left_commute
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
    43
  val add_0		= real_add_zero_left
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
    44
  val add_0_right	= real_add_zero_right
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
    45
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
    46
  val eq_diff_eq	= real_eq_diff_eq
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
    47
  val eqI_rules		= [real_less_eqI, real_eq_eqI, real_le_eqI]
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
    48
  fun dest_eqI th = 
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
    49
      #1 (HOLogic.dest_bin "op =" HOLogic.boolT 
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
    50
	      (HOLogic.dest_Trueprop (concl_of th)))
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
    51
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
    52
  val diff_def		= real_diff_def
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
    53
  val minus_add_distrib	= real_minus_add_distrib
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
    54
  val minus_minus	= real_minus_minus
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
    55
  val minus_0		= real_minus_zero
10606
e3229a37d53f converted rinv to inverse;
bauerg
parents: 10043
diff changeset
    56
  val add_inverses	= [real_add_minus, real_add_minus_left]
9043
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
    57
  val cancel_simps	= [real_add_minus_cancel, real_minus_add_cancel]
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
    58
end;
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
    59
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
    60
structure Real_Cancel = Abel_Cancel (Real_Cancel_Data);
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
    61
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
    62
Addsimprocs [Real_Cancel.sum_conv, Real_Cancel.rel_conv];
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
    63
9081
d54b2c41fe0e new lemma real_minus_diff_eq
paulson
parents: 9069
diff changeset
    64
Goal "- (z - y) = y - (z::real)";
d54b2c41fe0e new lemma real_minus_diff_eq
paulson
parents: 9069
diff changeset
    65
by (Simp_tac 1);
d54b2c41fe0e new lemma real_minus_diff_eq
paulson
parents: 9069
diff changeset
    66
qed "real_minus_diff_eq";
d54b2c41fe0e new lemma real_minus_diff_eq
paulson
parents: 9069
diff changeset
    67
Addsimps [real_minus_diff_eq];
d54b2c41fe0e new lemma real_minus_diff_eq
paulson
parents: 9069
diff changeset
    68
9043
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
    69
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
    70
(**** Theorems about the ordering ****)
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
    71
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
    72
Goal "(0 < x) = (EX y. x = real_of_preal y)";
7334
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
    73
by (auto_tac (claset(), simpset() addsimps [real_of_preal_zero_less]));
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
    74
by (cut_inst_tac [("x","x")] real_of_preal_trichotomy 1);
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
    75
by (blast_tac (claset() addSEs [real_less_irrefl,
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
    76
				real_of_preal_not_minus_gt_zero RS notE]) 1);
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
    77
qed "real_gt_zero_preal_Ex";
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
    78
9043
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
    79
Goal "real_of_preal z < x ==> EX y. x = real_of_preal y";
7334
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
    80
by (blast_tac (claset() addSDs [real_of_preal_zero_less RS real_less_trans]
10752
c4f1bf2acf4c tidying, and separation of HOL-Hyperreal from HOL-Real
paulson
parents: 10712
diff changeset
    81
                        addIs [real_gt_zero_preal_Ex RS iffD1]) 1);
7334
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
    82
qed "real_gt_preal_preal_Ex";
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
    83
9043
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
    84
Goal "real_of_preal z <= x ==> EX y. x = real_of_preal y";
10752
c4f1bf2acf4c tidying, and separation of HOL-Hyperreal from HOL-Real
paulson
parents: 10712
diff changeset
    85
by (blast_tac (claset() addDs [order_le_imp_less_or_eq,
7334
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
    86
			       real_gt_preal_preal_Ex]) 1);
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
    87
qed "real_ge_preal_preal_Ex";
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
    88
9043
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
    89
Goal "y <= 0 ==> ALL x. y < real_of_preal x";
10752
c4f1bf2acf4c tidying, and separation of HOL-Hyperreal from HOL-Real
paulson
parents: 10712
diff changeset
    90
by (auto_tac (claset() addEs [order_le_imp_less_or_eq RS disjE]
7334
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
    91
                       addIs [real_of_preal_zero_less RSN(2,real_less_trans)],
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
    92
              simpset() addsimps [real_of_preal_zero_less]));
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
    93
qed "real_less_all_preal";
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
    94
9043
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
    95
Goal "~ 0 < y ==> ALL x. y < real_of_preal x";
7334
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
    96
by (blast_tac (claset() addSIs [real_less_all_preal,real_leI]) 1);
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
    97
qed "real_less_all_real2";
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
    98
9043
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
    99
Goal "[| R + L = S;  (0::real) < L |] ==> R < S";
7334
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   100
by (rtac (real_less_sum_gt_0_iff RS iffD1) 1);
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   101
by (auto_tac (claset(), simpset() addsimps real_add_ac));
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   102
qed "real_lemma_add_positive_imp_less";
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   103
9043
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
   104
Goal "EX T::real. 0 < T & R + T = S ==> R < S";
7334
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   105
by (blast_tac (claset() addIs [real_lemma_add_positive_imp_less]) 1);
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   106
qed "real_ex_add_positive_left_less";
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   107
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   108
(*Alternative definition for real_less.  NOT for rewriting*)
9043
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
   109
Goal "(R < S) = (EX T::real. 0 < T & R + T = S)";
7334
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   110
by (blast_tac (claset() addSIs [real_less_add_positive_left_Ex,
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   111
				real_ex_add_positive_left_less]) 1);
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   112
qed "real_less_iff_add";
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   113
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   114
Goal "(real_of_preal m1 <= real_of_preal m2) = (m1 <= m2)";
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   115
by (auto_tac (claset() addSIs [preal_leI],
10752
c4f1bf2acf4c tidying, and separation of HOL-Hyperreal from HOL-Real
paulson
parents: 10712
diff changeset
   116
              simpset() addsimps [real_less_le_iff RS sym]));
c4f1bf2acf4c tidying, and separation of HOL-Hyperreal from HOL-Real
paulson
parents: 10712
diff changeset
   117
by (dtac order_le_less_trans 1 THEN assume_tac 1);
7334
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   118
by (etac preal_less_irrefl 1);
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   119
qed "real_of_preal_le_iff";
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   120
9043
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
   121
Goal "[| 0 < x; 0 < y |] ==> (0::real) < x * y";
7334
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   122
by (auto_tac (claset(), simpset() addsimps [real_gt_zero_preal_Ex]));  
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   123
by (res_inst_tac [("x","y*ya")] exI 1);
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   124
by (full_simp_tac (simpset() addsimps [real_of_preal_mult]) 1);
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   125
qed "real_mult_order";
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   126
9043
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
   127
Goal "[| x < 0; y < 0 |] ==> (0::real) < x * y";
7334
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   128
by (REPEAT(dtac (real_minus_zero_less_iff RS iffD2) 1));
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   129
by (dtac real_mult_order 1 THEN assume_tac 1);
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   130
by (Asm_full_simp_tac 1);
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   131
qed "real_mult_less_zero1";
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   132
9043
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
   133
Goal "[| 0 < x; y < 0 |] ==> x*y < (0::real)";
7334
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   134
by (dtac (real_minus_zero_less_iff RS iffD2) 1);
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   135
by (dtac real_mult_order 1 THEN assume_tac 1);
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   136
by (rtac (real_minus_zero_less_iff RS iffD1) 1);
9043
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
   137
by (Asm_full_simp_tac 1);
7334
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   138
qed "real_mult_less_zero";
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   139
9043
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
   140
Goalw [real_one_def] "0 < 1r";
7334
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   141
by (auto_tac (claset() addIs [real_gt_zero_preal_Ex RS iffD2],
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   142
	      simpset() addsimps [real_of_preal_def]));
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   143
qed "real_zero_less_one";
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   144
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   145
(*** Monotonicity results ***)
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   146
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   147
Goal "(v+z < w+z) = (v < (w::real))";
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   148
by (Simp_tac 1);
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   149
qed "real_add_right_cancel_less";
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   150
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   151
Goal "(z+v < z+w) = (v < (w::real))";
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   152
by (Simp_tac 1);
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   153
qed "real_add_left_cancel_less";
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   154
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   155
Addsimps [real_add_right_cancel_less, real_add_left_cancel_less];
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   156
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   157
Goal "(v+z <= w+z) = (v <= (w::real))";
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   158
by (Simp_tac 1);
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   159
qed "real_add_right_cancel_le";
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   160
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   161
Goal "(z+v <= z+w) = (v <= (w::real))";
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   162
by (Simp_tac 1);
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   163
qed "real_add_left_cancel_le";
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   164
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   165
Addsimps [real_add_right_cancel_le, real_add_left_cancel_le];
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   166
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   167
(*"v<=w ==> v+z <= w+z"*)
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   168
bind_thm ("real_add_less_mono1", real_add_right_cancel_less RS iffD2);
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   169
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   170
(*"v<=w ==> v+z <= w+z"*)
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   171
bind_thm ("real_add_le_mono1", real_add_right_cancel_le RS iffD2);
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   172
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   173
Goal "!!z z'::real. [| w'<w; z'<=z |] ==> w' + z' < w + z";
10752
c4f1bf2acf4c tidying, and separation of HOL-Hyperreal from HOL-Real
paulson
parents: 10712
diff changeset
   174
by (etac (real_add_less_mono1 RS order_less_le_trans) 1);
7334
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   175
by (Simp_tac 1);
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   176
qed "real_add_less_le_mono";
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   177
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   178
Goal "!!z z'::real. [| w'<=w; z'<z |] ==> w' + z' < w + z";
10752
c4f1bf2acf4c tidying, and separation of HOL-Hyperreal from HOL-Real
paulson
parents: 10712
diff changeset
   179
by (etac (real_add_le_mono1 RS order_le_less_trans) 1);
7334
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   180
by (Simp_tac 1);
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   181
qed "real_add_le_less_mono";
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   182
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   183
Goal "!!(A::real). A < B ==> C + A < C + B";
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   184
by (Simp_tac 1);
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   185
qed "real_add_less_mono2";
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   186
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   187
Goal "!!(A::real). A + C < B + C ==> A < B";
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   188
by (Full_simp_tac 1);
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   189
qed "real_less_add_right_cancel";
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   190
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   191
Goal "!!(A::real). C + A < C + B ==> A < B";
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   192
by (Full_simp_tac 1);
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   193
qed "real_less_add_left_cancel";
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   194
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   195
Goal "!!(A::real). A + C <= B + C ==> A <= B";
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   196
by (Full_simp_tac 1);
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   197
qed "real_le_add_right_cancel";
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   198
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   199
Goal "!!(A::real). C + A <= C + B ==> A <= B";
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   200
by (Full_simp_tac 1);
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   201
qed "real_le_add_left_cancel";
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   202
9043
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
   203
Goal "[| 0 < x; 0 < y |] ==> (0::real) < x + y";
10752
c4f1bf2acf4c tidying, and separation of HOL-Hyperreal from HOL-Real
paulson
parents: 10712
diff changeset
   204
by (etac order_less_trans 1);
7334
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   205
by (dtac real_add_less_mono2 1);
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   206
by (Full_simp_tac 1);
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   207
qed "real_add_order";
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   208
9043
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
   209
Goal "[| 0 <= x; 0 <= y |] ==> (0::real) <= x + y";
10752
c4f1bf2acf4c tidying, and separation of HOL-Hyperreal from HOL-Real
paulson
parents: 10712
diff changeset
   210
by (REPEAT(dtac order_le_imp_less_or_eq 1));
c4f1bf2acf4c tidying, and separation of HOL-Hyperreal from HOL-Real
paulson
parents: 10712
diff changeset
   211
by (auto_tac (claset() addIs [real_add_order, order_less_imp_le],
7334
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   212
	      simpset()));
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   213
qed "real_le_add_order";
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   214
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   215
Goal "[| R1 < S1; R2 < S2 |] ==> R1 + R2 < S1 + (S2::real)";
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   216
by (dtac real_add_less_mono1 1);
10752
c4f1bf2acf4c tidying, and separation of HOL-Hyperreal from HOL-Real
paulson
parents: 10712
diff changeset
   217
by (etac order_less_trans 1);
7334
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   218
by (etac real_add_less_mono2 1);
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   219
qed "real_add_less_mono";
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   220
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   221
Goal "!!(q1::real). q1 <= q2  ==> x + q1 <= x + q2";
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   222
by (Simp_tac 1);
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   223
qed "real_add_left_le_mono1";
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   224
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   225
Goal "[|i<=j;  k<=l |] ==> i + k <= j + (l::real)";
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   226
by (dtac real_add_le_mono1 1);
10752
c4f1bf2acf4c tidying, and separation of HOL-Hyperreal from HOL-Real
paulson
parents: 10712
diff changeset
   227
by (etac order_trans 1);
7334
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   228
by (Simp_tac 1);
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   229
qed "real_add_le_mono";
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   230
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   231
Goal "EX (x::real). x < y";
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   232
by (rtac (real_add_zero_right RS subst) 1);
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   233
by (res_inst_tac [("x","y + (-1r)")] exI 1);
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   234
by (auto_tac (claset() addSIs [real_add_less_mono2],
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   235
	  simpset() addsimps [real_minus_zero_less_iff2, real_zero_less_one]));
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   236
qed "real_less_Ex";
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   237
9043
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
   238
Goal "(0::real) < r ==>  u + (-r) < u";
7334
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   239
by (res_inst_tac [("C","r")] real_less_add_right_cancel 1);
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   240
by (simp_tac (simpset() addsimps [real_add_assoc]) 1);
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   241
qed "real_add_minus_positive_less_self";
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   242
10699
f0c3da8477e9 more tidying
paulson
parents: 10677
diff changeset
   243
Goal "(-s <= -r) = ((r::real) <= s)";
f0c3da8477e9 more tidying
paulson
parents: 10677
diff changeset
   244
by (rtac sym 1);
7334
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   245
by (Step_tac 1);
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   246
by (dres_inst_tac [("x","-s")] real_add_left_le_mono1 1);
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   247
by (dres_inst_tac [("x","r")] real_add_left_le_mono1 2);
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   248
by Auto_tac;
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   249
by (dres_inst_tac [("z","-r")] real_add_le_mono1 1);
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   250
by (dres_inst_tac [("z","s")] real_add_le_mono1 2);
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   251
by (auto_tac (claset(), simpset() addsimps [real_add_assoc]));
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   252
qed "real_le_minus_iff";
10699
f0c3da8477e9 more tidying
paulson
parents: 10677
diff changeset
   253
Addsimps [real_le_minus_iff];
7334
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   254
9043
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
   255
Goal "(0::real) <= x*x";
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
   256
by (res_inst_tac [("R2.0","0"),("R1.0","x")] real_linear_less2 1);
7334
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   257
by (auto_tac (claset() addIs [real_mult_order,
10752
c4f1bf2acf4c tidying, and separation of HOL-Hyperreal from HOL-Real
paulson
parents: 10712
diff changeset
   258
			      real_mult_less_zero1,order_less_imp_le],
7334
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   259
	      simpset()));
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   260
qed "real_le_square";
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   261
Addsimps [real_le_square];
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   262
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   263
(*----------------------------------------------------------------------------
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   264
             An embedding of the naturals in the reals
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   265
 ----------------------------------------------------------------------------*)
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   266
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   267
Goalw [real_of_posnat_def] "real_of_posnat 0 = 1r";
10778
2c6605049646 more tidying, especially to remove real_of_posnat
paulson
parents: 10752
diff changeset
   268
by (simp_tac (simpset() addsimps [pnat_one_iff RS sym,real_of_preal_def,
2c6605049646 more tidying, especially to remove real_of_posnat
paulson
parents: 10752
diff changeset
   269
                       symmetric real_one_def]) 1);
7334
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   270
qed "real_of_posnat_one";
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   271
11464
ddea204de5bc turned translation for 1::nat into def.
nipkow
parents: 10919
diff changeset
   272
Goalw [real_of_posnat_def] "real_of_posnat 1' = 1r + 1r";
10778
2c6605049646 more tidying, especially to remove real_of_posnat
paulson
parents: 10752
diff changeset
   273
by (simp_tac (simpset() addsimps [real_of_preal_def,real_one_def,
2c6605049646 more tidying, especially to remove real_of_posnat
paulson
parents: 10752
diff changeset
   274
                               pnat_two_eq,real_add,prat_of_pnat_add RS sym,
2c6605049646 more tidying, especially to remove real_of_posnat
paulson
parents: 10752
diff changeset
   275
                               preal_of_prat_add RS sym] @ pnat_add_ac) 1);
7334
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   276
qed "real_of_posnat_two";
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   277
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   278
Goalw [real_of_posnat_def]
10778
2c6605049646 more tidying, especially to remove real_of_posnat
paulson
parents: 10752
diff changeset
   279
     "real_of_posnat n1 + real_of_posnat n2 = real_of_posnat (n1 + n2) + 1r";
7334
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   280
by (full_simp_tac (simpset() addsimps [real_of_posnat_one RS sym,
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   281
    real_of_posnat_def,real_of_preal_add RS sym,preal_of_prat_add RS sym,
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   282
    prat_of_pnat_add RS sym,pnat_of_nat_add]) 1);
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   283
qed "real_of_posnat_add";
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   284
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   285
Goal "real_of_posnat (n + 1) = real_of_posnat n + 1r";
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   286
by (res_inst_tac [("x1","1r")] (real_add_right_cancel RS iffD1) 1);
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   287
by (rtac (real_of_posnat_add RS subst) 1);
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   288
by (full_simp_tac (simpset() addsimps [real_of_posnat_two,real_add_assoc]) 1);
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   289
qed "real_of_posnat_add_one";
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   290
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   291
Goal "real_of_posnat (Suc n) = real_of_posnat n + 1r";
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   292
by (stac (real_of_posnat_add_one RS sym) 1);
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   293
by (Simp_tac 1);
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   294
qed "real_of_posnat_Suc";
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   295
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   296
Goal "inj(real_of_posnat)";
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   297
by (rtac injI 1);
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   298
by (rewtac real_of_posnat_def);
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   299
by (dtac (inj_real_of_preal RS injD) 1);
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   300
by (dtac (inj_preal_of_prat RS injD) 1);
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   301
by (dtac (inj_prat_of_pnat RS injD) 1);
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   302
by (etac (inj_pnat_of_nat RS injD) 1);
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   303
qed "inj_real_of_posnat";
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   304
10919
144ede948e58 renamings: real_of_nat, real_of_int -> (overloaded) real
paulson
parents: 10784
diff changeset
   305
Goalw [real_of_nat_def] "real (0::nat) = 0";
10778
2c6605049646 more tidying, especially to remove real_of_posnat
paulson
parents: 10752
diff changeset
   306
by (simp_tac (simpset() addsimps [real_of_posnat_one]) 1);
2c6605049646 more tidying, especially to remove real_of_posnat
paulson
parents: 10752
diff changeset
   307
qed "real_of_nat_zero";
2c6605049646 more tidying, especially to remove real_of_posnat
paulson
parents: 10752
diff changeset
   308
11464
ddea204de5bc turned translation for 1::nat into def.
nipkow
parents: 10919
diff changeset
   309
Goalw [real_of_nat_def] "real (1') = 1r";
10778
2c6605049646 more tidying, especially to remove real_of_posnat
paulson
parents: 10752
diff changeset
   310
by (simp_tac (simpset() addsimps [real_of_posnat_two, real_add_assoc]) 1);
2c6605049646 more tidying, especially to remove real_of_posnat
paulson
parents: 10752
diff changeset
   311
qed "real_of_nat_one";
2c6605049646 more tidying, especially to remove real_of_posnat
paulson
parents: 10752
diff changeset
   312
Addsimps [real_of_nat_zero, real_of_nat_one];
2c6605049646 more tidying, especially to remove real_of_posnat
paulson
parents: 10752
diff changeset
   313
2c6605049646 more tidying, especially to remove real_of_posnat
paulson
parents: 10752
diff changeset
   314
Goalw [real_of_nat_def]
10919
144ede948e58 renamings: real_of_nat, real_of_int -> (overloaded) real
paulson
parents: 10784
diff changeset
   315
     "real (m + n) = real (m::nat) + real n";
10778
2c6605049646 more tidying, especially to remove real_of_posnat
paulson
parents: 10752
diff changeset
   316
by (simp_tac (simpset() addsimps 
2c6605049646 more tidying, especially to remove real_of_posnat
paulson
parents: 10752
diff changeset
   317
              [real_of_posnat_add,real_add_assoc RS sym]) 1);
2c6605049646 more tidying, especially to remove real_of_posnat
paulson
parents: 10752
diff changeset
   318
qed "real_of_nat_add";
10784
27e4d90b35b5 more removal of obsolete rules
paulson
parents: 10778
diff changeset
   319
Addsimps [real_of_nat_add];
7334
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   320
10778
2c6605049646 more tidying, especially to remove real_of_posnat
paulson
parents: 10752
diff changeset
   321
(*Not for addsimps: often the LHS is used to represent a positive natural*)
10919
144ede948e58 renamings: real_of_nat, real_of_int -> (overloaded) real
paulson
parents: 10784
diff changeset
   322
Goalw [real_of_nat_def] "real (Suc n) = real n + 1r";
10778
2c6605049646 more tidying, especially to remove real_of_posnat
paulson
parents: 10752
diff changeset
   323
by (simp_tac (simpset() addsimps [real_of_posnat_Suc] @ real_add_ac) 1);
2c6605049646 more tidying, especially to remove real_of_posnat
paulson
parents: 10752
diff changeset
   324
qed "real_of_nat_Suc";
2c6605049646 more tidying, especially to remove real_of_posnat
paulson
parents: 10752
diff changeset
   325
2c6605049646 more tidying, especially to remove real_of_posnat
paulson
parents: 10752
diff changeset
   326
Goalw [real_of_nat_def, real_of_posnat_def]
10919
144ede948e58 renamings: real_of_nat, real_of_int -> (overloaded) real
paulson
parents: 10784
diff changeset
   327
     "(real (n::nat) < real m) = (n < m)";
10778
2c6605049646 more tidying, especially to remove real_of_posnat
paulson
parents: 10752
diff changeset
   328
by Auto_tac;
2c6605049646 more tidying, especially to remove real_of_posnat
paulson
parents: 10752
diff changeset
   329
qed "real_of_nat_less_iff";
2c6605049646 more tidying, especially to remove real_of_posnat
paulson
parents: 10752
diff changeset
   330
AddIffs [real_of_nat_less_iff];
7334
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   331
10919
144ede948e58 renamings: real_of_nat, real_of_int -> (overloaded) real
paulson
parents: 10784
diff changeset
   332
Goal "(real (n::nat) <= real m) = (n <= m)";
10778
2c6605049646 more tidying, especially to remove real_of_posnat
paulson
parents: 10752
diff changeset
   333
by (simp_tac (simpset() addsimps [linorder_not_less RS sym]) 1); 
2c6605049646 more tidying, especially to remove real_of_posnat
paulson
parents: 10752
diff changeset
   334
qed "real_of_nat_le_iff";
2c6605049646 more tidying, especially to remove real_of_posnat
paulson
parents: 10752
diff changeset
   335
AddIffs [real_of_nat_le_iff];
2c6605049646 more tidying, especially to remove real_of_posnat
paulson
parents: 10752
diff changeset
   336
10919
144ede948e58 renamings: real_of_nat, real_of_int -> (overloaded) real
paulson
parents: 10784
diff changeset
   337
Goal "inj (real :: nat => real)";
10778
2c6605049646 more tidying, especially to remove real_of_posnat
paulson
parents: 10752
diff changeset
   338
by (rtac injI 1);
2c6605049646 more tidying, especially to remove real_of_posnat
paulson
parents: 10752
diff changeset
   339
by (auto_tac (claset() addSIs [inj_real_of_posnat RS injD],
2c6605049646 more tidying, especially to remove real_of_posnat
paulson
parents: 10752
diff changeset
   340
	      simpset() addsimps [real_of_nat_def,real_add_right_cancel]));
2c6605049646 more tidying, especially to remove real_of_posnat
paulson
parents: 10752
diff changeset
   341
qed "inj_real_of_nat";
2c6605049646 more tidying, especially to remove real_of_posnat
paulson
parents: 10752
diff changeset
   342
10919
144ede948e58 renamings: real_of_nat, real_of_int -> (overloaded) real
paulson
parents: 10784
diff changeset
   343
Goal "0 <= real (n::nat)";
7334
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   344
by (induct_tac "n" 1);
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   345
by (auto_tac (claset(),
10778
2c6605049646 more tidying, especially to remove real_of_posnat
paulson
parents: 10752
diff changeset
   346
              simpset () addsimps [real_of_nat_Suc]));
2c6605049646 more tidying, especially to remove real_of_posnat
paulson
parents: 10752
diff changeset
   347
by (dtac real_add_le_less_mono 1); 
2c6605049646 more tidying, especially to remove real_of_posnat
paulson
parents: 10752
diff changeset
   348
by (rtac real_zero_less_one 1); 
2c6605049646 more tidying, especially to remove real_of_posnat
paulson
parents: 10752
diff changeset
   349
by (asm_full_simp_tac (simpset() addsimps [order_less_imp_le]) 1); 
2c6605049646 more tidying, especially to remove real_of_posnat
paulson
parents: 10752
diff changeset
   350
qed "real_of_nat_ge_zero";
2c6605049646 more tidying, especially to remove real_of_posnat
paulson
parents: 10752
diff changeset
   351
AddIffs [real_of_nat_ge_zero];
2c6605049646 more tidying, especially to remove real_of_posnat
paulson
parents: 10752
diff changeset
   352
10919
144ede948e58 renamings: real_of_nat, real_of_int -> (overloaded) real
paulson
parents: 10784
diff changeset
   353
Goal "real (m * n) = real (m::nat) * real n";
10778
2c6605049646 more tidying, especially to remove real_of_posnat
paulson
parents: 10752
diff changeset
   354
by (induct_tac "m" 1);
2c6605049646 more tidying, especially to remove real_of_posnat
paulson
parents: 10752
diff changeset
   355
by (auto_tac (claset(),
10784
27e4d90b35b5 more removal of obsolete rules
paulson
parents: 10778
diff changeset
   356
	      simpset() addsimps [real_of_nat_Suc,
10778
2c6605049646 more tidying, especially to remove real_of_posnat
paulson
parents: 10752
diff changeset
   357
				  real_add_mult_distrib, real_add_commute]));
2c6605049646 more tidying, especially to remove real_of_posnat
paulson
parents: 10752
diff changeset
   358
qed "real_of_nat_mult";
10784
27e4d90b35b5 more removal of obsolete rules
paulson
parents: 10778
diff changeset
   359
Addsimps [real_of_nat_mult];
10778
2c6605049646 more tidying, especially to remove real_of_posnat
paulson
parents: 10752
diff changeset
   360
10919
144ede948e58 renamings: real_of_nat, real_of_int -> (overloaded) real
paulson
parents: 10784
diff changeset
   361
Goal "(real (n::nat) = real m) = (n = m)";
10778
2c6605049646 more tidying, especially to remove real_of_posnat
paulson
parents: 10752
diff changeset
   362
by (auto_tac (claset() addDs [inj_real_of_nat RS injD], simpset()));
11599
12cc28aafb4d renamed real_of_nat_eq_cancel to real_of_nat_inject, and declared as iff rule;
wenzelm
parents: 11464
diff changeset
   363
qed "real_of_nat_inject";
12cc28aafb4d renamed real_of_nat_eq_cancel to real_of_nat_inject, and declared as iff rule;
wenzelm
parents: 11464
diff changeset
   364
AddIffs [real_of_nat_inject];
7334
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   365
10919
144ede948e58 renamings: real_of_nat, real_of_int -> (overloaded) real
paulson
parents: 10784
diff changeset
   366
Goal "n <= m --> real (m - n) = real (m::nat) - real n";
10778
2c6605049646 more tidying, especially to remove real_of_posnat
paulson
parents: 10752
diff changeset
   367
by (induct_tac "m" 1);
2c6605049646 more tidying, especially to remove real_of_posnat
paulson
parents: 10752
diff changeset
   368
by (auto_tac (claset(),
10919
144ede948e58 renamings: real_of_nat, real_of_int -> (overloaded) real
paulson
parents: 10784
diff changeset
   369
	      simpset() addsimps [real_diff_def, Suc_diff_le, le_Suc_eq, 
144ede948e58 renamings: real_of_nat, real_of_int -> (overloaded) real
paulson
parents: 10784
diff changeset
   370
                          real_of_nat_Suc, real_of_nat_zero] @ real_add_ac));
144ede948e58 renamings: real_of_nat, real_of_int -> (overloaded) real
paulson
parents: 10784
diff changeset
   371
qed_spec_mp "real_of_nat_diff";
10778
2c6605049646 more tidying, especially to remove real_of_posnat
paulson
parents: 10752
diff changeset
   372
10919
144ede948e58 renamings: real_of_nat, real_of_int -> (overloaded) real
paulson
parents: 10784
diff changeset
   373
Goal "(real (n::nat) = 0) = (n = 0)";
11599
12cc28aafb4d renamed real_of_nat_eq_cancel to real_of_nat_inject, and declared as iff rule;
wenzelm
parents: 11464
diff changeset
   374
by (auto_tac ((claset() addIs [inj_real_of_nat RS injD], simpset()) delIffs [real_of_nat_inject]));
10778
2c6605049646 more tidying, especially to remove real_of_posnat
paulson
parents: 10752
diff changeset
   375
qed "real_of_nat_zero_iff";
2c6605049646 more tidying, especially to remove real_of_posnat
paulson
parents: 10752
diff changeset
   376
10919
144ede948e58 renamings: real_of_nat, real_of_int -> (overloaded) real
paulson
parents: 10784
diff changeset
   377
Goal "neg z ==> real (nat z) = 0";
10778
2c6605049646 more tidying, especially to remove real_of_posnat
paulson
parents: 10752
diff changeset
   378
by (asm_simp_tac (simpset() addsimps [neg_nat, real_of_nat_zero]) 1);
2c6605049646 more tidying, especially to remove real_of_posnat
paulson
parents: 10752
diff changeset
   379
qed "real_of_nat_neg_int";
2c6605049646 more tidying, especially to remove real_of_posnat
paulson
parents: 10752
diff changeset
   380
Addsimps [real_of_nat_neg_int];
2c6605049646 more tidying, especially to remove real_of_posnat
paulson
parents: 10752
diff changeset
   381
2c6605049646 more tidying, especially to remove real_of_posnat
paulson
parents: 10752
diff changeset
   382
2c6605049646 more tidying, especially to remove real_of_posnat
paulson
parents: 10752
diff changeset
   383
(*----------------------------------------------------------------------------
2c6605049646 more tidying, especially to remove real_of_posnat
paulson
parents: 10752
diff changeset
   384
     inverse, etc.
2c6605049646 more tidying, especially to remove real_of_posnat
paulson
parents: 10752
diff changeset
   385
 ----------------------------------------------------------------------------*)
7334
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   386
10606
e3229a37d53f converted rinv to inverse;
bauerg
parents: 10043
diff changeset
   387
Goal "0 < x ==> 0 < inverse (x::real)";
7334
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   388
by (EVERY1[rtac ccontr, dtac real_leI]);
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   389
by (forward_tac [real_minus_zero_less_iff2 RS iffD2] 1);
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   390
by (forward_tac [real_not_refl2 RS not_sym] 1);
10606
e3229a37d53f converted rinv to inverse;
bauerg
parents: 10043
diff changeset
   391
by (dtac (real_not_refl2 RS not_sym RS real_inverse_not_zero) 1);
10752
c4f1bf2acf4c tidying, and separation of HOL-Hyperreal from HOL-Real
paulson
parents: 10712
diff changeset
   392
by (EVERY1[dtac order_le_imp_less_or_eq, Step_tac]); 
7334
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   393
by (dtac real_mult_less_zero1 1 THEN assume_tac 1);
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   394
by (auto_tac (claset() addIs [real_zero_less_one RS real_less_asym],
9053
paulson
parents: 9043
diff changeset
   395
	      simpset()));
10606
e3229a37d53f converted rinv to inverse;
bauerg
parents: 10043
diff changeset
   396
qed "real_inverse_gt_zero";
7334
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   397
10606
e3229a37d53f converted rinv to inverse;
bauerg
parents: 10043
diff changeset
   398
Goal "x < 0 ==> inverse (x::real) < 0";
7499
23e090051cb8 isatool expandshort;
wenzelm
parents: 7334
diff changeset
   399
by (ftac real_not_refl2 1);
7334
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   400
by (dtac (real_minus_zero_less_iff RS iffD2) 1);
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   401
by (rtac (real_minus_zero_less_iff RS iffD1) 1);
10648
a8c647cfa31f first stage in tidying up Real and Hyperreal.
paulson
parents: 10606
diff changeset
   402
by (stac (real_minus_inverse RS sym) 1);
10606
e3229a37d53f converted rinv to inverse;
bauerg
parents: 10043
diff changeset
   403
by (auto_tac (claset() addIs [real_inverse_gt_zero], simpset()));
e3229a37d53f converted rinv to inverse;
bauerg
parents: 10043
diff changeset
   404
qed "real_inverse_less_zero";
7334
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   405
9043
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
   406
Goal "[| (0::real) < z; x < y |] ==> x*z < y*z";       
7334
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   407
by (rotate_tac 1 1);
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   408
by (dtac real_less_sum_gt_zero 1);
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   409
by (rtac real_sum_gt_zero_less 1);
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   410
by (dtac real_mult_order 1 THEN assume_tac 1);
10752
c4f1bf2acf4c tidying, and separation of HOL-Hyperreal from HOL-Real
paulson
parents: 10712
diff changeset
   411
by (asm_full_simp_tac
c4f1bf2acf4c tidying, and separation of HOL-Hyperreal from HOL-Real
paulson
parents: 10712
diff changeset
   412
    (simpset() addsimps [real_add_mult_distrib2, real_mult_commute ]) 1);
7334
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   413
qed "real_mult_less_mono1";
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   414
10606
e3229a37d53f converted rinv to inverse;
bauerg
parents: 10043
diff changeset
   415
Goal "[| (0::real) < z; x < y |] ==> z * x < z * y";       
10752
c4f1bf2acf4c tidying, and separation of HOL-Hyperreal from HOL-Real
paulson
parents: 10712
diff changeset
   416
by (asm_simp_tac
c4f1bf2acf4c tidying, and separation of HOL-Hyperreal from HOL-Real
paulson
parents: 10712
diff changeset
   417
    (simpset() addsimps [real_mult_commute,real_mult_less_mono1]) 1);
7334
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   418
qed "real_mult_less_mono2";
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   419
10606
e3229a37d53f converted rinv to inverse;
bauerg
parents: 10043
diff changeset
   420
Goal "[| (0::real) < z; x * z < y * z |] ==> x < y";
10752
c4f1bf2acf4c tidying, and separation of HOL-Hyperreal from HOL-Real
paulson
parents: 10712
diff changeset
   421
by (forw_inst_tac [("x","x*z")] 
c4f1bf2acf4c tidying, and separation of HOL-Hyperreal from HOL-Real
paulson
parents: 10712
diff changeset
   422
    (real_inverse_gt_zero RS real_mult_less_mono1) 1);
7334
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   423
by (auto_tac (claset(),
10752
c4f1bf2acf4c tidying, and separation of HOL-Hyperreal from HOL-Real
paulson
parents: 10712
diff changeset
   424
	      simpset() addsimps [real_mult_assoc,real_not_refl2 RS not_sym]));
7334
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   425
qed "real_mult_less_cancel1";
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   426
9043
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
   427
Goal "[| (0::real) < z; z*x < z*y |] ==> x < y";
7334
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   428
by (etac real_mult_less_cancel1 1);
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   429
by (asm_full_simp_tac (simpset() addsimps [real_mult_commute]) 1);
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   430
qed "real_mult_less_cancel2";
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   431
9043
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
   432
Goal "(0::real) < z ==> (x*z < y*z) = (x < y)";
10752
c4f1bf2acf4c tidying, and separation of HOL-Hyperreal from HOL-Real
paulson
parents: 10712
diff changeset
   433
by (blast_tac
c4f1bf2acf4c tidying, and separation of HOL-Hyperreal from HOL-Real
paulson
parents: 10712
diff changeset
   434
    (claset() addIs [real_mult_less_mono1, real_mult_less_cancel1]) 1);
7334
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   435
qed "real_mult_less_iff1";
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   436
9043
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
   437
Goal "(0::real) < z ==> (z*x < z*y) = (x < y)";
10752
c4f1bf2acf4c tidying, and separation of HOL-Hyperreal from HOL-Real
paulson
parents: 10712
diff changeset
   438
by (blast_tac
c4f1bf2acf4c tidying, and separation of HOL-Hyperreal from HOL-Real
paulson
parents: 10712
diff changeset
   439
    (claset() addIs [real_mult_less_mono2, real_mult_less_cancel2]) 1);
7334
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   440
qed "real_mult_less_iff2";
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   441
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   442
Addsimps [real_mult_less_iff1,real_mult_less_iff2];
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   443
9013
9dd0274f76af Updated files to remove 0r and 1r from theorems in descendant theories
fleuriot
parents: 8867
diff changeset
   444
(* 05/00 *)
9043
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
   445
Goalw [real_le_def] "(0::real) < z ==> (x*z <= y*z) = (x <= y)";
9013
9dd0274f76af Updated files to remove 0r and 1r from theorems in descendant theories
fleuriot
parents: 8867
diff changeset
   446
by (Auto_tac);
9dd0274f76af Updated files to remove 0r and 1r from theorems in descendant theories
fleuriot
parents: 8867
diff changeset
   447
qed "real_mult_le_cancel_iff1";
9dd0274f76af Updated files to remove 0r and 1r from theorems in descendant theories
fleuriot
parents: 8867
diff changeset
   448
9043
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
   449
Goalw [real_le_def] "(0::real) < z ==> (z*x <= z*y) = (x <= y)";
9013
9dd0274f76af Updated files to remove 0r and 1r from theorems in descendant theories
fleuriot
parents: 8867
diff changeset
   450
by (Auto_tac);
9dd0274f76af Updated files to remove 0r and 1r from theorems in descendant theories
fleuriot
parents: 8867
diff changeset
   451
qed "real_mult_le_cancel_iff2";
9dd0274f76af Updated files to remove 0r and 1r from theorems in descendant theories
fleuriot
parents: 8867
diff changeset
   452
9dd0274f76af Updated files to remove 0r and 1r from theorems in descendant theories
fleuriot
parents: 8867
diff changeset
   453
Addsimps [real_mult_le_cancel_iff1,real_mult_le_cancel_iff2];
9dd0274f76af Updated files to remove 0r and 1r from theorems in descendant theories
fleuriot
parents: 8867
diff changeset
   454
9dd0274f76af Updated files to remove 0r and 1r from theorems in descendant theories
fleuriot
parents: 8867
diff changeset
   455
9043
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
   456
Goal "[| (0::real) <= z; x < y |] ==> x*z <= y*z";
10752
c4f1bf2acf4c tidying, and separation of HOL-Hyperreal from HOL-Real
paulson
parents: 10712
diff changeset
   457
by (EVERY1 [rtac real_less_or_eq_imp_le, dtac order_le_imp_less_or_eq]);
7334
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   458
by (auto_tac (claset() addIs [real_mult_less_mono1],simpset()));
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   459
qed "real_mult_le_less_mono1";
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   460
10752
c4f1bf2acf4c tidying, and separation of HOL-Hyperreal from HOL-Real
paulson
parents: 10712
diff changeset
   461
Goal "[| u<v;  x<y;  (0::real) < v;  0 < x |] ==> u*x < v* y";
c4f1bf2acf4c tidying, and separation of HOL-Hyperreal from HOL-Real
paulson
parents: 10712
diff changeset
   462
by (etac (real_mult_less_mono1 RS order_less_trans) 1);
c4f1bf2acf4c tidying, and separation of HOL-Hyperreal from HOL-Real
paulson
parents: 10712
diff changeset
   463
by (assume_tac 1);
c4f1bf2acf4c tidying, and separation of HOL-Hyperreal from HOL-Real
paulson
parents: 10712
diff changeset
   464
by (etac real_mult_less_mono2 1);
c4f1bf2acf4c tidying, and separation of HOL-Hyperreal from HOL-Real
paulson
parents: 10712
diff changeset
   465
by (assume_tac 1);
7334
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   466
qed "real_mult_less_mono";
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   467
10784
27e4d90b35b5 more removal of obsolete rules
paulson
parents: 10778
diff changeset
   468
(*Variant of the theorem above; sometimes it's stronger*)
10752
c4f1bf2acf4c tidying, and separation of HOL-Hyperreal from HOL-Real
paulson
parents: 10712
diff changeset
   469
Goal "[| x < y;  r1 < r2;  (0::real) <= r1;  0 <= x|] ==> r1 * x < r2 * y";
c4f1bf2acf4c tidying, and separation of HOL-Hyperreal from HOL-Real
paulson
parents: 10712
diff changeset
   470
by (subgoal_tac "0<r2" 1);
c4f1bf2acf4c tidying, and separation of HOL-Hyperreal from HOL-Real
paulson
parents: 10712
diff changeset
   471
by (blast_tac (claset() addIs [order_le_less_trans]) 2); 
c4f1bf2acf4c tidying, and separation of HOL-Hyperreal from HOL-Real
paulson
parents: 10712
diff changeset
   472
by (case_tac "x=0" 1);
c4f1bf2acf4c tidying, and separation of HOL-Hyperreal from HOL-Real
paulson
parents: 10712
diff changeset
   473
by (auto_tac (claset() addSDs [order_le_imp_less_or_eq] 
c4f1bf2acf4c tidying, and separation of HOL-Hyperreal from HOL-Real
paulson
parents: 10712
diff changeset
   474
	               addIs [real_mult_less_mono, real_mult_order],
c4f1bf2acf4c tidying, and separation of HOL-Hyperreal from HOL-Real
paulson
parents: 10712
diff changeset
   475
	      simpset()));
c4f1bf2acf4c tidying, and separation of HOL-Hyperreal from HOL-Real
paulson
parents: 10712
diff changeset
   476
qed "real_mult_less_mono'";
7334
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   477
9043
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
   478
Goal "1r <= x ==> 0 < x";
7334
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   479
by (rtac ccontr 1 THEN dtac real_leI 1);
10752
c4f1bf2acf4c tidying, and separation of HOL-Hyperreal from HOL-Real
paulson
parents: 10712
diff changeset
   480
by (dtac order_trans 1 THEN assume_tac 1);
c4f1bf2acf4c tidying, and separation of HOL-Hyperreal from HOL-Real
paulson
parents: 10712
diff changeset
   481
by (auto_tac (claset() addDs [real_zero_less_one RSN (2,order_le_less_trans)],
c4f1bf2acf4c tidying, and separation of HOL-Hyperreal from HOL-Real
paulson
parents: 10712
diff changeset
   482
	      simpset()));
7334
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   483
qed "real_gt_zero";
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   484
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   485
Goal "[| 1r < r; 1r <= x |]  ==> x <= r * x";
10752
c4f1bf2acf4c tidying, and separation of HOL-Hyperreal from HOL-Real
paulson
parents: 10712
diff changeset
   486
by (dtac (real_gt_zero RS order_less_imp_le) 1);
7334
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   487
by (auto_tac (claset() addSDs [real_mult_le_less_mono1],
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   488
    simpset()));
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   489
qed "real_mult_self_le";
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   490
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   491
Goal "[| 1r <= r; 1r <= x |]  ==> x <= r * x";
10752
c4f1bf2acf4c tidying, and separation of HOL-Hyperreal from HOL-Real
paulson
parents: 10712
diff changeset
   492
by (dtac order_le_imp_less_or_eq 1);
c4f1bf2acf4c tidying, and separation of HOL-Hyperreal from HOL-Real
paulson
parents: 10712
diff changeset
   493
by (auto_tac (claset() addIs [real_mult_self_le], simpset()));
7334
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   494
qed "real_mult_self_le2";
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   495
10606
e3229a37d53f converted rinv to inverse;
bauerg
parents: 10043
diff changeset
   496
Goal "[| 0 < r; r < x |] ==> inverse x < inverse (r::real)";
10752
c4f1bf2acf4c tidying, and separation of HOL-Hyperreal from HOL-Real
paulson
parents: 10712
diff changeset
   497
by (ftac order_less_trans 1 THEN assume_tac 1);
10606
e3229a37d53f converted rinv to inverse;
bauerg
parents: 10043
diff changeset
   498
by (ftac real_inverse_gt_zero 1);
e3229a37d53f converted rinv to inverse;
bauerg
parents: 10043
diff changeset
   499
by (forw_inst_tac [("x","x")] real_inverse_gt_zero 1);
e3229a37d53f converted rinv to inverse;
bauerg
parents: 10043
diff changeset
   500
by (forw_inst_tac [("x","r"),("z","inverse r")] real_mult_less_mono1 1);
7334
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   501
by (assume_tac 1);
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   502
by (asm_full_simp_tac (simpset() addsimps [real_not_refl2 RS 
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   503
					   not_sym RS real_mult_inv_right]) 1);
10606
e3229a37d53f converted rinv to inverse;
bauerg
parents: 10043
diff changeset
   504
by (ftac real_inverse_gt_zero 1);
e3229a37d53f converted rinv to inverse;
bauerg
parents: 10043
diff changeset
   505
by (forw_inst_tac [("x","1r"),("z","inverse x")] real_mult_less_mono2 1);
7334
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   506
by (assume_tac 1);
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   507
by (asm_full_simp_tac (simpset() addsimps [real_not_refl2 RS 
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   508
         not_sym RS real_mult_inv_left,real_mult_assoc RS sym]) 1);
10606
e3229a37d53f converted rinv to inverse;
bauerg
parents: 10043
diff changeset
   509
qed "real_inverse_less_swap";
7334
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   510
9043
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
   511
Goal "(x*y = 0) = (x = 0 | y = (0::real))";
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
   512
by Auto_tac;
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
   513
by (blast_tac (claset() addIs [ccontr] addDs [real_mult_not_zero]) 1);
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
   514
qed "real_mult_is_0";
10712
351ba950d4d9 further tidying
paulson
parents: 10699
diff changeset
   515
AddIffs [real_mult_is_0];
7334
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   516
10778
2c6605049646 more tidying, especially to remove real_of_posnat
paulson
parents: 10752
diff changeset
   517
Goal "[| x ~= 0; y ~= 0 |] \
2c6605049646 more tidying, especially to remove real_of_posnat
paulson
parents: 10752
diff changeset
   518
\     ==> inverse x + inverse y = (x + y) * inverse (x * (y::real))";
7334
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   519
by (asm_full_simp_tac (simpset() addsimps 
10606
e3229a37d53f converted rinv to inverse;
bauerg
parents: 10043
diff changeset
   520
		       [real_inverse_distrib,real_add_mult_distrib,
7334
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   521
			real_mult_assoc RS sym]) 1);
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   522
by (stac real_mult_assoc 1);
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   523
by (rtac (real_mult_left_commute RS subst) 1);
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   524
by (asm_full_simp_tac (simpset() addsimps [real_add_commute]) 1);
10606
e3229a37d53f converted rinv to inverse;
bauerg
parents: 10043
diff changeset
   525
qed "real_inverse_add";
7334
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   526
9013
9dd0274f76af Updated files to remove 0r and 1r from theorems in descendant theories
fleuriot
parents: 8867
diff changeset
   527
(* 05/00 *)
9043
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
   528
Goal "(0 <= -R) = (R <= (0::real))";
9013
9dd0274f76af Updated files to remove 0r and 1r from theorems in descendant theories
fleuriot
parents: 8867
diff changeset
   529
by (auto_tac (claset() addDs [sym],
9dd0274f76af Updated files to remove 0r and 1r from theorems in descendant theories
fleuriot
parents: 8867
diff changeset
   530
    simpset() addsimps [real_le_less]));
9dd0274f76af Updated files to remove 0r and 1r from theorems in descendant theories
fleuriot
parents: 8867
diff changeset
   531
qed "real_minus_zero_le_iff";
9dd0274f76af Updated files to remove 0r and 1r from theorems in descendant theories
fleuriot
parents: 8867
diff changeset
   532
9043
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
   533
Goal "(-R <= 0) = ((0::real) <= R)";
9013
9dd0274f76af Updated files to remove 0r and 1r from theorems in descendant theories
fleuriot
parents: 8867
diff changeset
   534
by (auto_tac (claset(),simpset() addsimps 
9dd0274f76af Updated files to remove 0r and 1r from theorems in descendant theories
fleuriot
parents: 8867
diff changeset
   535
    [real_minus_zero_less_iff2,real_le_less]));
9dd0274f76af Updated files to remove 0r and 1r from theorems in descendant theories
fleuriot
parents: 8867
diff changeset
   536
qed "real_minus_zero_le_iff2";
9dd0274f76af Updated files to remove 0r and 1r from theorems in descendant theories
fleuriot
parents: 8867
diff changeset
   537
9053
paulson
parents: 9043
diff changeset
   538
Addsimps [real_minus_zero_le_iff, real_minus_zero_le_iff2];
9013
9dd0274f76af Updated files to remove 0r and 1r from theorems in descendant theories
fleuriot
parents: 8867
diff changeset
   539
9043
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
   540
Goal "x * x + y * y = 0 ==> x = (0::real)";
9013
9dd0274f76af Updated files to remove 0r and 1r from theorems in descendant theories
fleuriot
parents: 8867
diff changeset
   541
by (dtac real_add_minus_eq_minus 1);
9dd0274f76af Updated files to remove 0r and 1r from theorems in descendant theories
fleuriot
parents: 8867
diff changeset
   542
by (cut_inst_tac [("x","x")] real_le_square 1);
9dd0274f76af Updated files to remove 0r and 1r from theorems in descendant theories
fleuriot
parents: 8867
diff changeset
   543
by (Auto_tac THEN dtac real_le_anti_sym 1);
9043
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
   544
by Auto_tac;
9013
9dd0274f76af Updated files to remove 0r and 1r from theorems in descendant theories
fleuriot
parents: 8867
diff changeset
   545
qed "real_sum_squares_cancel";
9dd0274f76af Updated files to remove 0r and 1r from theorems in descendant theories
fleuriot
parents: 8867
diff changeset
   546
9043
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
   547
Goal "x * x + y * y = 0 ==> y = (0::real)";
9013
9dd0274f76af Updated files to remove 0r and 1r from theorems in descendant theories
fleuriot
parents: 8867
diff changeset
   548
by (res_inst_tac [("y","x")] real_sum_squares_cancel 1);
9dd0274f76af Updated files to remove 0r and 1r from theorems in descendant theories
fleuriot
parents: 8867
diff changeset
   549
by (asm_full_simp_tac (simpset() addsimps [real_add_commute]) 1);
9dd0274f76af Updated files to remove 0r and 1r from theorems in descendant theories
fleuriot
parents: 8867
diff changeset
   550
qed "real_sum_squares_cancel2";
9dd0274f76af Updated files to remove 0r and 1r from theorems in descendant theories
fleuriot
parents: 8867
diff changeset
   551
7334
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   552
(*----------------------------------------------------------------------------
9043
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
   553
     Some convenient biconditionals for products of signs (lcp)
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
   554
 ----------------------------------------------------------------------------*)
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
   555
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
   556
Goal "((0::real) < x*y) = (0 < x & 0 < y | x < 0 & y < 0)";
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
   557
by (auto_tac (claset(), 
9069
e8d530582061 a big tidy-up
paulson
parents: 9053
diff changeset
   558
              simpset() addsimps [order_le_less, linorder_not_less, 
9043
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
   559
                                  real_mult_order, real_mult_less_zero1]));
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
   560
by (ALLGOALS (rtac ccontr)); 
9069
e8d530582061 a big tidy-up
paulson
parents: 9053
diff changeset
   561
by (auto_tac (claset(), simpset() addsimps [order_le_less, linorder_not_less]));
9043
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
   562
by (ALLGOALS (etac rev_mp)); 
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
   563
by (ALLGOALS (dtac real_mult_less_zero THEN' assume_tac));
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
   564
by (auto_tac (claset() addDs [order_less_not_sym], 
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
   565
              simpset() addsimps [real_mult_commute]));  
9069
e8d530582061 a big tidy-up
paulson
parents: 9053
diff changeset
   566
qed "real_zero_less_mult_iff";
9043
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
   567
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
   568
Goal "((0::real) <= x*y) = (0 <= x & 0 <= y | x <= 0 & y <= 0)";
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
   569
by (auto_tac (claset(), 
9069
e8d530582061 a big tidy-up
paulson
parents: 9053
diff changeset
   570
              simpset() addsimps [order_le_less, linorder_not_less,  
e8d530582061 a big tidy-up
paulson
parents: 9053
diff changeset
   571
                                  real_zero_less_mult_iff]));
e8d530582061 a big tidy-up
paulson
parents: 9053
diff changeset
   572
qed "real_zero_le_mult_iff";
9043
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
   573
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
   574
Goal "(x*y < (0::real)) = (0 < x & y < 0 | x < 0 & 0 < y)";
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
   575
by (auto_tac (claset(), 
9069
e8d530582061 a big tidy-up
paulson
parents: 9053
diff changeset
   576
              simpset() addsimps [real_zero_le_mult_iff, 
9043
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
   577
                                  linorder_not_le RS sym]));
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
   578
by (auto_tac (claset() addDs [order_less_not_sym],  
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
   579
              simpset() addsimps [linorder_not_le]));
9069
e8d530582061 a big tidy-up
paulson
parents: 9053
diff changeset
   580
qed "real_mult_less_zero_iff";
9043
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
   581
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
   582
Goal "(x*y <= (0::real)) = (0 <= x & y <= 0 | x <= 0 & 0 <= y)";
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
   583
by (auto_tac (claset() addDs [order_less_not_sym], 
9069
e8d530582061 a big tidy-up
paulson
parents: 9053
diff changeset
   584
              simpset() addsimps [real_zero_less_mult_iff, 
9043
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
   585
                                  linorder_not_less RS sym]));
9069
e8d530582061 a big tidy-up
paulson
parents: 9053
diff changeset
   586
qed "real_mult_le_zero_iff";
9043
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
   587