src/HOL/Real/RealDef.thy
author paulson
Tue, 02 Dec 2003 11:48:15 +0100
changeset 14270 342451d763f9
parent 14269 502a7c95de73
child 14329 ff3210fe968f
permissions -rw-r--r--
More re-organising of numerical theorems
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
5588
a3ab526bb891 Revised version with Abelian group simprocs
paulson
parents:
diff changeset
     1
(*  Title       : Real/RealDef.thy
7219
4e3f386c2e37 inserted Id: lines
paulson
parents: 7127
diff changeset
     2
    ID          : $Id$
5588
a3ab526bb891 Revised version with Abelian group simprocs
paulson
parents:
diff changeset
     3
    Author      : Jacques D. Fleuriot
a3ab526bb891 Revised version with Abelian group simprocs
paulson
parents:
diff changeset
     4
    Copyright   : 1998  University of Cambridge
a3ab526bb891 Revised version with Abelian group simprocs
paulson
parents:
diff changeset
     5
    Description : The reals
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
     6
*)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
     7
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
     8
theory RealDef = PReal:
5588
a3ab526bb891 Revised version with Abelian group simprocs
paulson
parents:
diff changeset
     9
14270
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
    10
(*MOVE TO THEORY PREAL*)
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
    11
instance preal :: order
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
    12
proof qed
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
    13
 (assumption |
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
    14
  rule preal_le_refl preal_le_trans preal_le_anti_sym preal_less_le)+
10752
c4f1bf2acf4c tidying, and separation of HOL-Hyperreal from HOL-Real
paulson
parents: 10648
diff changeset
    15
14270
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
    16
instance preal :: order
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
    17
  by (intro_classes,
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
    18
      (assumption | 
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
    19
       rule preal_le_refl preal_le_trans preal_le_anti_sym preal_less_le)+)
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
    20
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
    21
lemma preal_le_linear: "x <= y | y <= (x::preal)"
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
    22
apply (insert preal_linear [of x y]) 
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
    23
apply (auto simp add: order_less_le) 
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
    24
done
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
    25
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
    26
instance preal :: linorder
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
    27
  by (intro_classes, rule preal_le_linear)
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
    28
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
    29
5588
a3ab526bb891 Revised version with Abelian group simprocs
paulson
parents:
diff changeset
    30
constdefs
a3ab526bb891 Revised version with Abelian group simprocs
paulson
parents:
diff changeset
    31
  realrel   ::  "((preal * preal) * (preal * preal)) set"
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
    32
  "realrel == {p. \<exists>x1 y1 x2 y2. p = ((x1,y1),(x2,y2)) & x1+y2 = x2+y1}"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
    33
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
    34
typedef (REAL)  real = "UNIV//realrel"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
    35
  by (auto simp add: quotient_def)
5588
a3ab526bb891 Revised version with Abelian group simprocs
paulson
parents:
diff changeset
    36
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
    37
instance real :: ord ..
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
    38
instance real :: zero ..
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
    39
instance real :: one ..
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
    40
instance real :: plus ..
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
    41
instance real :: times ..
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
    42
instance real :: minus ..
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
    43
instance real :: inverse ..
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
    44
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
    45
consts
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
    46
   (*Overloaded constants denoting the Nat and Real subsets of enclosing
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
    47
     types such as hypreal and complex*)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
    48
   Nats  :: "'a set"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
    49
   Reals :: "'a set"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
    50
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
    51
   (*overloaded constant for injecting other types into "real"*)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
    52
   real :: "'a => real"
5588
a3ab526bb891 Revised version with Abelian group simprocs
paulson
parents:
diff changeset
    53
a3ab526bb891 Revised version with Abelian group simprocs
paulson
parents:
diff changeset
    54
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
    55
defs (overloaded)
5588
a3ab526bb891 Revised version with Abelian group simprocs
paulson
parents:
diff changeset
    56
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
    57
  real_zero_def:
12018
ec054019c910 Numerals and simprocs for types real and hypreal. The abstract
paulson
parents: 11713
diff changeset
    58
  "0 == Abs_REAL(realrel``{(preal_of_prat(prat_of_pnat 1),
ec054019c910 Numerals and simprocs for types real and hypreal. The abstract
paulson
parents: 11713
diff changeset
    59
			    preal_of_prat(prat_of_pnat 1))})"
ec054019c910 Numerals and simprocs for types real and hypreal. The abstract
paulson
parents: 11713
diff changeset
    60
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
    61
  real_one_def:
12018
ec054019c910 Numerals and simprocs for types real and hypreal. The abstract
paulson
parents: 11713
diff changeset
    62
  "1 == Abs_REAL(realrel``
ec054019c910 Numerals and simprocs for types real and hypreal. The abstract
paulson
parents: 11713
diff changeset
    63
               {(preal_of_prat(prat_of_pnat 1) + preal_of_prat(prat_of_pnat 1),
ec054019c910 Numerals and simprocs for types real and hypreal. The abstract
paulson
parents: 11713
diff changeset
    64
		 preal_of_prat(prat_of_pnat 1))})"
5588
a3ab526bb891 Revised version with Abelian group simprocs
paulson
parents:
diff changeset
    65
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
    66
  real_minus_def:
10919
144ede948e58 renamings: real_of_nat, real_of_int -> (overloaded) real
paulson
parents: 10834
diff changeset
    67
  "- R ==  Abs_REAL(UN (x,y):Rep_REAL(R). realrel``{(y,x)})"
10606
e3229a37d53f converted rinv to inverse;
bauerg
parents: 9391
diff changeset
    68
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
    69
  real_diff_def:
10606
e3229a37d53f converted rinv to inverse;
bauerg
parents: 9391
diff changeset
    70
  "R - (S::real) == R + - S"
