src/HOL/Real/RealOrd.ML
author paulson
Wed, 14 Jun 2000 18:21:25 +0200
changeset 9069 e8d530582061
parent 9053 80fca868ec4c
child 9081 d54b2c41fe0e
permissions -rw-r--r--
a big tidy-up
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
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
    33
  val thy		= RealDef.thy
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
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
    56
  val add_inverses	= [real_add_minus, real_add_minus_left];
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
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
    64
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
    65
(**** Theorems about the ordering ****)
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
    66
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
    67
Goal "(0 < x) = (EX y. x = real_of_preal y)";
7334
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
    68
by (auto_tac (claset(), simpset() addsimps [real_of_preal_zero_less]));
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
    69
by (cut_inst_tac [("x","x")] real_of_preal_trichotomy 1);
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
    70
by (blast_tac (claset() addSEs [real_less_irrefl,
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
    71
				real_of_preal_not_minus_gt_zero RS notE]) 1);
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
    72
qed "real_gt_zero_preal_Ex";
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
    73
9043
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
    74
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
    75
by (blast_tac (claset() addSDs [real_of_preal_zero_less RS real_less_trans]
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
    76
               addIs [real_gt_zero_preal_Ex RS iffD1]) 1);
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
    77
qed "real_gt_preal_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() addDs [real_le_imp_less_or_eq,
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
    81
			       real_gt_preal_preal_Ex]) 1);
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
    82
qed "real_ge_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 "y <= 0 ==> ALL x. y < real_of_preal x";
7334
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
    85
by (auto_tac (claset() addEs [real_le_imp_less_or_eq RS disjE]
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
    86
                       addIs [real_of_preal_zero_less RSN(2,real_less_trans)],
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
    87
              simpset() addsimps [real_of_preal_zero_less]));
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
    88
qed "real_less_all_preal";
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
    89
9043
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
    90
Goal "~ 0 < y ==> ALL x. y < real_of_preal x";
7334
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
    91
by (blast_tac (claset() addSIs [real_less_all_preal,real_leI]) 1);
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
    92
qed "real_less_all_real2";
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
    93
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
    94
Goal "((x::real) < y) = (-y < -x)";
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
    95
by (rtac (real_less_sum_gt_0_iff RS subst) 1);
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
    96
by (res_inst_tac [("W1","x")] (real_less_sum_gt_0_iff RS subst) 1);
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
    97
by (simp_tac (simpset() addsimps [real_add_commute]) 1);
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
    98
qed "real_less_swap_iff";
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
    99
9043
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
   100
Goal "[| R + L = S;  (0::real) < L |] ==> R < S";
7334
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   101
by (rtac (real_less_sum_gt_0_iff RS iffD1) 1);
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   102
by (auto_tac (claset(), simpset() addsimps real_add_ac));
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   103
qed "real_lemma_add_positive_imp_less";
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   104
9043
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
   105
Goal "EX T::real. 0 < T & R + T = S ==> R < S";
7334
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   106
by (blast_tac (claset() addIs [real_lemma_add_positive_imp_less]) 1);
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   107
qed "real_ex_add_positive_left_less";
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   108
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   109
(*Alternative definition for real_less.  NOT for rewriting*)
9043
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
   110
Goal "(R < S) = (EX T::real. 0 < T & R + T = S)";
7334
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   111
by (blast_tac (claset() addSIs [real_less_add_positive_left_Ex,
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   112
				real_ex_add_positive_left_less]) 1);
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   113
qed "real_less_iff_add";
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   114
9043
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
   115
Goal "((0::real) < x) = (-x < x)";
7334
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   116
by Safe_tac;
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   117
by (rtac ccontr 2 THEN forward_tac 
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   118
    [real_leI RS real_le_imp_less_or_eq] 2);
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   119
by (Step_tac 2);
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   120
by (dtac (real_minus_zero_less_iff RS iffD2) 2);
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   121
by (blast_tac (claset() addIs [real_less_trans]) 2);
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   122
by (auto_tac (claset(),
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   123
	      simpset() addsimps 
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   124
 	        [real_gt_zero_preal_Ex,real_of_preal_minus_less_self]));
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   125
qed "real_gt_zero_iff";
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::real)) = (x < -x)";
7334
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   128
by (rtac (real_minus_zero_less_iff RS subst) 1);
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   129
by (stac real_gt_zero_iff 1);
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   130
by (Full_simp_tac 1);
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   131
qed "real_lt_zero_iff";
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
Goalw [real_le_def] "((0::real) <= x) = (-x <= x)";
7334
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   134
by (auto_tac (claset(), simpset() addsimps [real_lt_zero_iff RS sym]));
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   135
qed "real_ge_zero_iff";
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   136
9043
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
   137
Goalw [real_le_def] "(x <= (0::real)) = (x <= -x)";
7334
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   138
by (auto_tac (claset(), simpset() addsimps [real_gt_zero_iff RS sym]));
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   139
qed "real_le_zero_iff";
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   140
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   141
Goal "(real_of_preal m1 <= real_of_preal m2) = (m1 <= m2)";
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   142
by (auto_tac (claset() addSIs [preal_leI],
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   143
    simpset() addsimps [real_less_le_iff RS sym]));
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   144
by (dtac preal_le_less_trans 1 THEN assume_tac 1);
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   145
by (etac preal_less_irrefl 1);
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   146
qed "real_of_preal_le_iff";
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   147
9043
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
   148
Goal "[| 0 < x; 0 < y |] ==> (0::real) < x * y";
7334
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   149
by (auto_tac (claset(), simpset() addsimps [real_gt_zero_preal_Ex]));  
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   150
by (res_inst_tac [("x","y*ya")] exI 1);
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   151
by (full_simp_tac (simpset() addsimps [real_of_preal_mult]) 1);
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   152
qed "real_mult_order";
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   153
9043
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
   154
Goal "[| x < 0; y < 0 |] ==> (0::real) < x * y";
7334
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   155
by (REPEAT(dtac (real_minus_zero_less_iff RS iffD2) 1));
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   156
by (dtac real_mult_order 1 THEN assume_tac 1);
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   157
by (Asm_full_simp_tac 1);
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   158
qed "real_mult_less_zero1";
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   159
9043
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
   160
Goal "[| 0 <= x; 0 <= y |] ==> (0::real) <= x * y";
7334
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   161
by (REPEAT(dtac real_le_imp_less_or_eq 1));
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   162
by (auto_tac (claset() addIs [real_mult_order, real_less_imp_le],
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   163
	      simpset()));
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   164
qed "real_le_mult_order";
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   165
9043
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
   166
Goal "[| 0 < x; 0 <= y |] ==> (0::real) <= x * y";
7334
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   167
by (dtac real_le_imp_less_or_eq 1);
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   168
by (auto_tac (claset() addIs [real_mult_order,
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   169
			      real_less_imp_le],simpset()));
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   170
qed "real_less_le_mult_order";
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   171
9043
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
   172
Goal "[| x <= 0; y <= 0 |] ==> (0::real) <= x * y";
7334
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   173
by (rtac real_less_or_eq_imp_le 1);
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   174
by (dtac real_le_imp_less_or_eq 1 THEN etac disjE 1);
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   175
by Auto_tac;
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   176
by (dtac real_le_imp_less_or_eq 1);
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   177
by (auto_tac (claset() addDs [real_mult_less_zero1],simpset()));
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   178
qed "real_mult_le_zero1";
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   179
9043
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
   180
