src/HOL/Real/ex/BinEx.thy
author wenzelm
Sat, 03 Nov 2001 18:42:38 +0100
changeset 12038 343a9888e875
parent 12018 ec054019c910
child 12613 279facb4253a
permissions -rw-r--r--
proper use of bind_thm(s);
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
11595
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
     1
(*  Title:      HOL/Real/ex/BinEx.thy
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
     2
    ID:         $Id$
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
     3
    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
     4
    Copyright   1999  University of Cambridge
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
     5
*)
7577
644f9b4ae764 proper theory setup for Real/ex/BinEx;
wenzelm
parents:
diff changeset
     6
11595
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
     7
header {* Binary arithmetic examples *}
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
     8
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
     9
theory BinEx = Real:
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
    10
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
    11
text {*
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
    12
  Examples of performing binary arithmetic by simplification This time
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
    13
  we use the reals, though the representation is just of integers.
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
    14
*}
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
    15
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
    16
text {* \medskip Addition *}
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
    17
11704
3c50a2cd6f00 * sane numerals (stage 2): plain "num" syntax (removed "#");
wenzelm
parents: 11701
diff changeset
    18
lemma "(1359::real) + -2468 = -1109"
11595
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
    19
  by simp
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
    20
11704
3c50a2cd6f00 * sane numerals (stage 2): plain "num" syntax (removed "#");
wenzelm
parents: 11701
diff changeset
    21
lemma "(93746::real) + -46375 = 47371"
11595
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
    22
  by simp
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
    23
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
    24
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
    25
text {* \medskip Negation *}
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
    26
11704
3c50a2cd6f00 * sane numerals (stage 2): plain "num" syntax (removed "#");
wenzelm
parents: 11701
diff changeset
    27
lemma "- (65745::real) = -65745"
11595
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
    28
  by simp
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
    29
11704
3c50a2cd6f00 * sane numerals (stage 2): plain "num" syntax (removed "#");
wenzelm
parents: 11701
diff changeset
    30
lemma "- (-54321::real) = 54321"
11595
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
    31
  by simp
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
    32
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
    33
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
    34
text {* \medskip Multiplication *}
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
    35
11704
3c50a2cd6f00 * sane numerals (stage 2): plain "num" syntax (removed "#");
wenzelm
parents: 11701
diff changeset
    36
lemma "(-84::real) * 51 = -4284"
11595
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
    37
  by simp
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
    38
11704
3c50a2cd6f00 * sane numerals (stage 2): plain "num" syntax (removed "#");
wenzelm
parents: 11701
diff changeset
    39
lemma "(255::real) * 255 = 65025"
11595
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
    40
  by simp
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
    41
11704
3c50a2cd6f00 * sane numerals (stage 2): plain "num" syntax (removed "#");
wenzelm
parents: 11701
diff changeset
    42
lemma "(1359::real) * -2468 = -3354012"
11595
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
    43
  by simp
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
    44
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
    45
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
    46
text {* \medskip Inequalities *}
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
    47
11704
3c50a2cd6f00 * sane numerals (stage 2): plain "num" syntax (removed "#");
wenzelm
parents: 11701
diff changeset
    48
lemma "(89::real) * 10 \<noteq> 889"
11595
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
    49
  by simp
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
    50
11704
3c50a2cd6f00 * sane numerals (stage 2): plain "num" syntax (removed "#");
wenzelm
parents: 11701
diff changeset
    51
lemma "(13::real) < 18 - 4"
11595
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
    52
  by simp
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
    53
11704
3c50a2cd6f00 * sane numerals (stage 2): plain "num" syntax (removed "#");
wenzelm
parents: 11701
diff changeset
    54
lemma "(-345::real) < -242 + -100"
11595
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
    55
  by simp
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
    56
11704
3c50a2cd6f00 * sane numerals (stage 2): plain "num" syntax (removed "#");
wenzelm
parents: 11701
diff changeset
    57
lemma "(13557456::real) < 18678654"
11595
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
    58
  by simp
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
    59
12018
ec054019c910 Numerals and simprocs for types real and hypreal. The abstract
paulson
parents: 11704
diff changeset
    60