5588
a3ab526bb891 Revised version with Abelian group simprocs
paulson
parents:
diff changeset
    71
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
    72
  real_inverse_def:
11713
883d559b0b8c sane numerals (stage 3): provide generic "1" on all number types;
wenzelm
parents: 11701
diff changeset
    73
  "inverse (R::real) == (SOME S. (R = 0 & S = 0) | S * R = 1)"
5588
a3ab526bb891 Revised version with Abelian group simprocs
paulson
parents:
diff changeset
    74
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
    75
  real_divide_def:
10606
e3229a37d53f converted rinv to inverse;
bauerg
parents: 9391
diff changeset
    76
  "R / (S::real) == R * inverse S"
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
    77
5588
a3ab526bb891 Revised version with Abelian group simprocs
paulson
parents:
diff changeset
    78
constdefs
a3ab526bb891 Revised version with Abelian group simprocs
paulson
parents:
diff changeset
    79
12018
ec054019c910 Numerals and simprocs for types real and hypreal. The abstract
paulson
parents: 11713
diff changeset
    80
  (** these don't use the overloaded "real" function: users don't see them **)
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
    81
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
    82
  real_of_preal :: "preal => real"
7077
60b098bb8b8a heavily revised by Jacques: coercions have alphabetic names;
paulson
parents: 5787
diff changeset
    83
  "real_of_preal m     ==
12018
ec054019c910 Numerals and simprocs for types real and hypreal. The abstract
paulson
parents: 11713
diff changeset
    84
           Abs_REAL(realrel``{(m + preal_of_prat(prat_of_pnat 1),
ec054019c910 Numerals and simprocs for types real and hypreal. The abstract
paulson
parents: 11713
diff changeset
    85
                               preal_of_prat(prat_of_pnat 1))})"
5588
a3ab526bb891 Revised version with Abelian group simprocs
paulson
parents:
diff changeset
    86
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
    87
  real_of_posnat :: "nat => real"
7077
60b098bb8b8a heavily revised by Jacques: coercions have alphabetic names;
paulson
parents: 5787
diff changeset
    88
  "real_of_posnat n == real_of_preal(preal_of_prat(prat_of_pnat(pnat_of_nat n)))"
60b098bb8b8a heavily revised by Jacques: coercions have alphabetic names;
paulson
parents: 5787
diff changeset
    89
5588
a3ab526bb891 Revised version with Abelian group simprocs
paulson
parents:
diff changeset
    90
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
    91
defs (overloaded)
5588
a3ab526bb891 Revised version with Abelian group simprocs
paulson
parents:
diff changeset
    92
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
    93
  real_of_nat_def:   "real n == real_of_posnat n + (- 1)"
10919
144ede948e58 renamings: real_of_nat, real_of_int -> (overloaded) real
paulson
parents: 10834
diff changeset
    94
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
    95
  real_add_def:
10919
144ede948e58 renamings: real_of_nat, real_of_int -> (overloaded) real
paulson
parents: 10834
diff changeset
    96
  "P+Q == Abs_REAL(UN p1:Rep_REAL(P). UN p2:Rep_REAL(Q).
10834
a7897aebbffc *** empty log message ***
nipkow
parents: 10797
diff changeset
    97
                   (%(x1,y1). (%(x2,y2). realrel``{(x1+x2, y1+y2)}) p2) p1)"
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
    98
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
    99
  real_mult_def:
10919
144ede948e58 renamings: real_of_nat, real_of_int -> (overloaded) real
paulson
parents: 10834
diff changeset
   100
  "P*Q == Abs_REAL(UN p1:Rep_REAL(P). UN p2:Rep_REAL(Q).
10834
a7897aebbffc *** empty log message ***
nipkow
parents: 10797
diff changeset
   101
                   (%(x1,y1). (%(x2,y2). realrel``{(x1*x2+y1*y2,x1*y2+x2*y1)})
10752
c4f1bf2acf4c tidying, and separation of HOL-Hyperreal from HOL-Real
paulson
parents: 10648
diff changeset
   102
		   p2) p1)"
5588
a3ab526bb891 Revised version with Abelian group simprocs
paulson
parents:
diff changeset
   103
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   104
  real_less_def:
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   105
  "P<Q == \<exists>x1 y1 x2 y2. x1 + y2 < x2 + y1 &
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   106
                            (x1,y1):Rep_REAL(P) & (x2,y2):Rep_REAL(Q)"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   107
  real_le_def:
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   108
  "P \<le> (Q::real) == ~(Q < P)"
5588
a3ab526bb891 Revised version with Abelian group simprocs
paulson
parents:
diff changeset
   109
12114
a8e860c86252 eliminated old "symbols" syntax, use "xsymbols" instead;
wenzelm
parents: 12018
diff changeset
   110
syntax (xsymbols)
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   111
  Reals     :: "'a set"                   ("\<real>")
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   112
  Nats      :: "'a set"                   ("\<nat>")
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   113
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   114
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   115
(*** Proving that realrel is an equivalence relation ***)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   116
14270
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   117
lemma preal_trans_lemma:
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   118
     "[| (x1::preal) + y2 = x2 + y1; x2 + y3 = x3 + y2 |]
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   119
      ==> x1 + y3 = x3 + y1"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   120