Goal "[| 0 <= x; y < 0 |] ==> x * y <= (0::real)";
7334
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   181
by (rtac real_less_or_eq_imp_le 1);
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   182
by (dtac real_le_imp_less_or_eq 1 THEN etac disjE 1);
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   183
by Auto_tac;
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   184
by (dtac (real_minus_zero_less_iff RS iffD2) 1);
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   185
by (rtac (real_minus_zero_less_iff RS subst) 1);
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   186
by (blast_tac (claset() addDs [real_mult_order] 
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   187
	                addIs [real_minus_mult_eq2 RS ssubst]) 1);
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   188
qed "real_mult_le_zero";
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   189
9043
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
   190
Goal "[| 0 < x; y < 0 |] ==> x*y < (0::real)";
7334
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   191
by (dtac (real_minus_zero_less_iff RS iffD2) 1);
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   192
by (dtac real_mult_order 1 THEN assume_tac 1);
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   193
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
   194
by (Asm_full_simp_tac 1);
7334
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   195
qed "real_mult_less_zero";
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   196
9043
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
   197
Goalw [real_one_def] "0 < 1r";
7334
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   198
by (auto_tac (claset() addIs [real_gt_zero_preal_Ex RS iffD2],
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   199
	      simpset() addsimps [real_of_preal_def]));
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   200
qed "real_zero_less_one";
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   201
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   202
(*** Monotonicity results ***)
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   203
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   204
Goal "(v+z < w+z) = (v < (w::real))";
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   205
by (Simp_tac 1);
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   206
qed "real_add_right_cancel_less";
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   207
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   208
Goal "(z+v < z+w) = (v < (w::real))";
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   209
by (Simp_tac 1);
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   210
qed "real_add_left_cancel_less";
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   211
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   212
Addsimps [real_add_right_cancel_less, real_add_left_cancel_less];
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   213
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   214
Goal "(v+z <= w+z) = (v <= (w::real))";
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   215
by (Simp_tac 1);
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   216
qed "real_add_right_cancel_le";
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   217
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   218
Goal "(z+v <= z+w) = (v <= (w::real))";
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   219
by (Simp_tac 1);
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   220
qed "real_add_left_cancel_le";
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   221
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   222
Addsimps [real_add_right_cancel_le, real_add_left_cancel_le];
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   223
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   224
(*"v<=w ==> v+z <= w+z"*)
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   225
bind_thm ("real_add_less_mono1", real_add_right_cancel_less RS iffD2);
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   226
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   227
(*"v<=w ==> v+z <= w+z"*)
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   228
bind_thm ("real_add_le_mono1", real_add_right_cancel_le RS iffD2);
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   229
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   230
Goal "!!z z'::real. [| w'<w; z'<=z |] ==> w' + z' < w + z";
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   231
by (etac (real_add_less_mono1 RS real_less_le_trans) 1);
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   232
by (Simp_tac 1);
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   233
qed "real_add_less_le_mono";
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   234
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   235
Goal "!!z z'::real. [| w'<=w; z'<z |] ==> w' + z' < w + z";
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   236
by (etac (real_add_le_mono1 RS real_le_less_trans) 1);
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   237
by (Simp_tac 1);
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   238
qed "real_add_le_less_mono";
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   239
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   240
Goal "!!(A::real). A < B ==> C + A < C + B";
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   241
by (Simp_tac 1);
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   242
qed "real_add_less_mono2";
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   243
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   244
Goal "!!(A::real). A + C < B + C ==> A < B";
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   245
by (Full_simp_tac 1);
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   246
qed "real_less_add_right_cancel";
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   247
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   248
Goal "!!(A::real). C + A < C + B ==> A < B";
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   249
by (Full_simp_tac 1);
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   250
qed "real_less_add_left_cancel";
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   251
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   252
Goal "!!(A::real). A + C <= B + C ==> A <= B";
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   253
by (Full_simp_tac 1);
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   254
qed "real_le_add_right_cancel";
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   255
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   256
Goal "!!(A::real). C + A <= C + B ==> A <= B";
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   257
by (Full_simp_tac 1);
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   258
qed "real_le_add_left_cancel";
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   259
9043
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
   260
Goal "[| 0 < x; 0 < y |] ==> (0::real) < x + y";
7334
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   261
by (etac real_less_trans 1);
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   262
by (dtac real_add_less_mono2 1);
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   263
by (Full_simp_tac 1);
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   264
qed "real_add_order";
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   265
9043
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
   266
Goal "[| 0 <= x; 0 <= y |] ==> (0::real) <= x + y";
7334
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   267
by (REPEAT(dtac real_le_imp_less_or_eq 1));
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   268
by (auto_tac (claset() addIs [real_add_order, real_less_imp_le],
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   269
	      simpset()));
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   270
qed "real_le_add_order";
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   271
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   272
Goal "[| R1 < S1; R2 < S2 |] ==> R1 + R2 < S1 + (S2::real)";
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   273
by (dtac real_add_less_mono1 1);
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   274
by (etac real_less_trans 1);
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   275
by (etac real_add_less_mono2 1);
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   276
qed "real_add_less_mono";
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   277
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   278
Goal "!!(q1::real). q1 <= q2  ==> x + q1 <= x + q2";
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   279
by (Simp_tac 1);
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   280
qed "real_add_left_le_mono1";
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   281
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   282
Goal "[|i<=j;  k<=l |] ==> i + k <= j + (l::real)";
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   283
by (dtac real_add_le_mono1 1);
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   284
by (etac real_le_trans 1);
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   285
by (Simp_tac 1);
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   286
qed "real_add_le_mono";
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   287
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   288
Goal "EX (x::real). x < y";
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   289
by (rtac (real_add_zero_right RS subst) 1);
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   290
by (res_inst_tac [("x","y + (-1r)")] exI 1);
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   291
by (auto_tac (claset() addSIs [real_add_less_mono2],
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   292
	  simpset() addsimps [real_minus_zero_less_iff2, real_zero_less_one]));
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   293
qed "real_less_Ex";
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   294
9043
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
   295
Goal "(0::real) < r ==>  u + (-r) < u";
7334
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   296
by (res_inst_tac [("C","r")] real_less_add_right_cancel 1);
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   297
by (simp_tac (simpset() addsimps [real_add_assoc]) 1);
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   298
qed "real_add_minus_positive_less_self";
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   299
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   300
Goal "((r::real) <= s) = (-s <= -r)";
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   301
by (Step_tac 1);
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   302
by (dres_inst_tac [("x","-s")] real_add_left_le_mono1 1);
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   303
by (dres_inst_tac [("x","r")] real_add_left_le_mono1 2);
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   304
by Auto_tac;
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   305
by (dres_inst_tac [("z","-r")] real_add_le_mono1 1);
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   306
by (dres_inst_tac [("z","s")] real_add_le_mono1 2);
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   307
by (auto_tac (claset(), simpset() addsimps [real_add_assoc]));
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   308
qed "real_le_minus_iff";
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   309
Addsimps [real_le_minus_iff RS sym];
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   310
          
9043
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
   311
Goal "0 <= 1r";
7334
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   312
by (rtac (real_zero_less_one RS real_less_imp_le) 1);
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   313
qed "real_zero_le_one";
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   314
Addsimps [real_zero_le_one];
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   315
9043
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
   316
Goal "(0::real) <= x*x";
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
   317
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
   318