lemma "(999999::real) \<le> (1000001 + 1) - 2"
11595
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
    61
  by simp
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
    62
11704
3c50a2cd6f00 * sane numerals (stage 2): plain "num" syntax (removed "#");
wenzelm
parents: 11701
diff changeset
    63
lemma "(1234567::real) \<le> 1234567"
11595
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
    64
  by simp
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
    65
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
    66
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
    67
text {* \medskip Tests *}
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
    68
12018
ec054019c910 Numerals and simprocs for types real and hypreal. The abstract
paulson
parents: 11704
diff changeset
    69
lemma "(x + y = x) = (y = (0::real))"
11595
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
    70
  by arith
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
    71
12018
ec054019c910 Numerals and simprocs for types real and hypreal. The abstract
paulson
parents: 11704
diff changeset
    72
lemma "(x + y = y) = (x = (0::real))"
11595
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
    73
  by arith
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
    74
12018
ec054019c910 Numerals and simprocs for types real and hypreal. The abstract
paulson
parents: 11704
diff changeset
    75
lemma "(x + y = (0::real)) = (x = -y)"
11595
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
    76
  by arith
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
    77
12018
ec054019c910 Numerals and simprocs for types real and hypreal. The abstract
paulson
parents: 11704
diff changeset
    78
lemma "(x + y = (0::real)) = (y = -x)"
11595
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
    79
  by arith
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
    80
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
    81
lemma "((x + y) < (x + z)) = (y < (z::real))"
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
    82
  by arith
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
    83
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
    84
lemma "((x + z) < (y + z)) = (x < (y::real))"
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
    85
  by arith
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
    86
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
    87
lemma "(\<not> x < y) = (y \<le> (x::real))"
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
    88
  by arith
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
    89
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
    90
lemma "\<not> (x < y \<and> y < (x::real))"
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
    91
  by arith
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
    92
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
    93
lemma "(x::real) < y ==> \<not> y < x"
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
    94
  by arith
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
    95
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
    96
lemma "((x::real) \<noteq> y) = (x < y \<or> y < x)"
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
    97
  by arith
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
    98
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
    99
lemma "(\<not> x \<le> y) = (y < (x::real))"
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
   100
  by arith
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
   101
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
   102
lemma "x \<le> y \<or> y \<le> (x::real)"
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
   103
  by arith
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
   104
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
   105
lemma "x \<le> y \<or> y < (x::real)"
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
   106
  by arith
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
   107
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
   108
lemma "x < y \<or> y \<le> (x::real)"
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
   109
  by arith
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
   110
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
   111
lemma "x \<le> (x::real)"
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
   112
  by arith
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
   113
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
   114
lemma "((x::real) \<le> y) = (x < y \<or> x = y)"
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
   115
  by arith
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
   116
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
   117
lemma "((x::real) \<le> y \<and> y \<le> x) = (x = y)"
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
   118
  by arith
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
   119
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
   120
lemma "\<not>(x < y \<and> y \<le> (x::real))"
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
   121
  by arith
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
   122
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
   123
lemma "\<not>(x \<le> y \<and> y < (x::real))"
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
   124
  by arith
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
   125
12018
ec054019c910 Numerals and simprocs for types real and hypreal. The abstract
paulson
parents: 11704
diff changeset
   126
lemma "(-x < (0::real)) = (0 < x)"
11595
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
   127
  by arith
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
   128
12018
ec054019c910 Numerals and simprocs for types real and hypreal. The abstract
paulson
parents: 11704
diff changeset
   129
lemma "((0::real) < -x) = (x < 0)"
11595
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
   130
  by arith
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
   131
12018
ec054019c910 Numerals and simprocs for types real and hypreal. The abstract
paulson
parents: 11704
diff changeset
   132
lemma "(-x \<le> (0::real)) = (0 \<le> x)"
11595
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
   133
  by arith
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
   134
12018
ec054019c910 Numerals and simprocs for types real and hypreal. The abstract
paulson
parents: 11704
diff changeset
   135
lemma "((0::real) \<le> -x) = (x \<le> 0)"
11595
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
   136
  by arith
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
   137
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
   138