apply (rule_tac C = y2 in preal_add_right_cancel)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   121
apply (rotate_tac 1, drule sym)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   122
apply (simp add: preal_add_ac)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   123
apply (rule preal_add_left_commute [THEN subst])
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   124
apply (rule_tac x1 = x1 in preal_add_assoc [THEN subst])
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   125
apply (simp add: preal_add_ac)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   126
done
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   127
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   128
lemma realrel_iff [simp]: "(((x1,y1),(x2,y2)): realrel) = (x1 + y2 = x2 + y1)"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   129
by (unfold realrel_def, blast)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   130
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   131
lemma realrel_refl: "(x,x): realrel"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   132
apply (case_tac "x")
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   133
apply (simp add: realrel_def)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   134
done
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   135
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   136
lemma equiv_realrel: "equiv UNIV realrel"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   137
apply (unfold equiv_def refl_def sym_def trans_def realrel_def)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   138
apply (fast elim!: sym preal_trans_lemma)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   139
done
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   140
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   141
(* (realrel `` {x} = realrel `` {y}) = ((x,y) : realrel) *)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   142
lemmas equiv_realrel_iff = 
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   143
       eq_equiv_class_iff [OF equiv_realrel UNIV_I UNIV_I]
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   144
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   145
declare equiv_realrel_iff [simp]
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   146
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   147
lemma realrel_in_real [simp]: "realrel``{(x,y)}: REAL"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   148
by (unfold REAL_def realrel_def quotient_def, blast)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   149
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   150
lemma inj_on_Abs_REAL: "inj_on Abs_REAL REAL"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   151
apply (rule inj_on_inverseI)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   152
apply (erule Abs_REAL_inverse)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   153
done
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   154
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   155
declare inj_on_Abs_REAL [THEN inj_on_iff, simp]
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   156
declare Abs_REAL_inverse [simp]
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   157
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   158
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   159
lemmas eq_realrelD = equiv_realrel [THEN [2] eq_equiv_class]
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   160
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   161
lemma inj_Rep_REAL: "inj Rep_REAL"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   162
apply (rule inj_on_inverseI)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   163
apply (rule Rep_REAL_inverse)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   164
done
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   165
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   166
(** real_of_preal: the injection from preal to real **)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   167
lemma inj_real_of_preal: "inj(real_of_preal)"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   168
apply (rule inj_onI)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   169
apply (unfold real_of_preal_def)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   170
apply (drule inj_on_Abs_REAL [THEN inj_onD])
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   171
apply (rule realrel_in_real)+
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   172
apply (drule eq_equiv_class)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   173
apply (rule equiv_realrel, blast)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   174
apply (simp add: realrel_def)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   175
done
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   176
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   177
lemma eq_Abs_REAL: 
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   178
    "(!!x y. z = Abs_REAL(realrel``{(x,y)}) ==> P) ==> P"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   179
apply (rule_tac x1 = z in Rep_REAL [unfolded REAL_def, THEN quotientE])
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   180
apply (drule_tac f = Abs_REAL in arg_cong)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   181
apply (case_tac "x")
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   182
apply (simp add: Rep_REAL_inverse)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   183
done
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   184
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   185
(**** real_minus: additive inverse on real ****)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   186
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   187
lemma real_minus_congruent:
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   188
  "congruent realrel (%p. (%(x,y). realrel``{(y,x)}) p)"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   189
apply (unfold congruent_def, clarify)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   190
apply (simp add: preal_add_commute)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   191
done
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   192
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   193
lemma real_minus:
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   194
      "- (Abs_REAL(realrel``{(x,y)})) = Abs_REAL(realrel `` {(y,x)})"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   195
apply (unfold real_minus_def)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   196
apply (rule_tac f = Abs_REAL in arg_cong)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   197
apply (simp add: realrel_in_real [THEN Abs_REAL_inverse] 
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   198
            UN_equiv_class [OF equiv_realrel real_minus_congruent])
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   199
done
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   200
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   201
lemma real_minus_minus: "- (- z) = (z::real)"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   202
apply (rule_tac z = z in eq_Abs_REAL)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   203
apply (simp add: real_minus)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   204
done
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   205
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   206
declare real_minus_minus [simp]
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   207
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   208
lemma inj_real_minus: "inj(%r::real. -r)"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   209
apply (rule inj_onI)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   210
apply (drule_tac f = uminus in arg_cong)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   211
apply (simp add: real_minus_minus)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   212
done
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   213
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   214
lemma real_minus_zero: "- 0 = (0::real)"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   215
apply (unfold real_zero_def)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   216
apply (simp add: real_minus)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   217
done
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   218
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   219
declare real_minus_zero [simp]
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   220
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   221
lemma real_minus_zero_iff: "(-x = 0) = (x = (0::real))"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   222
apply (rule_tac z = x in eq_Abs_REAL)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   223
apply (auto simp add: real_zero_def real_minus preal_add_ac)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   224
done
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   225
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   226
declare real_minus_zero_iff [simp]
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   227
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   228
(*** Congruence property for addition ***)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   229
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   230
lemma real_add_congruent2_lemma:
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   231
     "[|a + ba = aa + b; ab + bc = ac + bb|]
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   232
      ==> a + ab + (ba + bc) = aa + ac + (b + (bb::preal))"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   233
apply (simp add: preal_add_assoc) 
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   234
apply (rule preal_add_left_commute [of ab, THEN ssubst])
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   235
apply (simp add: preal_add_assoc [symmetric])
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   236
apply (simp add: preal_add_ac)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   237
done
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   238
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   239
lemma real_add:
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   240
  "Abs_REAL(realrel``{(x1,y1)}) + Abs_REAL(realrel``{(x2,y2)}) =
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   241
   Abs_REAL(realrel``{(x1+x2, y1+y2)})"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   242