by (auto_tac (claset() addIs [real_mult_order,
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   319
			      real_mult_less_zero1,real_less_imp_le],
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   320
	      simpset()));
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   321
qed "real_le_square";
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   322
Addsimps [real_le_square];
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   323
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   324
(*----------------------------------------------------------------------------
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   325
             An embedding of the naturals in the reals
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   326
 ----------------------------------------------------------------------------*)
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   327
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   328
Goalw [real_of_posnat_def] "real_of_posnat 0 = 1r";
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   329
by (full_simp_tac (simpset() addsimps [pnat_one_iff RS sym,real_of_preal_def]) 1);
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   330
by (fold_tac [real_one_def]);
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   331
by (rtac refl 1);
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   332
qed "real_of_posnat_one";
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   333
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   334
Goalw [real_of_posnat_def] "real_of_posnat 1 = 1r + 1r";
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   335
by (full_simp_tac (simpset() addsimps [real_of_preal_def,real_one_def,
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   336
    pnat_two_eq,real_add,prat_of_pnat_add RS sym,preal_of_prat_add RS sym
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   337
    ] @ pnat_add_ac) 1);
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   338
qed "real_of_posnat_two";
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   339
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   340
Goalw [real_of_posnat_def]
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   341
    "real_of_posnat n1 + real_of_posnat n2 = real_of_posnat (n1 + n2) + 1r";
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   342
by (full_simp_tac (simpset() addsimps [real_of_posnat_one RS sym,
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   343
    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
   344
    prat_of_pnat_add RS sym,pnat_of_nat_add]) 1);
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   345
qed "real_of_posnat_add";
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   346
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   347
Goal "real_of_posnat (n + 1) = real_of_posnat n + 1r";
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   348
by (res_inst_tac [("x1","1r")] (real_add_right_cancel RS iffD1) 1);
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   349
by (rtac (real_of_posnat_add RS subst) 1);
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   350
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
   351
qed "real_of_posnat_add_one";
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   352
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   353
Goal "real_of_posnat (Suc n) = real_of_posnat n + 1r";
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   354
by (stac (real_of_posnat_add_one RS sym) 1);
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   355
by (Simp_tac 1);
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   356
qed "real_of_posnat_Suc";
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   357
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   358
Goal "inj(real_of_posnat)";
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   359
by (rtac injI 1);
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   360
by (rewtac real_of_posnat_def);
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   361
by (dtac (inj_real_of_preal RS injD) 1);
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   362
by (dtac (inj_preal_of_prat RS injD) 1);
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   363
by (dtac (inj_prat_of_pnat RS injD) 1);
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   364
by (etac (inj_pnat_of_nat RS injD) 1);
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   365
qed "inj_real_of_posnat";
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   366
9043
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
   367
Goalw [real_of_posnat_def] "0 < real_of_posnat n";
7334
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   368
by (rtac (real_gt_zero_preal_Ex RS iffD2) 1);
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   369
by (Blast_tac 1);
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   370
qed "real_of_posnat_less_zero";
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   371
9043
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
   372
Goal "real_of_posnat n ~= 0";
7334
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   373
by (rtac (real_of_posnat_less_zero RS real_not_refl2 RS not_sym) 1);
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   374
qed "real_of_posnat_not_eq_zero";
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   375
Addsimps[real_of_posnat_not_eq_zero];
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   376
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   377
Goal "1r <= real_of_posnat n";
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   378
by (simp_tac (simpset() addsimps [real_of_posnat_one RS sym]) 1);
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   379
by (induct_tac "n" 1);
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   380
by (auto_tac (claset(),
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   381
	      simpset () addsimps [real_of_posnat_Suc,real_of_posnat_one,
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   382
			   real_of_posnat_less_zero, real_less_imp_le]));
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   383
qed "real_of_posnat_less_one";
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   384
Addsimps [real_of_posnat_less_one];
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   385
9043
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
   386
Goal "rinv(real_of_posnat n) ~= 0";
7334
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   387
by (rtac ((real_of_posnat_less_zero RS 
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   388
    real_not_refl2 RS not_sym) RS rinv_not_zero) 1);
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   389
qed "real_of_posnat_rinv_not_zero";
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   390
Addsimps [real_of_posnat_rinv_not_zero];
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   391
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   392
Goal "rinv(real_of_posnat x) = rinv(real_of_posnat y) ==> x = y";
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   393
by (rtac (inj_real_of_posnat RS injD) 1);
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   394
by (res_inst_tac [("n2","x")] 
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   395
    (real_of_posnat_rinv_not_zero RS real_mult_left_cancel RS iffD1) 1);
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   396
by (full_simp_tac (simpset() addsimps [(real_of_posnat_less_zero RS 
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   397
    real_not_refl2 RS not_sym) RS real_mult_inv_left]) 1);
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   398
by (asm_full_simp_tac (simpset() addsimps [(real_of_posnat_less_zero RS 
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   399
    real_not_refl2 RS not_sym)]) 1);
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   400
qed "real_of_posnat_rinv_inj";
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   401
9043
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
   402
Goal "0 < x ==> 0 < rinv x";
7334
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   403
by (EVERY1[rtac ccontr, dtac real_leI]);
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   404
by (forward_tac [real_minus_zero_less_iff2 RS iffD2] 1);
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   405
by (forward_tac [real_not_refl2 RS not_sym] 1);
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   406
by (dtac (real_not_refl2 RS not_sym RS rinv_not_zero) 1);
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   407
by (EVERY1[dtac real_le_imp_less_or_eq, Step_tac]); 
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   408
by (dtac real_mult_less_zero1 1 THEN assume_tac 1);
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   409
by (auto_tac (claset() addIs [real_zero_less_one RS real_less_asym],
9053
paulson
parents: 9043
diff changeset
   410
	      simpset()));
7334
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   411
qed "real_rinv_gt_zero";
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   412
9043
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
   413
Goal "x < 0 ==> rinv x < 0";
7499
23e090051cb8 isatool expandshort;
wenzelm
parents: 7334
diff changeset
   414
by (ftac real_not_refl2 1);
7334
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   415
by (dtac (real_minus_zero_less_iff RS iffD2) 1);
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   416
by (rtac (real_minus_zero_less_iff RS iffD1) 1);
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   417
by (dtac (real_minus_rinv RS sym) 1);
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   418
by (auto_tac (claset() addIs [real_rinv_gt_zero], simpset()));
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   419
qed "real_rinv_less_zero";
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   420
9043
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
   421
Goal "0 < rinv(real_of_posnat n)";
7334
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   422
by (rtac (real_of_posnat_less_zero RS real_rinv_gt_zero) 1);
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   423
qed "real_of_posnat_rinv_gt_zero";
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   424
Addsimps [real_of_posnat_rinv_gt_zero];
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   425
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   426
Goal "x+x = x*(1r+1r)";
9043
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
   427
by (simp_tac (simpset() addsimps [real_add_mult_distrib2]) 1);
7334
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   428
qed "real_add_self";
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   429
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   430
Goal "x < x + 1r";
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   431
by (rtac (real_less_sum_gt_0_iff RS iffD1) 1);
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   432
by (full_simp_tac (simpset() addsimps [real_zero_less_one,
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   433
				real_add_assoc, real_add_left_commute]) 1);
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   434
qed "real_self_less_add_one";
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   435
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   436
Goal "1r < 1r + 1r";
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   437
by (rtac real_self_less_add_one 1);
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   438
qed "real_one_less_two";
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   439
9043
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
   440
Goal "0 < 1r + 1r";
7334
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   441
by (rtac ([real_zero_less_one,
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   442
	   real_one_less_two] MRS real_less_trans) 1);
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   443
qed "real_zero_less_two";
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   444
9043
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
   445
