src/HOL/Complex/ex/BinEx.thy
author paulson
Tue, 03 Feb 2004 11:06:36 +0100
changeset 14373 67a628beb981
parent 14051 4b61bbb0dcab
child 15149 c5c4884634b7
permissions -rw-r--r--
tidying of the complex numbers
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
14051
4b61bbb0dcab examples now use Complex_Main
paulson
parents: 13966
diff changeset
     1
(*  Title:      HOL/Complex/ex/BinEx.thy
13966
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
     2
    ID:         $Id$
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
     3
    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
     4
    Copyright   1999  University of Cambridge
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
     5
*)
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
     6
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
     7
header {* Binary arithmetic examples *}
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
     8
14051
4b61bbb0dcab examples now use Complex_Main
paulson
parents: 13966
diff changeset
     9
theory BinEx = Complex_Main:
13966
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
    10
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
    11
text {*
14051
4b61bbb0dcab examples now use Complex_Main
paulson
parents: 13966
diff changeset
    12
  Examples of performing binary arithmetic by simplification.  This time
13966
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
    13
  we use the reals, though the representation is just of integers.
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
    14
*}
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
    15
14051
4b61bbb0dcab examples now use Complex_Main
paulson
parents: 13966
diff changeset
    16
subsection{*Real Arithmetic*}
4b61bbb0dcab examples now use Complex_Main
paulson
parents: 13966
diff changeset
    17
4b61bbb0dcab examples now use Complex_Main
paulson
parents: 13966
diff changeset
    18
subsubsection {*Addition *}
13966
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
    19
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
    20
lemma "(1359::real) + -2468 = -1109"
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
    21
  by simp
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
    22
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
    23
lemma "(93746::real) + -46375 = 47371"
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
    24
  by simp
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
    25
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
    26
14051
4b61bbb0dcab examples now use Complex_Main
paulson
parents: 13966
diff changeset
    27
subsubsection {*Negation *}
13966
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
    28
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
    29
lemma "- (65745::real) = -65745"
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
    30
  by simp
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
    31
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
    32
lemma "- (-54321::real) = 54321"
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
    33
  by simp
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
    34
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
    35
14051
4b61bbb0dcab examples now use Complex_Main
paulson
parents: 13966
diff changeset
    36
subsubsection {*Multiplication *}
13966
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
    37
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
    38
lemma "(-84::real) * 51 = -4284"
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
    39
  by simp
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
    40
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
    41
lemma "(255::real) * 255 = 65025"
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
    42
  by simp
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
    43
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
    44
lemma "(1359::real) * -2468 = -3354012"
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
    45
  by simp
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
    46
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
    47
14051
4b61bbb0dcab examples now use Complex_Main
paulson
parents: 13966
diff changeset
    48
subsubsection {*Inequalities *}
13966
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
    49
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
    50
lemma "(89::real) * 10 \<noteq> 889"
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
    51
  by simp
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
    52
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
    53
lemma "(13::real) < 18 - 4"
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
    54
  by simp
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
    55
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
    56
lemma "(-345::real) < -242 + -100"
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
    57
  by simp
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
    58
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
    59
lemma "(13557456::real) < 18678654"
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
    60
  by simp
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
    61
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
    62
lemma "(999999::real) \<le> (1000001 + 1) - 2"
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
    63
  by simp
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
    64
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
    65
lemma "(1234567::real) \<le> 1234567"
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
    66
  by simp
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
    67
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
    68
14051
4b61bbb0dcab examples now use Complex_Main
paulson
parents: 13966
diff changeset
    69
subsubsection {*Powers *}
13966
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
    70
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
    71
lemma "2 ^ 15 = (32768::real)"
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
    72
  by simp
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
    73
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
    74
lemma "-3 ^ 7 = (-2187::real)"
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
    75
  by simp
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
    76
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
    77
lemma "13 ^ 7 = (62748517::real)"
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
    78
  by simp
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
    79
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
    80
lemma "3 ^ 15 = (14348907::real)"
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
    81
  by simp
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
    82
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
    83