apply (simp add: real_add_def UN_UN_split_split_eq)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   243
apply (subst equiv_realrel [THEN UN_equiv_class2])
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   244
apply (auto simp add: congruent2_def)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   245
apply (blast intro: real_add_congruent2_lemma) 
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   246
done
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   247
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   248
lemma real_add_commute: "(z::real) + w = w + z"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   249
apply (rule_tac z = z in eq_Abs_REAL)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   250
apply (rule_tac z = w in eq_Abs_REAL)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   251
apply (simp add: preal_add_ac real_add)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   252
done
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   253
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   254
lemma real_add_assoc: "((z1::real) + z2) + z3 = z1 + (z2 + z3)"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   255
apply (rule_tac z = z1 in eq_Abs_REAL)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   256
apply (rule_tac z = z2 in eq_Abs_REAL)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   257
apply (rule_tac z = z3 in eq_Abs_REAL)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   258
apply (simp add: real_add preal_add_assoc)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   259
done
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   260
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   261
(*For AC rewriting*)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   262
lemma real_add_left_commute: "(x::real)+(y+z)=y+(x+z)"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   263
  apply (rule mk_left_commute [of "op +"])
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   264
  apply (rule real_add_assoc)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   265
  apply (rule real_add_commute)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   266
  done
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   267
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   268
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   269
(* real addition is an AC operator *)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   270
lemmas real_add_ac = real_add_assoc real_add_commute real_add_left_commute
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   271
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   272
lemma real_add_zero_left: "(0::real) + z = z"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   273
apply (unfold real_of_preal_def real_zero_def)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   274
apply (rule_tac z = z in eq_Abs_REAL)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   275
apply (simp add: real_add preal_add_ac)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   276
done
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   277
declare real_add_zero_left [simp]
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   278
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   279
lemma real_add_zero_right: "z + (0::real) = z"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   280
by (simp add: real_add_commute)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   281
declare real_add_zero_right [simp]
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   282
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   283
instance real :: plus_ac0
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   284
  by (intro_classes,
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   285
      (assumption | 
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   286
       rule real_add_commute real_add_assoc real_add_zero_left)+)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   287
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   288
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   289
lemma real_add_minus: "z + (-z) = (0::real)"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   290
apply (unfold real_zero_def)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   291
apply (rule_tac z = z in eq_Abs_REAL)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   292
apply (simp add: real_minus real_add preal_add_commute)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   293
done
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   294
declare real_add_minus [simp]
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   295
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   296
lemma real_add_minus_left: "(-z) + z = (0::real)"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   297
by (simp add: real_add_commute)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   298
declare real_add_minus_left [simp]
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   299
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   300
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   301
(*** Congruence property for multiplication ***)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   302
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   303
lemma real_mult_congruent2_lemma: "!!(x1::preal). [| x1 + y2 = x2 + y1 |] ==>
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   304
          x * x1 + y * y1 + (x * y2 + x2 * y) =
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   305
          x * x2 + y * y2 + (x * y1 + x1 * y)"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   306
apply (simp add: preal_add_left_commute preal_add_assoc [symmetric] preal_add_mult_distrib2 [symmetric])
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   307
apply (rule preal_mult_commute [THEN subst])
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   308
apply (rule_tac y1 = x2 in preal_mult_commute [THEN subst])
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   309
apply (simp add: preal_add_assoc preal_add_mult_distrib2 [symmetric])
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   310
apply (simp add: preal_add_commute)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   311
done
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   312
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   313
lemma real_mult_congruent2:
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   314
    "congruent2 realrel (%p1 p2.
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   315
          (%(x1,y1). (%(x2,y2). realrel``{(x1*x2 + y1*y2, x1*y2+x2*y1)}) p2) p1)"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   316
apply (rule equiv_realrel [THEN congruent2_commuteI], clarify)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   317
apply (unfold split_def)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   318
apply (simp add: preal_mult_commute preal_add_commute)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   319
apply (auto simp add: real_mult_congruent2_lemma)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   320
done
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   321
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   322
lemma real_mult:
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   323
   "Abs_REAL((realrel``{(x1,y1)})) * Abs_REAL((realrel``{(x2,y2)})) =
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   324
    Abs_REAL(realrel `` {(x1*x2+y1*y2,x1*y2+x2*y1)})"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   325