Goal "1r + 1r ~= 0";
7334
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   446
by (rtac (real_zero_less_two RS real_not_refl2 RS not_sym) 1);
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   447
qed "real_two_not_zero";
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   448
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   449
Addsimps [real_two_not_zero];
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   450
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   451
Goal "x*rinv(1r + 1r) + x*rinv(1r + 1r) = x";
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   452
by (stac real_add_self 1);
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   453
by (full_simp_tac (simpset() addsimps [real_mult_assoc]) 1);
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   454
qed "real_sum_of_halves";
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
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";       
7334
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   457
by (rotate_tac 1 1);
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   458
by (dtac real_less_sum_gt_zero 1);
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   459
by (rtac real_sum_gt_zero_less 1);
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   460
by (dtac real_mult_order 1 THEN assume_tac 1);
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   461
by (asm_full_simp_tac (simpset() addsimps [real_add_mult_distrib2,
9053
paulson
parents: 9043
diff changeset
   462
					   real_mult_commute ]) 1);
7334
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   463
qed "real_mult_less_mono1";
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   464
9043
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
   465
Goal "[| (0::real) < z; x < y |] ==> z*x < z*y";       
7334
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   466
by (asm_simp_tac (simpset() addsimps [real_mult_commute,real_mult_less_mono1]) 1);
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   467
qed "real_mult_less_mono2";
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   468
9043
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
   469
Goal "[| (0::real) < z; x*z < y*z |] ==> x < y";
7334
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   470
by (forw_inst_tac [("x","x*z")] (real_rinv_gt_zero 
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   471
                       RS real_mult_less_mono1) 1);
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   472
by (auto_tac (claset(),
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   473
	      simpset() addsimps 
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   474
     [real_mult_assoc,real_not_refl2 RS not_sym]));
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   475
qed "real_mult_less_cancel1";
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   476
9043
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
   477
Goal "[| (0::real) < z; z*x < z*y |] ==> x < y";
7334
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   478
by (etac real_mult_less_cancel1 1);
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   479
by (asm_full_simp_tac (simpset() addsimps [real_mult_commute]) 1);
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   480
qed "real_mult_less_cancel2";
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   481
9043
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
   482
Goal "(0::real) < z ==> (x*z < y*z) = (x < y)";
7334
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   483
by (blast_tac (claset() addIs [real_mult_less_mono1,
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   484
    real_mult_less_cancel1]) 1);
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   485
qed "real_mult_less_iff1";
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   486
9043
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
   487
Goal "(0::real) < z ==> (z*x < z*y) = (x < y)";
7334
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   488
by (blast_tac (claset() addIs [real_mult_less_mono2,
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   489
    real_mult_less_cancel2]) 1);
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   490
qed "real_mult_less_iff2";
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   491
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   492
Addsimps [real_mult_less_iff1,real_mult_less_iff2];
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   493
9013
9dd0274f76af Updated files to remove 0r and 1r from theorems in descendant theories
fleuriot
parents: 8867
diff changeset
   494
(* 05/00 *)
9043
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
   495
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
   496
by (Auto_tac);
9dd0274f76af Updated files to remove 0r and 1r from theorems in descendant theories
fleuriot
parents: 8867
diff changeset
   497
qed "real_mult_le_cancel1";
9dd0274f76af Updated files to remove 0r and 1r from theorems in descendant theories
fleuriot
parents: 8867
diff changeset
   498
9043
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
   499
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
   500
by (Auto_tac);
9dd0274f76af Updated files to remove 0r and 1r from theorems in descendant theories
fleuriot
parents: 8867
diff changeset
   501
qed "real_mult_le_cancel2";
9dd0274f76af Updated files to remove 0r and 1r from theorems in descendant theories
fleuriot
parents: 8867
diff changeset
   502
9043
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
   503
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
   504
by (Auto_tac);
9dd0274f76af Updated files to remove 0r and 1r from theorems in descendant theories
fleuriot
parents: 8867
diff changeset
   505
qed "real_mult_le_cancel_iff1";
9dd0274f76af Updated files to remove 0r and 1r from theorems in descendant theories
fleuriot
parents: 8867
diff changeset
   506
9043
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
   507
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
   508
by (Auto_tac);
9dd0274f76af Updated files to remove 0r and 1r from theorems in descendant theories
fleuriot
parents: 8867
diff changeset
   509
qed "real_mult_le_cancel_iff2";
9dd0274f76af Updated files to remove 0r and 1r from theorems in descendant theories
fleuriot
parents: 8867
diff changeset
   510
9dd0274f76af Updated files to remove 0r and 1r from theorems in descendant theories
fleuriot
parents: 8867
diff changeset
   511
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
   512
9dd0274f76af Updated files to remove 0r and 1r from theorems in descendant theories
fleuriot
parents: 8867
diff changeset
   513
9043
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
   514
Goal "[| (0::real) <= z; x < y |] ==> x*z <= y*z";
7334
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   515
by (EVERY1 [rtac real_less_or_eq_imp_le, dtac real_le_imp_less_or_eq]);
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   516
by (auto_tac (claset() addIs [real_mult_less_mono1],simpset()));
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   517
qed "real_mult_le_less_mono1";
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   518
9043
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
   519
Goal "[| (0::real) <= z; x < y |] ==> z*x <= z*y";
7334
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   520
by (asm_simp_tac (simpset() addsimps [real_mult_commute,real_mult_le_less_mono1]) 1);
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   521
qed "real_mult_le_less_mono2";
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   522
9043
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
   523
Goal "[| (0::real) <= z; x <= y |] ==> z*x <= z*y";
7334
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   524
by (dres_inst_tac [("x","x")] real_le_imp_less_or_eq 1);
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   525
by (auto_tac (claset() addIs [real_mult_le_less_mono2], simpset()));
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   526
qed "real_mult_le_le_mono1";
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   527
9043
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
   528
Goal "[| (0::real) < r1; r1 < r2; 0 < x; x < y|] ==> r1 * x < r2 * y";
7334
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   529
by (dres_inst_tac [("x","x")] real_mult_less_mono2 1);
9043
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
   530
by (dres_inst_tac [("R1.0","0")] real_less_trans 2);
7334
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   531
by (dres_inst_tac [("x","r1")] real_mult_less_mono1 3);
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   532
by Auto_tac;
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   533
by (blast_tac (claset() addIs [real_less_trans]) 1);
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   534
qed "real_mult_less_mono";
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   535
9043
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
   536
Goal "[| (0::real) < r1; r1  < r2;  0 < y|] ==> 0 < r2 * y";
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
   537
by (rtac real_mult_order 1); 
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
   538
by (assume_tac 2);
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
   539
by (blast_tac (claset() addIs [real_less_trans]) 1);
7334
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   540
qed "real_mult_order_trans";
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   541
9043
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
   542
Goal "[| (0::real) < r1; r1  < r2;  0 <= x; x < y|] ==> r1 * x < r2 * y";
7334
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   543
by (auto_tac (claset() addSDs [real_le_imp_less_or_eq] 
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   544
	               addIs [real_mult_less_mono,real_mult_order_trans],
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   545
              simpset()));
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   546
qed "real_mult_less_mono3";
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   547
9043
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
   548
Goal "[| (0::real) <= r1; r1  < r2;  0 <= x; x < y|] ==> r1 * x < r2 * y";
7334
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   549
by (auto_tac (claset() addSDs [real_le_imp_less_or_eq] 
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   550
	               addIs [real_mult_less_mono,real_mult_order_trans,
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   551
			      real_mult_order],
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   552
	      simpset()));
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   553
by (dres_inst_tac [("R2.0","x")] real_less_trans 1);
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   554
by (assume_tac 1);
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   555
by (blast_tac (claset() addIs [real_mult_order]) 1);
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   556
qed "real_mult_less_mono4";
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   557
9043
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
   558