lemma "(x::real) = y \<or> x < y \<or> y < x"
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
   139
  by arith
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
   140
12018
ec054019c910 Numerals and simprocs for types real and hypreal. The abstract
paulson
parents: 11704
diff changeset
   141
lemma "(x::real) = 0 \<or> 0 < x \<or> 0 < -x"
11595
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
   142
  by arith
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
   143
12018
ec054019c910 Numerals and simprocs for types real and hypreal. The abstract
paulson
parents: 11704
diff changeset
   144
lemma "(0::real) \<le> x \<or> 0 \<le> -x"
11595
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
   145
  by arith
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
   146
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
   147
lemma "((x::real) + y \<le> x + z) = (y \<le> z)"
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
   148
  by arith
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
   149
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
   150
lemma "((x::real) + z \<le> y + z) = (x \<le> y)"
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
   151
  by arith
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
   152
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
   153
lemma "(w::real) < x \<and> y < z ==> w + y < x + z"
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
   154
  by arith
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
   155
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
   156
lemma "(w::real) \<le> x \<and> y \<le> z ==> w + y \<le> x + z"
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
   157
  by arith
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
   158
12018
ec054019c910 Numerals and simprocs for types real and hypreal. The abstract
paulson
parents: 11704
diff changeset
   159
lemma "(0::real) \<le> x \<and> 0 \<le> y ==> 0 \<le> x + y"
11595
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
   160
  by arith
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
   161
12018
ec054019c910 Numerals and simprocs for types real and hypreal. The abstract
paulson
parents: 11704
diff changeset
   162
lemma "(0::real) < x \<and> 0 < y ==> 0 < x + y"
11595
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
   163
  by arith
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
   164
12018
ec054019c910 Numerals and simprocs for types real and hypreal. The abstract
paulson
parents: 11704
diff changeset
   165
lemma "(-x < y) = (0 < x + (y::real))"
11595
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
   166
  by arith
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
   167
12018
ec054019c910 Numerals and simprocs for types real and hypreal. The abstract
paulson
parents: 11704
diff changeset
   168
lemma "(x < -y) = (x + y < (0::real))"
11595
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
   169
  by arith
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
   170
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
   171
lemma "(y < x + -z) = (y + z < (x::real))"
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
   172
  by arith
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
   173
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
   174
lemma "(x + -y < z) = (x < z + (y::real))"
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
   175
  by arith
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
   176
12018
ec054019c910 Numerals and simprocs for types real and hypreal. The abstract
paulson
parents: 11704
diff changeset
   177
lemma "x \<le> y ==> x < y + (1::real)"
11595
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
   178
  by arith
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
   179
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
   180
lemma "(x - y) + y = (x::real)"
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
   181
  by arith
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
   182
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
   183
lemma "y + (x - y) = (x::real)"
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
   184
  by arith
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
   185
12018
ec054019c910 Numerals and simprocs for types real and hypreal. The abstract
paulson
parents: 11704
diff changeset
   186
lemma "x - x = (0::real)"
11595
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
   187
  by arith
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
   188
12018
ec054019c910 Numerals and simprocs for types real and hypreal. The abstract
paulson
parents: 11704
diff changeset
   189
lemma "(x - y = 0) = (x = (y::real))"
11595
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
   190
  by arith
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
   191
12018
ec054019c910 Numerals and simprocs for types real and hypreal. The abstract
paulson
parents: 11704
diff changeset
   192
lemma "((0::real) \<le> x + x) = (0 \<le> x)"
11595
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
   193
  by arith
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
   194
12018
ec054019c910 Numerals and simprocs for types real and hypreal. The abstract
paulson
parents: 11704
diff changeset
   195
lemma "(-x \<le> x) = ((0::real) \<le> x)"
11595
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
   196
  by arith
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
   197
12018
ec054019c910 Numerals and simprocs for types real and hypreal. The abstract
paulson
parents: 11704
diff changeset
   198
lemma "(x \<le> -x) = (x \<le> (0::real))"
11595
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
   199
  by arith
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
   200
12018
ec054019c910 Numerals and simprocs for types real and hypreal. The abstract
paulson
parents: 11704
diff changeset
   201