apply (unfold real_mult_def)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   326
apply (simp add: equiv_realrel [THEN UN_equiv_class2] real_mult_congruent2)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   327
done
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   328
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   329
lemma real_mult_commute: "(z::real) * w = w * z"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   330
apply (rule_tac z = z in eq_Abs_REAL)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   331
apply (rule_tac z = w in eq_Abs_REAL)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   332
apply (simp add: real_mult preal_add_ac preal_mult_ac)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   333
done
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   334
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   335
lemma real_mult_assoc: "((z1::real) * z2) * z3 = z1 * (z2 * z3)"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   336
apply (rule_tac z = z1 in eq_Abs_REAL)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   337
apply (rule_tac z = z2 in eq_Abs_REAL)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   338
apply (rule_tac z = z3 in eq_Abs_REAL)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   339
apply (simp add: preal_add_mult_distrib2 real_mult preal_add_ac preal_mult_ac)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   340
done
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   341
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   342
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   343
(*For AC rewriting*)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   344
lemma real_mult_left_commute: "(x::real)*(y*z)=y*(x*z)"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   345
  apply (rule mk_left_commute [of "op *"])
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   346
  apply (rule real_mult_assoc)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   347
  apply (rule real_mult_commute)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   348
  done
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   349
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   350
(* real multiplication is an AC operator *)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   351
lemmas real_mult_ac = real_mult_assoc real_mult_commute real_mult_left_commute
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   352
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   353
lemma real_mult_1: "(1::real) * z = z"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   354
apply (unfold real_one_def pnat_one_def)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   355
apply (rule_tac z = z in eq_Abs_REAL)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   356
apply (simp add: real_mult preal_add_mult_distrib2 preal_mult_1_right preal_mult_ac preal_add_ac)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   357
done
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   358
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   359
declare real_mult_1 [simp]
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   360
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   361
lemma real_mult_1_right: "z * (1::real) = z"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   362
by (simp add: real_mult_commute)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   363
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   364
declare real_mult_1_right [simp]
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   365
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   366
lemma real_mult_0: "0 * z = (0::real)"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   367
apply (unfold real_zero_def pnat_one_def)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   368
apply (rule_tac z = z in eq_Abs_REAL)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   369
apply (simp add: real_mult preal_add_mult_distrib2 preal_mult_1_right preal_mult_ac preal_add_ac)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   370
done
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   371
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   372
lemma real_mult_0_right: "z * 0 = (0::real)"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   373
by (simp add: real_mult_commute real_mult_0)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   374
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   375
declare real_mult_0_right [simp] real_mult_0 [simp]
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   376
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   377
lemma real_mult_minus_eq1: "(-x) * (y::real) = -(x * y)"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   378
apply (rule_tac z = x in eq_Abs_REAL)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   379
apply (rule_tac z = y in eq_Abs_REAL)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   380
apply (auto simp add: real_minus real_mult preal_mult_ac preal_add_ac)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   381
done
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   382
declare real_mult_minus_eq1 [simp]
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   383
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   384
lemmas real_minus_mult_eq1 = real_mult_minus_eq1 [symmetric, standard]
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   385
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   386
lemma real_mult_minus_eq2: "x * (- y :: real) = -(x * y)"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   387
by (simp add: real_mult_commute [of x])
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   388
declare real_mult_minus_eq2 [simp]
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   389
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   390
lemmas real_minus_mult_eq2 = real_mult_minus_eq2 [symmetric, standard]
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   391
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   392
lemma real_mult_minus_1: "(- (1::real)) * z = -z"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   393
by simp
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   394
declare real_mult_minus_1 [simp]
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   395
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   396
lemma real_mult_minus_1_right: "z * (- (1::real)) = -z"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   397
by (subst real_mult_commute, simp)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   398
declare real_mult_minus_1_right [simp]
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   399
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   400
lemma real_minus_mult_cancel: "(-x) * (-y) = x * (y::real)"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   401
by simp
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   402
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   403
declare real_minus_mult_cancel [simp]
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   404
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   405
lemma real_minus_mult_commute: "(-x) * y = x * (- y :: real)"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   406
by simp
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   407
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   408
(** Lemmas **)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   409
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   410
lemma real_add_assoc_cong: "(z::real) + v = z' + v' ==> z + (v + w) = z' + (v' + w)"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   411
by (simp add: real_add_assoc [symmetric])
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   412
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   413
lemma real_add_mult_distrib: "((z1::real) + z2) * w = (z1 * w) + (z2 * w)"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   414
apply (rule_tac z = z1 in eq_Abs_REAL)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   415
apply (rule_tac z = z2 in eq_Abs_REAL)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   416
apply (rule_tac z = w in eq_Abs_REAL)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   417
apply (simp add: preal_add_mult_distrib2 real_add real_mult preal_add_ac preal_mult_ac)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   418
done
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   419
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   420
lemma real_add_mult_distrib2: "(w::real) * (z1 + z2) = (w * z1) + (w * z2)"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   421
by (simp add: real_mult_commute [of w] real_add_mult_distrib)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   422
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   423
lemma real_diff_mult_distrib: "((z1::real) - z2) * w = (z1 * w) - (z2 * w)"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   424
apply (unfold real_diff_def)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   425
apply (simp add: real_add_mult_distrib)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   426
done
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   427
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   428
lemma real_diff_mult_distrib2: "(w::real) * (z1 - z2) = (w * z1) - (w * z2)"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   429
by (simp add: real_mult_commute [of w] real_diff_mult_distrib)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   430
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   431
(*** one and zero are distinct ***)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   432
lemma real_zero_not_eq_one: "0 ~= (1::real)"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   433
apply (unfold real_zero_def real_one_def)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   434
apply (auto simp add: preal_self_less_add_left [THEN preal_not_refl2])
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   435
done
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   436
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   437
(*** existence of inverse ***)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   438
(** lemma -- alternative definition of 0 **)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   439
lemma real_zero_iff: "0 = Abs_REAL (realrel `` {(x, x)})"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   440
apply (unfold real_zero_def)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   441
apply (auto simp add: preal_add_commute)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   442
done
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   443
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   444
lemma real_mult_inv_right_ex:
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   445
          "!!(x::real). x ~= 0 ==> \<exists>y. x*y = (1::real)"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   446
apply (unfold real_zero_def real_one_def)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   447
apply (rule_tac z = x in eq_Abs_REAL)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   448
apply (cut_tac x = xa and y = y in linorder_less_linear)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   449
apply (auto dest!: preal_less_add_left_Ex simp add: real_zero_iff [symmetric])
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   450
apply (rule_tac x = "Abs_REAL (realrel `` { (preal_of_prat (prat_of_pnat 1), pinv (D) + preal_of_prat (prat_of_pnat 1))}) " in exI)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   451
apply (rule_tac [2] x = "Abs_REAL (realrel `` { (pinv (D) + preal_of_prat (prat_of_pnat 1), preal_of_prat (prat_of_pnat 1))}) " in exI)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   452
apply (auto simp add: real_mult pnat_one_def preal_mult_1_right preal_add_mult_distrib2 preal_add_mult_distrib preal_mult_1 preal_mult_inv_right preal_add_ac preal_mult_ac)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   453
done
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   454
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   455
lemma real_mult_inv_left_ex: "x ~= 0 ==> \<exists>y. y*x = (1::real)"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   456
apply (drule real_mult_inv_right_ex)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   457
apply (auto simp add: real_mult_commute)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   458
done
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   459
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   460
lemma real_mult_inv_left: "x ~= 0 ==> inverse(x)*x = (1::real)"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   461
apply (unfold real_inverse_def)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   462
apply (frule real_mult_inv_left_ex, safe)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   463
apply (rule someI2, auto)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   464
done
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   465
declare real_mult_inv_left [simp]
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   466
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   467
lemma real_mult_inv_right: "x ~= 0 ==> x*inverse(x) = (1::real)"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   468
apply (subst real_mult_commute)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   469
apply (auto simp add: real_mult_inv_left)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   470
done
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   471
declare real_mult_inv_right [simp]
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   472
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   473
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   474
(*---------------------------------------------------------
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   475
     Theorems for ordering
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   476
 --------------------------------------------------------*)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   477