Goal "[| (0::real) < r1; r1 <= r2;  0 <= x; x <= y |] ==> r1 * x <= r2 * y";
7334
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   559
by (rtac real_less_or_eq_imp_le 1);
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   560
by (REPEAT(dtac real_le_imp_less_or_eq 1));
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   561
by (auto_tac (claset() addIs [real_mult_less_mono,
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   562
			      real_mult_order_trans,real_mult_order],
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   563
	      simpset()));
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   564
qed "real_mult_le_mono";
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   565
9043
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
   566
Goal "[| (0::real) < r1; r1 < r2;  0 <= x; x <= y |] ==> r1 * x <= r2 * y";
7334
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   567
by (rtac real_less_or_eq_imp_le 1);
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   568
by (REPEAT(dtac real_le_imp_less_or_eq 1));
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   569
by (auto_tac (claset() addIs [real_mult_less_mono, real_mult_order_trans,
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   570
			      real_mult_order],
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   571
	      simpset()));
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   572
qed "real_mult_le_mono2";
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   573
9043
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
   574
Goal "[| (0::real) <= r1; r1 < r2;  0 <= x; x <= y |] ==> r1 * x <= r2 * y";
7334
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   575
by (dtac real_le_imp_less_or_eq 1);
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   576
by (auto_tac (claset() addIs [real_mult_le_mono2],simpset()));
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   577
by (dtac real_le_trans 1 THEN assume_tac 1);
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   578
by (auto_tac (claset() addIs [real_less_le_mult_order], simpset()));
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   579
qed "real_mult_le_mono3";
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   580
9043
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
   581
Goal "[| (0::real) <= r1; r1 <= r2;  0 <= x; x <= y |] ==> r1 * x <= r2 * y";
7334
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   582
by (dres_inst_tac [("x","r1")] real_le_imp_less_or_eq 1);
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   583
by (auto_tac (claset() addIs [real_mult_le_mono3, real_mult_le_le_mono1],
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   584
	      simpset()));
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   585
qed "real_mult_le_mono4";
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   586
9043
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
   587
Goal "1r <= x ==> 0 < x";
7334
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   588
by (rtac ccontr 1 THEN dtac real_leI 1);
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   589
by (dtac real_le_trans 1 THEN assume_tac 1);
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   590
by (auto_tac (claset() addDs [real_zero_less_one RSN (2,real_le_less_trans)],
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   591
	      simpset() addsimps [real_less_not_refl]));
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   592
qed "real_gt_zero";
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   593
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   594
Goal "[| 1r < r; 1r <= x |]  ==> x <= r * x";
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   595
by (dtac (real_gt_zero RS real_less_imp_le) 1);
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   596
by (auto_tac (claset() addSDs [real_mult_le_less_mono1],
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   597
    simpset()));
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   598
qed "real_mult_self_le";
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   599
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   600
Goal "[| 1r <= r; 1r <= x |]  ==> x <= r * x";
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   601
by (dtac real_le_imp_less_or_eq 1);
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   602
by (auto_tac (claset() addIs [real_mult_self_le],
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   603
	      simpset() addsimps [real_le_refl]));
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   604
qed "real_mult_self_le2";
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   605
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   606
Goal "x < y ==> x < (x + y)*rinv(1r + 1r)";
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   607
by (dres_inst_tac [("C","x")] real_add_less_mono2 1);
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   608
by (dtac (real_add_self RS subst) 1);
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   609
by (dtac (real_zero_less_two RS real_rinv_gt_zero RS 
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   610
          real_mult_less_mono1) 1);
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   611
by (asm_full_simp_tac (simpset() addsimps [real_mult_assoc]) 1);
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   612
qed "real_less_half_sum";
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   613
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   614
Goal "x < y ==> (x + y)*rinv(1r + 1r) < y";
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   615
by (dtac real_add_less_mono1 1);
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   616
by (dtac (real_add_self RS subst) 1);
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   617
by (dtac (real_zero_less_two RS real_rinv_gt_zero RS 
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   618
          real_mult_less_mono1) 1);
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   619
by (asm_full_simp_tac (simpset() addsimps [real_mult_assoc]) 1);
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   620
qed "real_gt_half_sum";
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   621
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   622
Goal "x < y ==> EX r::real. x < r & r < y";
9043
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
   623
by (blast_tac (claset() addSIs [real_less_half_sum, real_gt_half_sum]) 1);
7334
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   624
qed "real_dense";
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   625
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   626
Goal "(EX n. rinv(real_of_posnat n) < r) = (EX n. 1r < r * real_of_posnat n)";
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   627
by (Step_tac 1);
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   628
by (dres_inst_tac [("n1","n")] (real_of_posnat_less_zero 
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   629
				RS real_mult_less_mono1) 1);
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   630
by (dres_inst_tac [("n2","n")] (real_of_posnat_less_zero RS 
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   631
				real_rinv_gt_zero RS real_mult_less_mono1) 2);
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   632
by (auto_tac (claset(),
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   633
	      simpset() addsimps [(real_of_posnat_less_zero RS 
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   634
				   real_not_refl2 RS not_sym),
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   635
				  real_mult_assoc]));
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   636
qed "real_of_posnat_rinv_Ex_iff";
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   637
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   638
Goal "(rinv(real_of_posnat n) < r) = (1r < r * real_of_posnat n)";
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   639
by (Step_tac 1);
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   640
by (dres_inst_tac [("n1","n")] (real_of_posnat_less_zero 
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   641
                       RS real_mult_less_mono1) 1);
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   642
by (dres_inst_tac [("n2","n")] (real_of_posnat_less_zero RS 
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   643
				real_rinv_gt_zero RS real_mult_less_mono1) 2);
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   644
by (auto_tac (claset(), simpset() addsimps [real_mult_assoc]));
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   645
qed "real_of_posnat_rinv_iff";
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   646
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   647
Goal "(rinv(real_of_posnat n) <= r) = (1r <= r * real_of_posnat n)";
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   648
by (Step_tac 1);
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   649
by (dres_inst_tac [("n2","n")] (real_of_posnat_less_zero RS 
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   650
    real_less_imp_le RS real_mult_le_le_mono1) 1);
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   651
by (dres_inst_tac [("n3","n")] (real_of_posnat_less_zero RS 
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   652
        real_rinv_gt_zero RS real_less_imp_le RS 
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   653
        real_mult_le_le_mono1) 2);
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   654
by (auto_tac (claset(), simpset() addsimps real_mult_ac));
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   655
qed "real_of_posnat_rinv_le_iff";
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   656
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   657
Goalw [real_of_posnat_def] "(real_of_posnat n < real_of_posnat m) = (n < m)";
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   658
by Auto_tac;
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   659
qed "real_of_posnat_less_iff";
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   660
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   661
Addsimps [real_of_posnat_less_iff];
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   662
9043
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
   663