lemma "-5 ^ 11 = (-48828125::real)"
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
    84
  by simp
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
    85
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
    86
14051
4b61bbb0dcab examples now use Complex_Main
paulson
parents: 13966
diff changeset
    87
subsubsection {*Tests *}
13966
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
    88
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
    89
lemma "(x + y = x) = (y = (0::real))"
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
    90
  by arith
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
    91
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
    92
lemma "(x + y = y) = (x = (0::real))"
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
    93
  by arith
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
    94
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
    95
lemma "(x + y = (0::real)) = (x = -y)"
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
    96
  by arith
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
    97
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
    98
lemma "(x + y = (0::real)) = (y = -x)"
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
    99
  by arith
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
   100
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
   101
lemma "((x + y) < (x + z)) = (y < (z::real))"
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
   102
  by arith
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
   103
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
   104
lemma "((x + z) < (y + z)) = (x < (y::real))"
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
   105
  by arith
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
   106
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
   107
lemma "(\<not> x < y) = (y \<le> (x::real))"
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
   108
  by arith
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
   109
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
   110
lemma "\<not> (x < y \<and> y < (x::real))"
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
   111
  by arith
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
   112
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
   113
lemma "(x::real) < y ==> \<not> y < x"
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
   114
  by arith
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
   115
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
   116
lemma "((x::real) \<noteq> y) = (x < y \<or> y < x)"
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
   117
  by arith
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
   118
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
   119
lemma "(\<not> x \<le> y) = (y < (x::real))"
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
   120
  by arith
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
   121
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
   122
lemma "x \<le> y \<or> y \<le> (x::real)"
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
   123
  by arith
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
   124
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
   125
lemma "x \<le> y \<or> y < (x::real)"
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
   126
  by arith
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
   127
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
   128
lemma "x < y \<or> y \<le> (x::real)"
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
   129
  by arith
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
   130
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
   131
lemma "x \<le> (x::real)"
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
   132
  by arith
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
   133
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
   134
lemma "((x::real) \<le> y) = (x < y \<or> x = y)"
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
   135
  by arith
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
   136
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
   137
lemma "((x::real) \<le> y \<and> y \<le> x) = (x = y)"
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
   138
  by arith
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
   139
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
   140
lemma "\<not>(x < y \<and> y \<le> (x::real))"
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
   141
  by arith
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
   142
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
   143
lemma "\<not>(x \<le> y \<and> y < (x::real))"
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
   144
  by arith
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
   145
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
   146
lemma "(-x < (0::real)) = (0 < x)"
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
   147
  by arith
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
   148
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
   149
lemma "((0::real) < -x) = (x < 0)"
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
   150
  by arith
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
   151
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
   152
lemma "(-x \<le> (0::real)) = (0 \<le> x)"
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
   153
  by arith
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
   154
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
   155
lemma "((0::real) \<le> -x) = (x \<le> 0)"
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
   156
  by arith
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
   157
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
   158
lemma "(x::real) = y \<or> x < y \<or> y < x"
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
   159
  by arith
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
   160
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
   161
lemma "(x::real) = 0 \<or> 0 < x \<or> 0 < -x"
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
   162
  by arith
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
   163
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
   164
lemma "(0::real) \<le> x \<or> 0 \<le> -x"
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
   165
  by arith
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
   166
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
   167
lemma "((x::real) + y \<le> x + z) = (y \<le> z)"
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
   168
  by arith
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
   169
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
   170
lemma "((x::real) + z \<le> y + z) = (x \<le> y)"
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
   171
  by arith
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
   172
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
   173
lemma "(w::real) < x \<and> y < z ==> w + y < x + z"
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
   174
  by arith
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
   175
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
   176
lemma "(w::real) \<le> x \<and> y \<le> z ==> w + y \<le> x + z"
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
   177
  by arith
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
   178
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
   179
lemma "(0::real) \<le> x \<and> 0 \<le> y ==> 0 \<le> x + y"
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
   180
  by arith
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
   181
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
   182
lemma "(0::real) < x \<and> 0 < y ==> 0 < x + y"
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
   183
  by arith
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
   184
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
   185