(* prove introduction and elimination rules for real_less *)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   478
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   479
(* real_less is a strong order i.e. nonreflexive and transitive *)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   480
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   481
(*** lemmas ***)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   482
lemma preal_lemma_eq_rev_sum: "!!(x::preal). [| x = y; x1 = y1 |] ==> x + y1 = x1 + y"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   483
by (simp add: preal_add_commute)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   484
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   485
lemma preal_add_left_commute_cancel: "!!(b::preal). x + (b + y) = x1 + (b + y1) ==> x + y = x1 + y1"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   486
by (simp add: preal_add_ac)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   487
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   488
lemma preal_lemma_for_not_refl: "!!(x::preal). [| x + y2a = x2a + y;
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   489
                       x + y2b = x2b + y |]
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   490
                    ==> x2a + y2b = x2b + y2a"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   491
apply (drule preal_lemma_eq_rev_sum, assumption)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   492
apply (erule_tac V = "x + y2b = x2b + y" in thin_rl)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   493
apply (simp add: preal_add_ac)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   494
apply (drule preal_add_left_commute_cancel)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   495
apply (simp add: preal_add_ac)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   496
done
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   497
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   498
lemma real_less_not_refl: "~ (R::real) < R"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   499
apply (rule_tac z = R in eq_Abs_REAL)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   500
apply (auto simp add: real_less_def)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   501
apply (drule preal_lemma_for_not_refl, assumption, auto)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   502
done
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   503
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   504
(*** y < y ==> P ***)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   505
lemmas real_less_irrefl = real_less_not_refl [THEN notE, standard]
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   506
declare real_less_irrefl [elim!]
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   507
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   508
lemma real_not_refl2: "!!(x::real). x < y ==> x ~= y"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   509
by (auto simp add: real_less_not_refl)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   510
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   511
(* lemma re-arranging and eliminating terms *)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   512
lemma preal_lemma_trans: "!! (a::preal). [| a + b = c + d;
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   513
             x2b + d + (c + y2e) < a + y2b + (x2e + b) |]
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   514
          ==> x2b + y2e < x2e + y2b"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   515
apply (simp add: preal_add_ac)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   516
apply (rule_tac C = "c+d" in preal_add_left_less_cancel)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   517
apply (simp add: preal_add_assoc [symmetric])
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   518
done
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   519
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   520
(** A MESS!  heavy re-writing involved*)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   521
lemma real_less_trans: "!!(R1::real). [| R1 < R2; R2 < R3 |] ==> R1 < R3"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   522
apply (rule_tac z = R1 in eq_Abs_REAL)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   523
apply (rule_tac z = R2 in eq_Abs_REAL)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   524
apply (rule_tac z = R3 in eq_Abs_REAL)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   525
apply (auto simp add: real_less_def)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   526
apply (rule exI)+
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   527
apply (rule conjI, rule_tac [2] conjI)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   528
 prefer 2 apply blast 
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   529
 prefer 2 apply blast 
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   530
apply (drule preal_lemma_for_not_refl, assumption)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   531
apply (blast dest: preal_add_less_mono intro: preal_lemma_trans)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   532
done
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   533
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   534
lemma real_less_not_sym: "!! (R1::real). R1 < R2 ==> ~ (R2 < R1)"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   535
apply (rule notI)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   536
apply (drule real_less_trans, assumption)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   537
apply (simp add: real_less_not_refl)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   538
done
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   539
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   540
(* [| x < y;  ~P ==> y < x |] ==> P *)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   541
lemmas real_less_asym = real_less_not_sym [THEN contrapos_np, standard]
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   542
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   543
lemma real_of_preal_add:
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   544
     "real_of_preal ((z1::preal) + z2) =
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   545
      real_of_preal z1 + real_of_preal z2"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   546
apply (unfold real_of_preal_def)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   547
apply (simp add: real_add preal_add_mult_distrib preal_mult_1 add: preal_add_ac)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   548
done
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   549
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   550
lemma real_of_preal_mult:
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   551
     "real_of_preal ((z1::preal) * z2) =
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   552
      real_of_preal z1* real_of_preal z2"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   553
apply (unfold real_of_preal_def)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   554
apply (simp (no_asm_use) add: real_mult preal_add_mult_distrib2 preal_mult_1 preal_mult_1_right pnat_one_def preal_add_ac preal_mult_ac)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   555
done
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   556
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   557
lemma real_of_preal_ExI:
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   558
      "!!(x::preal). y < x ==>
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   559
       \<exists>m. Abs_REAL (realrel `` {(x,y)}) = real_of_preal m"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   560
apply (unfold real_of_preal_def)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   561
apply (auto dest!: preal_less_add_left_Ex simp add: preal_add_ac)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   562
done
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   563
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   564
lemma real_of_preal_ExD:
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   565
      "!!(x::preal). \<exists>m. Abs_REAL (realrel `` {(x,y)}) =
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   566
                     real_of_preal m ==> y < x"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   567
apply (unfold real_of_preal_def)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   568
apply (auto simp add: preal_add_commute preal_add_assoc)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   569
apply (simp add: preal_add_assoc [symmetric] preal_self_less_add_left)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   570
done
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   571
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   572
lemma real_of_preal_iff: "(\<exists>m. Abs_REAL (realrel `` {(x,y)}) = real_of_preal m) = (y < x)"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   573
by (blast intro!: real_of_preal_ExI real_of_preal_ExD)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   574
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   575
(*** Gleason prop 9-4.4 p 127 ***)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   576
lemma real_of_preal_trichotomy:
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   577
      "\<exists>m. (x::real) = real_of_preal m | x = 0 | x = -(real_of_preal m)"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   578