Goal "0 < u  ==> (u < rinv (real_of_posnat n)) = (real_of_posnat n < rinv(u))";
7334
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   664
by (Step_tac 1);
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   665
by (res_inst_tac [("n2","n")] (real_of_posnat_less_zero RS 
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   666
    real_rinv_gt_zero RS real_mult_less_cancel1) 1);
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   667
by (res_inst_tac [("x1","u")] ( real_rinv_gt_zero
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   668
   RS real_mult_less_cancel1) 2);
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   669
by (auto_tac (claset(),
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   670
	      simpset() addsimps [real_of_posnat_less_zero, 
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   671
    real_not_refl2 RS not_sym]));
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   672
by (res_inst_tac [("z","u")] real_mult_less_cancel2 1);
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   673
by (res_inst_tac [("n1","n")] (real_of_posnat_less_zero RS 
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   674
    real_mult_less_cancel2) 3);
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   675
by (auto_tac (claset(),
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   676
	      simpset() addsimps [real_of_posnat_less_zero, 
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   677
    real_not_refl2 RS not_sym,real_mult_assoc RS sym]));
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   678
qed "real_of_posnat_less_rinv_iff";
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   679
9043
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
   680
Goal "0 < u ==> (u = rinv(real_of_posnat n)) = (real_of_posnat n = rinv u)";
7334
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   681
by (auto_tac (claset(),
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   682
	      simpset() addsimps [real_rinv_rinv,
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   683
    real_of_posnat_less_zero,real_not_refl2 RS not_sym]));
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   684
qed "real_of_posnat_rinv_eq_iff";
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   685
9043
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
   686
Goal "[| 0 < r; r < x |] ==> rinv x < rinv r";
7499
23e090051cb8 isatool expandshort;
wenzelm
parents: 7334
diff changeset
   687
by (ftac real_less_trans 1 THEN assume_tac 1);
23e090051cb8 isatool expandshort;
wenzelm
parents: 7334
diff changeset
   688
by (ftac real_rinv_gt_zero 1);
7334
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   689
by (forw_inst_tac [("x","x")] real_rinv_gt_zero 1);
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   690
by (forw_inst_tac [("x","r"),("z","rinv r")] real_mult_less_mono1 1);
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   691
by (assume_tac 1);
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   692
by (asm_full_simp_tac (simpset() addsimps [real_not_refl2 RS 
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   693
					   not_sym RS real_mult_inv_right]) 1);
7499
23e090051cb8 isatool expandshort;
wenzelm
parents: 7334
diff changeset
   694
by (ftac real_rinv_gt_zero 1);
7334
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   695
by (forw_inst_tac [("x","1r"),("z","rinv x")] real_mult_less_mono2 1);
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   696
by (assume_tac 1);
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   697
by (asm_full_simp_tac (simpset() addsimps [real_not_refl2 RS 
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   698
         not_sym RS real_mult_inv_left,real_mult_assoc RS sym]) 1);
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   699
qed "real_rinv_less_swap";
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   700
9043
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
   701
Goal "[| 0 < r; 0 < x|] ==> (r < x) = (rinv x < rinv r)";
7334
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   702
by (auto_tac (claset() addIs [real_rinv_less_swap],simpset()));
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   703
by (res_inst_tac [("t","r")] (real_rinv_rinv RS subst) 1);
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   704
by (etac (real_not_refl2 RS not_sym) 1);
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   705
by (res_inst_tac [("t","x")] (real_rinv_rinv RS subst) 1);
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   706
by (etac (real_not_refl2 RS not_sym) 1);
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   707
by (auto_tac (claset() addIs [real_rinv_less_swap],
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   708
	      simpset() addsimps [real_rinv_gt_zero]));
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   709
qed "real_rinv_less_iff";
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   710
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   711
Goal "r < r + rinv(real_of_posnat n)";
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   712
by (res_inst_tac [("C","-r")] real_less_add_left_cancel 1);
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   713
by (full_simp_tac (simpset() addsimps [real_add_assoc RS sym]) 1);
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   714
qed "real_add_rinv_real_of_posnat_less";
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   715
Addsimps [real_add_rinv_real_of_posnat_less];
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   716
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   717
Goal "r <= r + rinv(real_of_posnat n)";
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   718
by (rtac real_less_imp_le 1);
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   719
by (Simp_tac 1);
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   720
qed "real_add_rinv_real_of_posnat_le";
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   721
Addsimps [real_add_rinv_real_of_posnat_le];
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   722
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   723
Goal "r + (-rinv(real_of_posnat n)) < r";
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   724
by (res_inst_tac [("C","-r")] real_less_add_left_cancel 1);
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   725
by (full_simp_tac (simpset() addsimps [real_add_assoc RS sym,
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   726
				       real_minus_zero_less_iff2]) 1);
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   727
qed "real_add_minus_rinv_real_of_posnat_less";
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   728
Addsimps [real_add_minus_rinv_real_of_posnat_less];
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   729
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   730
Goal "r + (-rinv(real_of_posnat n)) <= r";
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   731
by (rtac real_less_imp_le 1);
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   732
by (Simp_tac 1);
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   733
qed "real_add_minus_rinv_real_of_posnat_le";
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   734
Addsimps [real_add_minus_rinv_real_of_posnat_le];
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   735
9043
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
   736
Goal "0 < r ==> r*(1r + (-rinv(real_of_posnat n))) < r";
7334
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   737
by (simp_tac (simpset() addsimps [real_add_mult_distrib2]) 1);
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   738
by (res_inst_tac [("C","-r")] real_less_add_left_cancel 1);
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   739
by (auto_tac (claset() addIs [real_mult_order],
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   740
	      simpset() addsimps [real_add_assoc RS sym,
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   741
				  real_minus_zero_less_iff2]));
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   742
qed "real_mult_less_self";
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   743
9043
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
   744
Goal "0 <= 1r + (-rinv(real_of_posnat n))";
7334
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   745
by (res_inst_tac [("C","rinv(real_of_posnat n)")] real_le_add_right_cancel 1);
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   746
by (simp_tac (simpset() addsimps [real_add_assoc,
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   747
				  real_of_posnat_rinv_le_iff]) 1);
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   748
qed "real_add_one_minus_rinv_ge_zero";
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   749
9043
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
   750
Goal "0 < r ==> 0 <= r*(1r + (-rinv(real_of_posnat n)))";
7334
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   751
by (dtac (real_add_one_minus_rinv_ge_zero RS real_mult_le_less_mono1) 1);
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   752
by Auto_tac;
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   753
qed "real_mult_add_one_minus_ge_zero";
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   754
9043
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
   755
Goal "(x*y = 0) = (x = 0 | y = (0::real))";
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
   756
by Auto_tac;
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
   757
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
   758
qed "real_mult_is_0";
7334
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   759
9043
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
   760
Goal "(0 = x*y) = (0 = x | (0::real) = y)";
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
   761
by (stac eq_commute 1 THEN stac real_mult_is_0 1);
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
   762
by Auto_tac;
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
   763
qed "real_0_is_mult";
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
   764
AddIffs [real_mult_is_0, real_0_is_mult];
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
   765
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
   766
Goal "[| x ~= 1r; y * x = y |] ==> y = 0";
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
   767
by (subgoal_tac "y*(1r + -x) = 0" 1);
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
   768
by (stac real_add_mult_distrib2 2);
7334
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   769
by (auto_tac (claset(),
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   770
	      simpset() addsimps [real_eq_minus_iff2 RS sym]));
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   771
qed "real_mult_eq_self_zero";
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   772
Addsimps [real_mult_eq_self_zero];
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   773
9043
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
   774
Goal "[| x ~= 1r; y = y * x |] ==> y = 0";
7334
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   775
by (dtac sym 1);
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   776
by (Asm_full_simp_tac 1);
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   777
qed "real_mult_eq_self_zero2";
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   778
Addsimps [real_mult_eq_self_zero2];
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   779
9043
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
   780