lemma "(-x < y) = (0 < x + (y::real))"
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
   186
  by arith
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
   187
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
   188
lemma "(x < -y) = (x + y < (0::real))"
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
   189
  by arith
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
   190
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
   191
lemma "(y < x + -z) = (y + z < (x::real))"
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
   192
  by arith
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
   193
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
   194
lemma "(x + -y < z) = (x < z + (y::real))"
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
   195
  by arith
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
   196
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
   197
lemma "x \<le> y ==> x < y + (1::real)"
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
   198
  by arith
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
   199
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
   200
lemma "(x - y) + y = (x::real)"
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
   201
  by arith
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
   202
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
   203
lemma "y + (x - y) = (x::real)"
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
   204
  by arith
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
   205
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
   206
lemma "x - x = (0::real)"
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
   207
  by arith
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
   208
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
   209
lemma "(x - y = 0) = (x = (y::real))"
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
   210
  by arith
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
   211
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
   212
lemma "((0::real) \<le> x + x) = (0 \<le> x)"
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
   213
  by arith
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
   214
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
   215
lemma "(-x \<le> x) = ((0::real) \<le> x)"
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
   216
  by arith
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
   217
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
   218
lemma "(x \<le> -x) = (x \<le> (0::real))"
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
   219
  by arith
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
   220
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
   221
lemma "(-x = (0::real)) = (x = 0)"
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
   222
  by arith
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
   223
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
   224
lemma "-(x - y) = y - (x::real)"
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
   225
  by arith
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
   226
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
   227
lemma "((0::real) < x - y) = (y < x)"
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
   228
  by arith
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
   229
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
   230
lemma "((0::real) \<le> x - y) = (y \<le> x)"
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
   231
  by arith
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
   232
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
   233
lemma "(x + y) - x = (y::real)"
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
   234
  by arith
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
   235
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
   236
lemma "(-x = y) = (x = (-y::real))"
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
   237
  by arith
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
   238
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
   239
lemma "x < (y::real) ==> \<not>(x = y)"
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
   240
  by arith
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
   241
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
   242
lemma "(x \<le> x + y) = ((0::real) \<le> y)"
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
   243
  by arith
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
   244
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
   245
lemma "(y \<le> x + y) = ((0::real) \<le> x)"
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
   246
  by arith
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
   247
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
   248
lemma "(x < x + y) = ((0::real) < y)"
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
   249
  by arith
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
   250
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
   251
lemma "(y < x + y) = ((0::real) < x)"
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
   252
  by arith
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
   253
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
   254
lemma "(x - y) - x = (-y::real)"
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
   255
  by arith
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
   256
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
   257
lemma "(x + y < z) = (x < z - (y::real))"
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
   258
  by arith
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
   259
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
   260
lemma "(x - y < z) = (x < z + (y::real))"
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
   261
  by arith
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
   262
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
   263
lemma "(x < y - z) = (x + z < (y::real))"
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
   264
  by arith
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
   265
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
   266
lemma "(x \<le> y - z) = (x + z \<le> (y::real))"
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
   267
  by arith
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
   268
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
   269
lemma "(x - y \<le> z) = (x \<le> z + (y::real))"
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
   270
  by arith
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
   271
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
   272
lemma "(-x < -y) = (y < (x::real))"
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
   273
  by arith
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
   274
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
   275
lemma "(-x \<le> -y) = (y \<le> (x::real))"
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
   276
  by arith
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
   277
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
   278
lemma "(a + b) - (c + d) = (a - c) + (b - (d::real))"
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
   279
  by arith
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
   280
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
   281
lemma "(0::real) - x = -x"
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
   282
  by arith
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
   283
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
   284
lemma "x - (0::real) = x"
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
   285
  by arith
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
   286
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
   287
lemma "w \<le> x \<and> y < z ==> w + y < x + (z::real)"
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
   288
  by arith
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
   289
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
   290
lemma "w < x \<and> y \<le> z ==> w + y < x + (z::real)"
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
   291
  by arith
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
   292
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
   293
lemma "(0::real) \<le> x \<and> 0 < y ==> 0 < x + (y::real)"
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
   294
  by arith
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
   295
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
   296