apply (unfold real_of_preal_def real_zero_def)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   579
apply (rule_tac z = x in eq_Abs_REAL)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   580
apply (auto simp add: real_minus preal_add_ac)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   581
apply (cut_tac x = x and y = y in linorder_less_linear)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   582
apply (auto dest!: preal_less_add_left_Ex simp add: preal_add_assoc [symmetric])
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   583
apply (auto simp add: preal_add_commute)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   584
done
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   585
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   586
lemma real_of_preal_trichotomyE: "!!P. [| !!m. x = real_of_preal m ==> P;
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   587
              x = 0 ==> P;
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   588
              !!m. x = -(real_of_preal m) ==> P |] ==> P"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   589
apply (cut_tac x = x in real_of_preal_trichotomy, auto)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   590
done
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   591
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   592
lemma real_of_preal_lessD:
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   593
      "real_of_preal m1 < real_of_preal m2 ==> m1 < m2"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   594
apply (unfold real_of_preal_def)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   595
apply (auto simp add: real_less_def preal_add_ac)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   596
apply (auto simp add: preal_add_assoc [symmetric])
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   597
apply (auto simp add: preal_add_ac)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   598
done
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   599
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   600
lemma real_of_preal_lessI: "m1 < m2 ==> real_of_preal m1 < real_of_preal m2"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   601
apply (drule preal_less_add_left_Ex)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   602
apply (auto simp add: real_of_preal_add real_of_preal_def real_less_def)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   603
apply (rule exI)+
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   604
apply (rule conjI, rule_tac [2] conjI)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   605
 apply (rule_tac [2] refl)+
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   606
apply (simp add: preal_self_less_add_left del: preal_add_less_iff2)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   607
done
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   608
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   609
lemma real_of_preal_less_iff1: "(real_of_preal m1 < real_of_preal m2) = (m1 < m2)"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   610
by (blast intro: real_of_preal_lessI real_of_preal_lessD)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   611
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   612
declare real_of_preal_less_iff1 [simp]
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   613
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   614
lemma real_of_preal_minus_less_self: "- real_of_preal m < real_of_preal m"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   615
apply (auto simp add: real_of_preal_def real_less_def real_minus)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   616
apply (rule exI)+
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   617
apply (rule conjI, rule_tac [2] conjI)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   618
 apply (rule_tac [2] refl)+
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   619
apply (simp (no_asm_use) add: preal_add_ac)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   620
apply (simp (no_asm_use) add: preal_self_less_add_right preal_add_assoc [symmetric])
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   621
done
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   622
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   623
lemma real_of_preal_minus_less_zero: "- real_of_preal m < 0"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   624
apply (unfold real_zero_def)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   625
apply (auto simp add: real_of_preal_def real_less_def real_minus)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   626
apply (rule exI)+
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   627
apply (rule conjI, rule_tac [2] conjI)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   628
 apply (rule_tac [2] refl)+
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   629
apply (simp (no_asm_use) add: preal_self_less_add_right preal_add_ac)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   630
done
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   631
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   632
lemma real_of_preal_not_minus_gt_zero: "~ 0 < - real_of_preal m"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   633
apply (cut_tac real_of_preal_minus_less_zero)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   634
apply (fast dest: real_less_trans elim: real_less_irrefl)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   635
done
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   636
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   637
lemma real_of_preal_zero_less: "0 < real_of_preal m"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   638
apply (unfold real_zero_def)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   639
apply (auto simp add: real_of_preal_def real_less_def real_minus)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   640
apply (rule exI)+
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   641
apply (rule conjI, rule_tac [2] conjI)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   642
 apply (rule_tac [2] refl)+
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   643
apply (simp (no_asm_use) add: preal_self_less_add_right preal_add_ac)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   644
done
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   645
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   646
lemma real_of_preal_not_less_zero: "~ real_of_preal m < 0"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   647
apply (cut_tac real_of_preal_zero_less)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   648
apply (blast dest: real_less_trans elim: real_less_irrefl)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   649
done
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   650
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   651
lemma real_minus_minus_zero_less: "0 < - (- real_of_preal m)"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   652
by (simp add: real_of_preal_zero_less)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   653
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   654
(* another lemma *)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   655
lemma real_of_preal_sum_zero_less:
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   656
      "0 < real_of_preal m + real_of_preal m1"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   657
apply (unfold real_zero_def)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   658
apply (auto simp add: real_of_preal_def real_less_def real_add)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   659
apply (rule exI)+
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   660
apply (rule conjI, rule_tac [2] conjI)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   661
 apply (rule_tac [2] refl)+
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   662
apply (simp (no_asm_use) add: preal_add_ac)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   663
apply (simp (no_asm_use) add: preal_self_less_add_right preal_add_assoc [symmetric])
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   664
done
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   665
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   666
lemma real_of_preal_minus_less_all: "- real_of_preal m < real_of_preal m1"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   667
apply (auto simp add: real_of_preal_def real_less_def real_minus)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   668
apply (rule exI)+
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   669
apply (rule conjI, rule_tac [2] conjI)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   670
 apply (rule_tac [2] refl)+
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   671
apply (simp (no_asm_use) add: preal_add_ac)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   672
apply (simp (no_asm_use) add: preal_self_less_add_right preal_add_assoc [symmetric])
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   673
done
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   674
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   675
lemma real_of_preal_not_minus_gt_all: "~ real_of_preal m < - real_of_preal m1"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   676
apply (cut_tac real_of_preal_minus_less_all)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   677
apply (blast dest: real_less_trans elim: real_less_irrefl)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   678
done
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   679
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   680
lemma real_of_preal_minus_less_rev1: "- real_of_preal m1 < - real_of_preal m2
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   681
      ==> real_of_preal m2 < real_of_preal m1"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   682