Goal "[| 0 <= x*y; 0 < x |] ==> (0::real) <= y";
7499
23e090051cb8 isatool expandshort;
wenzelm
parents: 7334
diff changeset
   781
by (ftac real_rinv_gt_zero 1);
7334
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   782
by (dres_inst_tac [("x","rinv x")] real_less_le_mult_order 1);
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   783
by (dtac (real_not_refl2 RS not_sym RS real_mult_inv_left) 2);
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   784
by (auto_tac (claset(),
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   785
	      simpset() addsimps [real_mult_assoc RS sym]));
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   786
qed "real_mult_ge_zero_cancel";
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   787
9043
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
   788
Goal "[|x ~= 0; y ~= 0 |] ==> rinv(x) + rinv(y) = (x + y)*rinv(x*y)";
7334
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   789
by (asm_full_simp_tac (simpset() addsimps 
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   790
		       [real_rinv_distrib,real_add_mult_distrib,
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   791
			real_mult_assoc RS sym]) 1);
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   792
by (stac real_mult_assoc 1);
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   793
by (rtac (real_mult_left_commute RS subst) 1);
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   794
by (asm_full_simp_tac (simpset() addsimps [real_add_commute]) 1);
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   795
qed "real_rinv_add";
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   796
9013
9dd0274f76af Updated files to remove 0r and 1r from theorems in descendant theories
fleuriot
parents: 8867
diff changeset
   797
(* 05/00 *)
9043
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
   798
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
   799
by (auto_tac (claset() addDs [sym],
9dd0274f76af Updated files to remove 0r and 1r from theorems in descendant theories
fleuriot
parents: 8867
diff changeset
   800
    simpset() addsimps [real_le_less]));
9dd0274f76af Updated files to remove 0r and 1r from theorems in descendant theories
fleuriot
parents: 8867
diff changeset
   801
qed "real_minus_zero_le_iff";
9dd0274f76af Updated files to remove 0r and 1r from theorems in descendant theories
fleuriot
parents: 8867
diff changeset
   802
9043
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
   803
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
   804