lemma "(-x = (0::real)) = (x = 0)"
11595
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
   202
  by arith
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
   203
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
   204
lemma "-(x - y) = y - (x::real)"
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
   205
  by arith
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
   206
12018
ec054019c910 Numerals and simprocs for types real and hypreal. The abstract
paulson
parents: 11704
diff changeset
   207
lemma "((0::real) < x - y) = (y < x)"
11595
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
   208
  by arith
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
   209
12018
ec054019c910 Numerals and simprocs for types real and hypreal. The abstract
paulson
parents: 11704
diff changeset
   210
lemma "((0::real) \<le> x - y) = (y \<le> x)"
11595
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
   211
  by arith
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
   212
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
   213
lemma "(x + y) - x = (y::real)"
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
   214
  by arith
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
   215
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
   216
lemma "(-x = y) = (x = (-y::real))"
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
   217
  by arith
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
   218
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
   219
lemma "x < (y::real) ==> \<not>(x = y)"
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
   220
  by arith
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
   221
12018
ec054019c910 Numerals and simprocs for types real and hypreal. The abstract
paulson
parents: 11704
diff changeset
   222
lemma "(x \<le> x + y) = ((0::real) \<le> y)"
11595
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
   223
  by arith
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
   224
12018
ec054019c910 Numerals and simprocs for types real and hypreal. The abstract
paulson
parents: 11704
diff changeset
   225
lemma "(y \<le> x + y) = ((0::real) \<le> x)"
11595
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
   226
  by arith
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
   227
12018
ec054019c910 Numerals and simprocs for types real and hypreal. The abstract
paulson
parents: 11704
diff changeset
   228
lemma "(x < x + y) = ((0::real) < y)"
11595
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
   229
  by arith
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
   230
12018
ec054019c910 Numerals and simprocs for types real and hypreal. The abstract
paulson
parents: 11704
diff changeset
   231
lemma "(y < x + y) = ((0::real) < x)"
11595
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
   232
  by arith
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
   233
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
   234
lemma "(x - y) - x = (-y::real)"
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
   235
  by arith
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
   236
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
   237
lemma "(x + y < z) = (x < z - (y::real))"
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
   238
  by arith
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
   239
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
   240
lemma "(x - y < z) = (x < z + (y::real))"
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
   241
  by arith
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
   242
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
   243
lemma "(x < y - z) = (x + z < (y::real))"
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
   244
  by arith
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
   245
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
   246
lemma "(x \<le> y - z) = (x + z \<le> (y::real))"
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
   247
  by arith
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
   248
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
   249
lemma "(x - y \<le> z) = (x \<le> z + (y::real))"
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
   250
  by arith
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
   251
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
   252
lemma "(-x < -y) = (y < (x::real))"
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
   253
  by arith
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
   254
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
   255
lemma "(-x \<le> -y) = (y \<le> (x::real))"
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
   256
  by arith
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
   257
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
   258
lemma "(a + b) - (c + d) = (a - c) + (b - (d::real))"
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
   259
  by arith
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
   260
12018
ec054019c910 Numerals and simprocs for types real and hypreal. The abstract
paulson
parents: 11704
diff changeset
   261
lemma "(0::real) - x = -x"
11595
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
   262
  by arith
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
   263
12018
ec054019c910 Numerals and simprocs for types real and hypreal. The abstract
paulson
parents: 11704
diff changeset
   264
lemma "x - (0::real) = x"
11595
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
   265
  by arith
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
   266
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
   267
lemma "w \<le> x \<and> y < z ==> w + y < x + (z::real)"
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
   268
  by arith
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
   269
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
   270
lemma "w < x \<and> y \<le> z ==> w + y < x + (z::real)"
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
   271
  by arith
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
   272
12018
ec054019c910 Numerals and simprocs for types real and hypreal. The abstract
paulson
parents: 11704
diff changeset
   273
lemma "(0::real) \<le> x \<and> 0 < y ==> 0 < x + (y::real)"
11595
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
   274
  by arith
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
   275
12018
ec054019c910 Numerals and simprocs for types real and hypreal. The abstract
paulson
parents: 11704
diff changeset
   276