lemma "(0::real) < x \<and> 0 \<le> y ==> 0 < x + y"
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
   297
  by arith
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
   298
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
   299
lemma "-x - y = -(x + (y::real))"
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
   300
  by arith
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
   301
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
   302
lemma "x - (-y) = x + (y::real)"
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
   303
  by arith
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
   304
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
   305
lemma "-x - -y = y - (x::real)"
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
   306
  by arith
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
   307
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
   308
lemma "(a - b) + (b - c) = a - (c::real)"
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
   309
  by arith
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
   310
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
   311
lemma "(x = y - z) = (x + z = (y::real))"
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
   312
  by arith
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
   313
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
   314
lemma "(x - y = z) = (x = z + (y::real))"
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
   315
  by arith
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
   316
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
   317
lemma "x - (x - y) = (y::real)"
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
   318
  by arith
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
   319
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
   320
lemma "x - (x + y) = -(y::real)"
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
   321
  by arith
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
   322
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
   323
lemma "x = y ==> x \<le> (y::real)"
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
   324
  by arith
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
   325
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
   326
lemma "(0::real) < x ==> \<not>(x = 0)"
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
   327
  by arith
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
   328
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
   329
lemma "(x + y) * (x - y) = (x * x) - (y * y)"
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
   330
  oops
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
   331
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
   332
lemma "(-x = -y) = (x = (y::real))"
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
   333
  by arith
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
   334
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
   335
lemma "(-x < -y) = (y < (x::real))"
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
   336
  by arith
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
   337
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
   338
lemma "!!a::real. a \<le> b ==> c \<le> d ==> x + y < z ==> a + c \<le> b + d"
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
   339
  by (tactic "fast_arith_tac 1")
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
   340
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
   341
lemma "!!a::real. a < b ==> c < d ==> a - d \<le> b + (-c)"
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
   342
  by (tactic "fast_arith_tac 1")
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
   343
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
   344
lemma "!!a::real. a \<le> b ==> b + b \<le> c ==> a + a \<le> c"
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
   345
  by (tactic "fast_arith_tac 1")
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
   346
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
   347
lemma "!!a::real. a + b \<le> i + j ==> a \<le> b ==> i \<le> j ==> a + a \<le> j + j"
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
   348
  by (tactic "fast_arith_tac 1")
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
   349
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
   350
lemma "!!a::real. a + b < i + j ==> a < b ==> i < j ==> a + a < j + j"
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
   351
  by (tactic "fast_arith_tac 1")
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
   352
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
   353
lemma "!!a::real. a + b + c \<le> i + j + k \<and> a \<le> b \<and> b \<le> c \<and> i \<le> j \<and> j \<le> k --> a + a + a \<le> k + k + k"
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
   354
  by arith
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
   355
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
   356
lemma "!!a::real. a + b + c + d \<le> i + j + k + l ==> a \<le> b ==> b \<le> c
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
   357
    ==> c \<le> d ==> i \<le> j ==> j \<le> k ==> k \<le> l ==> a \<le> l"
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
   358
  by (tactic "fast_arith_tac 1")
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
   359
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
   360
lemma "!!a::real. a + b + c + d \<le> i + j + k + l ==> a \<le> b ==> b \<le> c
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
   361
    ==> c \<le> d ==> i \<le> j ==> j \<le> k ==> k \<le> l ==> a + a + a + a \<le> l + l + l + l"
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
   362
  by (tactic "fast_arith_tac 1")
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
   363
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
   364
lemma "!!a::real. a + b + c + d \<le> i + j + k + l ==> a \<le> b ==> b \<le> c
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
   365
    ==> c \<le> d ==> i \<le> j ==> j \<le> k ==> k \<le> l ==> a + a + a + a + a \<le> l + l + l + l + i"
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
   366
  by (tactic "fast_arith_tac 1")
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
   367
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
   368