by (auto_tac (claset(),simpset() addsimps 
9dd0274f76af Updated files to remove 0r and 1r from theorems in descendant theories
fleuriot
parents: 8867
diff changeset
   805
    [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
   806
qed "real_minus_zero_le_iff2";
9dd0274f76af Updated files to remove 0r and 1r from theorems in descendant theories
fleuriot
parents: 8867
diff changeset
   807
9053
paulson
parents: 9043
diff changeset
   808
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
   809
9043
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
   810
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
   811
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
   812
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
   813
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
   814
by Auto_tac;
9013
9dd0274f76af Updated files to remove 0r and 1r from theorems in descendant theories
fleuriot
parents: 8867
diff changeset
   815
qed "real_sum_squares_cancel";
9dd0274f76af Updated files to remove 0r and 1r from theorems in descendant theories
fleuriot
parents: 8867
diff changeset
   816
9043
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
   817
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
   818
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
   819
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
   820
qed "real_sum_squares_cancel2";
9dd0274f76af Updated files to remove 0r and 1r from theorems in descendant theories
fleuriot
parents: 8867
diff changeset
   821
7334
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   822
(*----------------------------------------------------------------------------
9043
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
   823
     Some convenient biconditionals for products of signs (lcp)
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
   824
 ----------------------------------------------------------------------------*)
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
   825
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
   826
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
   827
by (auto_tac (claset(), 
9069
e8d530582061 a big tidy-up
paulson
parents: 9053
diff changeset
   828
              simpset() addsimps [order_le_less, linorder_not_less, 
9043
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
   829
                                  real_mult_order, real_mult_less_zero1]));
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
   830
by (ALLGOALS (rtac ccontr)); 
9069
e8d530582061 a big tidy-up
paulson
parents: 9053
diff changeset
   831
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
   832
by (ALLGOALS (etac rev_mp)); 
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
   833
by (ALLGOALS (dtac real_mult_less_zero THEN' assume_tac));
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
   834
by (auto_tac (claset() addDs [order_less_not_sym], 
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
   835
              simpset() addsimps [real_mult_commute]));  
9069
e8d530582061 a big tidy-up
paulson
parents: 9053
diff changeset
   836
qed "real_zero_less_mult_iff";
9043
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
   837
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
   838
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
   839
by (auto_tac (claset(), 
9069
e8d530582061 a big tidy-up
paulson
parents: 9053
diff changeset
   840
              simpset() addsimps [order_le_less, linorder_not_less,  
e8d530582061 a big tidy-up
paulson
parents: 9053
diff changeset
   841
                                  real_zero_less_mult_iff]));
e8d530582061 a big tidy-up
paulson
parents: 9053
diff changeset
   842
qed "real_zero_le_mult_iff";
9043
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
   843
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
   844
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
   845
by (auto_tac (claset(), 
9069
e8d530582061 a big tidy-up
paulson
parents: 9053
diff changeset
   846
              simpset() addsimps [real_zero_le_mult_iff, 
9043
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
   847
                                  linorder_not_le RS sym]));
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
   848
by (auto_tac (claset() addDs [order_less_not_sym],  
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
   849
              simpset() addsimps [linorder_not_le]));
9069
e8d530582061 a big tidy-up
paulson
parents: 9053
diff changeset
   850
qed "real_mult_less_zero_iff";
9043
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
   851
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
   852
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
   853
by (auto_tac (claset() addDs [order_less_not_sym], 
9069
e8d530582061 a big tidy-up
paulson
parents: 9053
diff changeset
   854
              simpset() addsimps [real_zero_less_mult_iff, 
9043
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
   855
                                  linorder_not_less RS sym]));
9069
e8d530582061 a big tidy-up
paulson
parents: 9053
diff changeset
   856
qed "real_mult_le_zero_iff";
9043
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
   857
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
   858
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
   859
(*----------------------------------------------------------------------------
7334
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   860
     Another embedding of the naturals in the reals (see real_of_posnat)
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   861
 ----------------------------------------------------------------------------*)
9043
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
   862
Goalw [real_of_nat_def] "real_of_nat 0 = 0";
9069
e8d530582061 a big tidy-up
paulson
parents: 9053
diff changeset
   863
by (simp_tac (simpset() addsimps [real_of_posnat_one]) 1);
7334
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   864
qed "real_of_nat_zero";
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   865
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   866
Goalw [real_of_nat_def] "real_of_nat 1 = 1r";
9069
e8d530582061 a big tidy-up
paulson
parents: 9053
diff changeset
   867
by (simp_tac (simpset() addsimps [real_of_posnat_two, real_add_assoc]) 1);
7334
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   868
qed "real_of_nat_one";
9069
e8d530582061 a big tidy-up
paulson
parents: 9053
diff changeset
   869
Addsimps [real_of_nat_zero, real_of_nat_one];
7334
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   870
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   871
Goalw [real_of_nat_def]
9069
e8d530582061 a big tidy-up
paulson
parents: 9053
diff changeset
   872
     "real_of_nat (m + n) = real_of_nat m + real_of_nat n";
7334
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   873
by (simp_tac (simpset() addsimps 
9069
e8d530582061 a big tidy-up
paulson
parents: 9053
diff changeset
   874
              [real_of_posnat_add,real_add_assoc RS sym]) 1);
7334
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   875
qed "real_of_nat_add";
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   876
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   877
Goalw [real_of_nat_def] "real_of_nat (Suc n) = real_of_nat n + 1r";
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   878
by (simp_tac (simpset() addsimps [real_of_posnat_Suc] @ real_add_ac) 1);
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   879
qed "real_of_nat_Suc";
9069
e8d530582061 a big tidy-up
paulson
parents: 9053
diff changeset
   880
Addsimps [real_of_nat_Suc];
7334
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   881
    
9069
e8d530582061 a big tidy-up
paulson
parents: 9053
diff changeset
   882
Goalw [real_of_nat_def] "(real_of_nat n < real_of_nat m) = (n < m)";
7334
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   883
by Auto_tac;
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   884
qed "real_of_nat_less_iff";
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   885
9069
e8d530582061 a big tidy-up
paulson
parents: 9053
diff changeset
   886
AddIffs [real_of_nat_less_iff];
7334
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   887
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   888
Goal "inj real_of_nat";
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   889
by (rtac injI 1);
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   890
by (auto_tac (claset() addSIs [inj_real_of_posnat RS injD],
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   891
	      simpset() addsimps [real_of_nat_def,real_add_right_cancel]));
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   892
qed "inj_real_of_nat";
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   893
9043
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
   894
Goalw [real_of_nat_def] "0 <= real_of_nat n";
7334
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   895
by (res_inst_tac [("C","1r")] real_le_add_right_cancel 1);
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   896
by (asm_full_simp_tac (simpset() addsimps [real_add_assoc]) 1);
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   897
qed "real_of_nat_ge_zero";
9069
e8d530582061 a big tidy-up
paulson
parents: 9053
diff changeset
   898
AddIffs [real_of_nat_ge_zero];
7334
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   899
9069
e8d530582061 a big tidy-up
paulson
parents: 9053
diff changeset
   900
Goal "real_of_nat (m * n) = real_of_nat m * real_of_nat n";
e8d530582061 a big tidy-up
paulson
parents: 9053
diff changeset
   901
by (induct_tac "m" 1);
7334
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   902
by (auto_tac (claset(),
9069
e8d530582061 a big tidy-up
paulson
parents: 9053
diff changeset
   903
	      simpset() addsimps [real_of_nat_add,
7334
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   904
				  real_add_mult_distrib, real_add_commute]));
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   905
qed "real_of_nat_mult";
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   906
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   907
Goal "(real_of_nat n = real_of_nat m) = (n = m)";
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   908
by (auto_tac (claset() addDs [inj_real_of_nat RS injD],
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   909
              simpset()));
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   910
qed "real_of_nat_eq_cancel";
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   911
9069
e8d530582061 a big tidy-up
paulson
parents: 9053
diff changeset
   912
Goal "n <= m --> real_of_nat (m - n) = real_of_nat m + (-real_of_nat n)";
e8d530582061 a big tidy-up
paulson
parents: 9053
diff changeset
   913
by (induct_tac "m" 1);
7334
a90fc1e5fb19 Real/Real.thy main entry point;
wenzelm
parents:
diff changeset
   914
by (auto_tac (claset(),
8867
06dcd62f65ad deleted a lot of obsolete arithmetic lemmas
paulson
parents: 8856
diff changeset
   915
	      simpset() addsimps [Suc_diff_le, le_Suc_eq, real_of_nat_Suc, 
06dcd62f65ad deleted a lot of obsolete arithmetic lemmas
paulson
parents: 8856
diff changeset
   916
				  real_of_nat_zero] @ real_add_ac));
06dcd62f65ad deleted a lot of obsolete arithmetic lemmas
paulson
parents: 8856
diff changeset
   917
qed_spec_mp "real_of_nat_minus";
9013
9dd0274f76af Updated files to remove 0r and 1r from theorems in descendant theories
fleuriot
parents: 8867
diff changeset
   918
9dd0274f76af Updated files to remove 0r and 1r from theorems in descendant theories
fleuriot
parents: 8867
diff changeset
   919
(* 05/00 *)
9069
e8d530582061 a big tidy-up
paulson
parents: 9053
diff changeset
   920
Goal "n < m ==> real_of_nat (m - n) = \
e8d530582061 a big tidy-up
paulson
parents: 9053
diff changeset
   921
\     real_of_nat m + -real_of_nat n";
9013
9dd0274f76af Updated files to remove 0r and 1r from theorems in descendant theories
fleuriot
parents: 8867
diff changeset
   922
by (auto_tac (claset() addIs [real_of_nat_minus],simpset()));
9dd0274f76af Updated files to remove 0r and 1r from theorems in descendant theories
fleuriot
parents: 8867
diff changeset
   923
qed "real_of_nat_minus2";
9dd0274f76af Updated files to remove 0r and 1r from theorems in descendant theories
fleuriot
parents: 8867
diff changeset
   924
9069
e8d530582061 a big tidy-up
paulson
parents: 9053
diff changeset
   925
Goalw [real_diff_def]
e8d530582061 a big tidy-up
paulson
parents: 9053
diff changeset
   926
     "n < m ==> real_of_nat (m - n) = real_of_nat m - real_of_nat n";
9013
9dd0274f76af Updated files to remove 0r and 1r from theorems in descendant theories
fleuriot
parents: 8867
diff changeset
   927
by (etac real_of_nat_minus2 1);
9dd0274f76af Updated files to remove 0r and 1r from theorems in descendant theories
fleuriot
parents: 8867
diff changeset
   928
qed "real_of_nat_diff";
9dd0274f76af Updated files to remove 0r and 1r from theorems in descendant theories
fleuriot
parents: 8867
diff changeset
   929
9069
e8d530582061 a big tidy-up
paulson
parents: 9053
diff changeset
   930
Goalw [real_diff_def]
e8d530582061 a big tidy-up
paulson
parents: 9053
diff changeset
   931
     "n <= m ==> real_of_nat (m - n) = real_of_nat m - real_of_nat n";
9013
9dd0274f76af Updated files to remove 0r and 1r from theorems in descendant theories
fleuriot
parents: 8867
diff changeset
   932
by (etac real_of_nat_minus 1);
9dd0274f76af Updated files to remove 0r and 1r from theorems in descendant theories
fleuriot
parents: 8867
diff changeset
   933
qed "real_of_nat_diff2";
9dd0274f76af Updated files to remove 0r and 1r from theorems in descendant theories
fleuriot
parents: 8867
diff changeset
   934
9069
e8d530582061 a big tidy-up
paulson
parents: 9053
diff changeset
   935
Goal "(real_of_nat n = 0) = (n = 0)";
e8d530582061 a big tidy-up
paulson
parents: 9053
diff changeset
   936
by (auto_tac (claset() addIs [inj_real_of_nat RS injD], simpset()));
e8d530582061 a big tidy-up
paulson
parents: 9053
diff changeset
   937
qed "real_of_nat_zero_iff";
e8d530582061 a big tidy-up
paulson
parents: 9053
diff changeset
   938
AddIffs [real_of_nat_zero_iff];
9013
9dd0274f76af Updated files to remove 0r and 1r from theorems in descendant theories
fleuriot
parents: 8867
diff changeset
   939
9069
e8d530582061 a big tidy-up
paulson
parents: 9053
diff changeset
   940
Goal "neg z ==> real_of_nat (nat z) = 0";
e8d530582061 a big tidy-up
paulson
parents: 9053
diff changeset
   941
by (asm_simp_tac (simpset() addsimps [neg_nat, real_of_nat_zero]) 1);
e8d530582061 a big tidy-up
paulson
parents: 9053
diff changeset
   942
qed "real_of_nat_neg_int";
e8d530582061 a big tidy-up
paulson
parents: 9053
diff changeset
   943
Addsimps [real_of_nat_neg_int];
e8d530582061 a big tidy-up
paulson
parents: 9053
diff changeset
   944