lemma "(0::real) < x \<and> 0 \<le> y ==> 0 < x + y"
11595
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
   277
  by arith
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
   278
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
   279
lemma "-x - y = -(x + (y::real))"
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
   280
  by arith
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
   281
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
   282
lemma "x - (-y) = x + (y::real)"
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
   283
  by arith
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
   284
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
   285
lemma "-x - -y = y - (x::real)"
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
   286
  by arith
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
   287
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
   288
lemma "(a - b) + (b - c) = a - (c::real)"
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
   289
  by arith
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
   290
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
   291
lemma "(x = y - z) = (x + z = (y::real))"
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
   292
  by arith
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
   293
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
   294
lemma "(x - y = z) = (x = z + (y::real))"
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
   295
  by arith
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
   296
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
   297
lemma "x - (x - y) = (y::real)"
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
   298
  by arith
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
   299
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
   300
lemma "x - (x + y) = -(y::real)"
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
   301
  by arith
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
   302
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
   303
lemma "x = y ==> x \<le> (y::real)"
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
   304
  by arith
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
   305
12018
ec054019c910 Numerals and simprocs for types real and hypreal. The abstract
paulson
parents: 11704
diff changeset
   306
lemma "(0::real) < x ==> \<not>(x = 0)"
11595
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
   307
  by arith
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
   308
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
   309
lemma "(x + y) * (x - y) = (x * x) - (y * y)"
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
   310
  oops
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
   311
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
   312
lemma "(-x = -y) = (x = (y::real))"
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
   313
  by arith
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
   314
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
   315
lemma "(-x < -y) = (y < (x::real))"
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
   316
  by arith
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
   317
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
   318
lemma "!!a::real. a \<le> b ==> c \<le> d ==> x + y < z ==> a + c \<le> b + d"
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
   319
  by (tactic "fast_arith_tac 1")
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
   320
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
   321
lemma "!!a::real. a < b ==> c < d ==> a - d \<le> b + (-c)"
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
   322
  by (tactic "fast_arith_tac 1")
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
   323
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
   324
lemma "!!a::real. a \<le> b ==> b + b \<le> c ==> a + a \<le> c"
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
   325
  by (tactic "fast_arith_tac 1")
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
   326
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
   327
lemma "!!a::real. a + b \<le> i + j ==> a \<le> b ==> i \<le> j ==> a + a \<le> j + j"
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
   328
  by (tactic "fast_arith_tac 1")
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
   329
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
   330
lemma "!!a::real. a + b < i + j ==> a < b ==> i < j ==> a + a < j + j"
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
   331
  by (tactic "fast_arith_tac 1")
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
   332
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
   333
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"
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
   334
  by arith
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
   335
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
   336
lemma "!!a::real. a + b + c + d \<le> i + j + k + l ==> a \<le> b ==> b \<le> c
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
   337
    ==> c \<le> d ==> i \<le> j ==> j \<le> k ==> k \<le> l ==> a \<le> l"
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
   338
  by (tactic "fast_arith_tac 1")
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
   339
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
   340
lemma "!!a::real. a + b + c + d \<le> i + j + k + l ==> a \<le> b ==> b \<le> c
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
   341
    ==> c \<le> d ==> i \<le> j ==> j \<le> k ==> k \<le> l ==> a + a + a + a \<le> l + l + l + l"
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
   342
  by (tactic "fast_arith_tac 1")
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
   343
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
   344
lemma "!!a::real. a + b + c + d \<le> i + j + k + l ==> a \<le> b ==> b \<le> c
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
   345
    ==> c \<le> d ==> i \<le> j ==> j \<le> k ==> k \<le> l ==> a + a + a + a + a \<le> l + l + l + l + i"
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
   346
  by (tactic "fast_arith_tac 1")
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
   347
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
   348
lemma "!!a::real. a + b + c + d \<le> i + j + k + l ==> a \<le> b ==> b \<le> c
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
   349
    ==> 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"
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
   350
  by (tactic "fast_arith_tac 1")
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
   351
6ef2535fff93 new-style theory;
wenzelm
parents: 7577
diff changeset
   352
end