apply (auto simp add: real_of_preal_def real_less_def real_minus)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   683
apply (rule exI)+
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   684
apply (rule conjI, rule_tac [2] conjI)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   685
 apply (rule_tac [2] refl)+
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   686
apply (auto simp add: preal_add_ac)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   687
apply (simp add: preal_add_assoc [symmetric])
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   688
apply (auto simp add: preal_add_ac)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   689
done
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   690
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   691
lemma real_of_preal_minus_less_rev2: "real_of_preal m1 < real_of_preal m2
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   692
      ==> - real_of_preal m2 < - real_of_preal m1"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   693
apply (auto simp add: real_of_preal_def real_less_def real_minus)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   694
apply (rule exI)+
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   695
apply (rule conjI, rule_tac [2] conjI)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   696
 apply (rule_tac [2] refl)+
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   697
apply (auto simp add: preal_add_ac)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   698
apply (simp add: preal_add_assoc [symmetric])
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   699
apply (auto simp add: preal_add_ac)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   700
done
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   701
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   702
lemma real_of_preal_minus_less_rev_iff: "(- real_of_preal m1 < - real_of_preal m2) =
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   703
      (real_of_preal m2 < real_of_preal m1)"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   704
apply (blast intro!: real_of_preal_minus_less_rev1 real_of_preal_minus_less_rev2)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   705
done
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   706
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   707
declare real_of_preal_minus_less_rev_iff [simp]
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   708
14270
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   709
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   710
subsection{*Linearity of the Ordering*}
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   711
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   712
lemma real_linear: "(x::real) < y | x = y | y < x"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   713
apply (rule_tac x = x in real_of_preal_trichotomyE)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   714
apply (rule_tac [!] x = y in real_of_preal_trichotomyE)
14270
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   715
apply (auto dest!: preal_le_anti_sym 
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   716
            simp add: preal_less_le_iff real_of_preal_minus_less_zero 
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   717
                      real_of_preal_zero_less real_of_preal_minus_less_all)
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   718
done
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   719
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   720
lemma real_neq_iff: "!!w::real. (w ~= z) = (w<z | z<w)"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   721
by (cut_tac real_linear, blast)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   722
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   723
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   724
lemma real_linear_less2: "!!(R1::real). [| R1 < R2 ==> P;  R1 = R2 ==> P;
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   725
                       R2 < R1 ==> P |] ==> P"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   726
apply (cut_tac x = R1 and y = R2 in real_linear, auto)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   727
done
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   728
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   729
lemma real_minus_zero_less_iff: "(0 < -R) = (R < (0::real))"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   730
apply (rule_tac x = R in real_of_preal_trichotomyE)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   731
apply (auto simp add: real_of_preal_not_minus_gt_zero real_of_preal_not_less_zero real_of_preal_zero_less real_of_preal_minus_less_zero)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   732
done
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   733
declare real_minus_zero_less_iff [simp]
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   734
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   735
lemma real_minus_zero_less_iff2: "(-R < 0) = ((0::real) < R)"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   736
apply (rule_tac x = R in real_of_preal_trichotomyE)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   737
apply (auto simp add: real_of_preal_not_minus_gt_zero real_of_preal_not_less_zero real_of_preal_zero_less real_of_preal_minus_less_zero)
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   738
done
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   739
declare real_minus_zero_less_iff2 [simp]
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   740
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   741
ML
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   742
{*
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   743
val real_le_def = thm "real_le_def";
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   744
val real_diff_def = thm "real_diff_def";
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   745
val real_divide_def = thm "real_divide_def";
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   746
val real_of_nat_def = thm "real_of_nat_def";
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   747
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   748
val preal_trans_lemma = thm"preal_trans_lemma";
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   749
val realrel_iff = thm"realrel_iff";
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   750
val realrel_refl = thm"realrel_refl";
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   751
val equiv_realrel = thm"equiv_realrel";
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   752
val equiv_realrel_iff = thm"equiv_realrel_iff";
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   753
val realrel_in_real = thm"realrel_in_real";
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   754
val inj_on_Abs_REAL = thm"inj_on_Abs_REAL";
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   755
val eq_realrelD = thm"eq_realrelD";
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   756
val inj_Rep_REAL = thm"inj_Rep_REAL";
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   757
val inj_real_of_preal = thm"inj_real_of_preal";
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   758
val eq_Abs_REAL = thm"eq_Abs_REAL";
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   759
val real_minus_congruent = thm"real_minus_congruent";
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   760
val real_minus = thm"real_minus";
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   761
val real_minus_minus = thm"real_minus_minus";
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   762
val inj_real_minus = thm"inj_real_minus";
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   763
val real_minus_zero = thm"real_minus_zero";
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   764
val real_minus_zero_iff = thm"real_minus_zero_iff";
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   765
val real_add_congruent2_lemma = thm"real_add_congruent2_lemma";
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   766
val real_add = thm"real_add";
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   767
val real_add_commute = thm"real_add_commute";
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   768
val real_add_assoc = thm"real_add_assoc";
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   769
val real_add_left_commute = thm"real_add_left_commute";
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   770
val real_add_zero_left = thm"real_add_zero_left";
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   771
val real_add_zero_right = thm"real_add_zero_right";
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   772
val real_add_minus = thm"real_add_minus";
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   773
val real_add_minus_left = thm"real_add_minus_left";
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   774
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   775
val real_add_ac = thms"real_add_ac";
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   776
val real_mult_ac = thms"real_mult_ac";
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   777
*}
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 13487
diff changeset
   778
10752
c4f1bf2acf4c tidying, and separation of HOL-Hyperreal from HOL-Real
paulson
parents: 10648
diff changeset
   779
5588
a3ab526bb891 Revised version with Abelian group simprocs
paulson
parents:
diff changeset
   780
end