lemma "!!a::real. a + b + c + d \<le> i + j + k + l ==> a \<le> b ==> b \<le> c
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
   369
    ==> c \<le> d ==> i \<le> j ==> j \<le> k ==> k \<le> l ==> a + a + a + a + a + a \<le> l + l + l + l + i + l"
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
   370
  by (tactic "fast_arith_tac 1")
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
   371
14051
4b61bbb0dcab examples now use Complex_Main
paulson
parents: 13966
diff changeset
   372
4b61bbb0dcab examples now use Complex_Main
paulson
parents: 13966
diff changeset
   373
subsection{*Complex Arithmetic*}
4b61bbb0dcab examples now use Complex_Main
paulson
parents: 13966
diff changeset
   374
4b61bbb0dcab examples now use Complex_Main
paulson
parents: 13966
diff changeset
   375
lemma "(1359 + 93746*ii) - (2468 + 46375*ii) = -1109 + 47371*ii"
4b61bbb0dcab examples now use Complex_Main
paulson
parents: 13966
diff changeset
   376
  by simp
4b61bbb0dcab examples now use Complex_Main
paulson
parents: 13966
diff changeset
   377
4b61bbb0dcab examples now use Complex_Main
paulson
parents: 13966
diff changeset
   378
lemma "- (65745 + -47371*ii) = -65745 + 47371*ii"
4b61bbb0dcab examples now use Complex_Main
paulson
parents: 13966
diff changeset
   379
  by simp
4b61bbb0dcab examples now use Complex_Main
paulson
parents: 13966
diff changeset
   380
4b61bbb0dcab examples now use Complex_Main
paulson
parents: 13966
diff changeset
   381
text{*Multiplication requires distributive laws.  Perhaps versions instantiated
4b61bbb0dcab examples now use Complex_Main
paulson
parents: 13966
diff changeset
   382
to literal constants should be added to the simpset.*}
4b61bbb0dcab examples now use Complex_Main
paulson
parents: 13966
diff changeset
   383
14373
67a628beb981 tidying of the complex numbers
paulson
parents: 14051
diff changeset
   384
lemmas distrib = left_distrib right_distrib left_diff_distrib right_diff_distrib
14051
4b61bbb0dcab examples now use Complex_Main
paulson
parents: 13966
diff changeset
   385
4b61bbb0dcab examples now use Complex_Main
paulson
parents: 13966
diff changeset
   386
lemma "(1 + ii) * (1 - ii) = 2"
4b61bbb0dcab examples now use Complex_Main
paulson
parents: 13966
diff changeset
   387
by (simp add: distrib)
4b61bbb0dcab examples now use Complex_Main
paulson
parents: 13966
diff changeset
   388
4b61bbb0dcab examples now use Complex_Main
paulson
parents: 13966
diff changeset
   389
lemma "(1 + 2*ii) * (1 + 3*ii) = -5 + 5*ii"
4b61bbb0dcab examples now use Complex_Main
paulson
parents: 13966
diff changeset
   390
by (simp add: distrib)
4b61bbb0dcab examples now use Complex_Main
paulson
parents: 13966
diff changeset
   391
4b61bbb0dcab examples now use Complex_Main
paulson
parents: 13966
diff changeset
   392
lemma "(-84 + 255*ii) + (51 * 255*ii) = -84 + 13260 * ii"
4b61bbb0dcab examples now use Complex_Main
paulson
parents: 13966
diff changeset
   393
by (simp add: distrib)
4b61bbb0dcab examples now use Complex_Main
paulson
parents: 13966
diff changeset
   394
14373
67a628beb981 tidying of the complex numbers
paulson
parents: 14051
diff changeset
   395
text{*No inequalities or linear arithmetic: the complex numbers are unordered!*}
14051
4b61bbb0dcab examples now use Complex_Main
paulson
parents: 13966
diff changeset
   396
4b61bbb0dcab examples now use Complex_Main
paulson
parents: 13966
diff changeset
   397
text{*No powers (not supported yet)*}
4b61bbb0dcab examples now use Complex_Main
paulson
parents: 13966
diff changeset
   398
13966
2160abf7cfe7 removal of the image HOL-Real and merging of HOL-Real-ex with HOL-Complex-ex
paulson
parents:
diff changeset
   399
end