src/HOL/Ring_and_Field.thy
author haftmann
Wed, 24 Oct 2007 07:19:56 +0200
changeset 25166 d813a98a5a36
parent 25152 bfde2f8c0f63
child 25186 f4d1ebffd025
permissions -rw-r--r--
added subclass_rule
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
14265
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
     1
(*  Title:   HOL/Ring_and_Field.thy
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
     2
    ID:      $Id$
23477
f4b83f03cac9 tuned and renamed group_eq_simps and ring_eq_simps
nipkow
parents: 23413
diff changeset
     3
    Author:  Gertrud Bauer, Steven Obua, Tobias Nipkow, Lawrence C Paulson, and Markus Wenzel,
16775
c1b87ef4a1c3 added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
avigad
parents: 16568
diff changeset
     4
             with contributions by Jeremy Avigad
14265
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
     5
*)
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
     6
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
     7
header {* (Ordered) Rings and Fields *}
14265
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
     8
15229
1eb23f805c06 new simprules for abs and for things like a/b<1
paulson
parents: 15228
diff changeset
     9
theory Ring_and_Field
15140
322485b816ac import -> imports
nipkow
parents: 15131
diff changeset
    10
imports OrderedGroup
15131
c69542757a4d New theory header syntax.
nipkow
parents: 15077
diff changeset
    11
begin
14504
7a3d80e276d4 new type class abelian_group
paulson
parents: 14475
diff changeset
    12
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
    13
text {*
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
    14
  The theory of partially ordered rings is taken from the books:
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
    15
  \begin{itemize}
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
    16
  \item \emph{Lattice Theory} by Garret Birkhoff, American Mathematical Society 1979 
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
    17
  \item \emph{Partially Ordered Algebraic Systems}, Pergamon Press 1963
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
    18
  \end{itemize}
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
    19
  Most of the used notions can also be looked up in 
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
    20
  \begin{itemize}
14770
fe9504ba63d5 removed duplicate thms;
wenzelm
parents: 14754
diff changeset
    21
  \item \url{http://www.mathworld.com} by Eric Weisstein et. al.
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
    22
  \item \emph{Algebra I} by van der Waerden, Springer.
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
    23
  \end{itemize}
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
    24
*}
14504
7a3d80e276d4 new type class abelian_group
paulson
parents: 14475
diff changeset
    25
22390
378f34b1e380 now using "class"
haftmann
parents: 21328
diff changeset
    26
class semiring = ab_semigroup_add + semigroup_mult +
25062
af5ef0d4d655 global class syntax
haftmann
parents: 24748
diff changeset
    27
  assumes left_distrib: "(a + b) * c = a * c + b * c"
af5ef0d4d655 global class syntax
haftmann
parents: 24748
diff changeset
    28
  assumes right_distrib: "a * (b + c) = a * b + a * c"
25152
bfde2f8c0f63 partially localized
haftmann
parents: 25078
diff changeset
    29
begin
bfde2f8c0f63 partially localized
haftmann
parents: 25078
diff changeset
    30
bfde2f8c0f63 partially localized
haftmann
parents: 25078
diff changeset
    31
text{*For the @{text combine_numerals} simproc*}
bfde2f8c0f63 partially localized
haftmann
parents: 25078
diff changeset
    32
lemma combine_common_factor:
bfde2f8c0f63 partially localized
haftmann
parents: 25078
diff changeset
    33
  "a * e + (b * e + c) = (a + b) * e + c"
bfde2f8c0f63 partially localized
haftmann
parents: 25078
diff changeset
    34
  by (simp add: left_distrib add_ac)
bfde2f8c0f63 partially localized
haftmann
parents: 25078
diff changeset
    35
bfde2f8c0f63 partially localized
haftmann
parents: 25078
diff changeset
    36
end
14504
7a3d80e276d4 new type class abelian_group
paulson
parents: 14475
diff changeset
    37
22390
378f34b1e380 now using "class"
haftmann
parents: 21328
diff changeset
    38
class mult_zero = times + zero +
25062
af5ef0d4d655 global class syntax
haftmann
parents: 24748
diff changeset
    39
  assumes mult_zero_left [simp]: "0 * a = 0"
af5ef0d4d655 global class syntax
haftmann
parents: 24748
diff changeset
    40
  assumes mult_zero_right [simp]: "a * 0 = 0"
21199
2d83f93c3580 * Added annihilation axioms ("x * 0 = 0") to axclass semiring_0.
krauss
parents: 20633
diff changeset
    41
22390
378f34b1e380 now using "class"
haftmann
parents: 21328
diff changeset
    42
class semiring_0 = semiring + comm_monoid_add + mult_zero
21199
2d83f93c3580 * Added annihilation axioms ("x * 0 = 0") to axclass semiring_0.
krauss
parents: 20633
diff changeset
    43
22390
378f34b1e380 now using "class"
haftmann
parents: 21328
diff changeset
    44
class semiring_0_cancel = semiring + comm_monoid_add + cancel_ab_semigroup_add
14504
7a3d80e276d4 new type class abelian_group
paulson
parents: 14475
diff changeset
    45
21199
2d83f93c3580 * Added annihilation axioms ("x * 0 = 0") to axclass semiring_0.
krauss
parents: 20633
diff changeset
    46
instance semiring_0_cancel \<subseteq> semiring_0
2d83f93c3580 * Added annihilation axioms ("x * 0 = 0") to axclass semiring_0.
krauss
parents: 20633
diff changeset
    47
proof
2d83f93c3580 * Added annihilation axioms ("x * 0 = 0") to axclass semiring_0.
krauss
parents: 20633
diff changeset
    48
  fix a :: 'a
2d83f93c3580 * Added annihilation axioms ("x * 0 = 0") to axclass semiring_0.
krauss
parents: 20633
diff changeset
    49
  have "0 * a + 0 * a = 0 * a + 0"
2d83f93c3580 * Added annihilation axioms ("x * 0 = 0") to axclass semiring_0.
krauss
parents: 20633
diff changeset
    50
    by (simp add: left_distrib [symmetric])
2d83f93c3580 * Added annihilation axioms ("x * 0 = 0") to axclass semiring_0.
krauss
parents: 20633
diff changeset
    51
  thus "0 * a = 0"
2d83f93c3580 * Added annihilation axioms ("x * 0 = 0") to axclass semiring_0.
krauss
parents: 20633
diff changeset
    52
    by (simp only: add_left_cancel)
25152
bfde2f8c0f63 partially localized
haftmann
parents: 25078
diff changeset
    53
next
bfde2f8c0f63 partially localized
haftmann
parents: 25078
diff changeset
    54
  fix a :: 'a
21199
2d83f93c3580 * Added annihilation axioms ("x * 0 = 0") to axclass semiring_0.
krauss
parents: 20633
diff changeset
    55
  have "a * 0 + a * 0 = a * 0 + 0"
2d83f93c3580 * Added annihilation axioms ("x * 0 = 0") to axclass semiring_0.
krauss
parents: 20633
diff changeset
    56
    by (simp add: right_distrib [symmetric])
2d83f93c3580 * Added annihilation axioms ("x * 0 = 0") to axclass semiring_0.
krauss
parents: 20633
diff changeset
    57
  thus "a * 0 = 0"
2d83f93c3580 * Added annihilation axioms ("x * 0 = 0") to axclass semiring_0.
krauss
parents: 20633
diff changeset
    58
    by (simp only: add_left_cancel)
2d83f93c3580 * Added annihilation axioms ("x * 0 = 0") to axclass semiring_0.
krauss
parents: 20633
diff changeset
    59
qed
14940
b9ab8babd8b3 Further development of matrix theory
obua
parents: 14770
diff changeset
    60
25152
bfde2f8c0f63 partially localized
haftmann
parents: 25078
diff changeset
    61
interpretation semiring_0_cancel \<subseteq> semiring_0
bfde2f8c0f63 partially localized
haftmann
parents: 25078
diff changeset
    62
proof unfold_locales
bfde2f8c0f63 partially localized
haftmann
parents: 25078
diff changeset
    63
  fix a :: 'a
bfde2f8c0f63 partially localized
haftmann
parents: 25078
diff changeset
    64
  have "plus (times zero a) (times zero a) = plus (times zero a) zero"
bfde2f8c0f63 partially localized
haftmann
parents: 25078
diff changeset
    65
    by (simp add: left_distrib [symmetric])
bfde2f8c0f63 partially localized
haftmann
parents: 25078
diff changeset
    66
  thus "times zero a = zero"
bfde2f8c0f63 partially localized
haftmann
parents: 25078
diff changeset
    67
    by (simp only: add_left_cancel)
bfde2f8c0f63 partially localized
haftmann
parents: 25078
diff changeset
    68
next
bfde2f8c0f63 partially localized
haftmann
parents: 25078
diff changeset
    69
  fix a :: 'a
bfde2f8c0f63 partially localized
haftmann
parents: 25078
diff changeset
    70
  have "plus (times a zero) (times a zero) = plus (times a zero) zero"
bfde2f8c0f63 partially localized
haftmann
parents: 25078
diff changeset
    71
    by (simp add: right_distrib [symmetric])
bfde2f8c0f63 partially localized
haftmann
parents: 25078
diff changeset
    72
  thus "times a zero = zero"
bfde2f8c0f63 partially localized
haftmann
parents: 25078
diff changeset
    73
    by (simp only: add_left_cancel)
bfde2f8c0f63 partially localized
haftmann
parents: 25078
diff changeset
    74
qed
bfde2f8c0f63 partially localized
haftmann
parents: 25078
diff changeset
    75
22390
378f34b1e380 now using "class"
haftmann
parents: 21328
diff changeset
    76
class comm_semiring = ab_semigroup_add + ab_semigroup_mult +
25062
af5ef0d4d655 global class syntax
haftmann
parents: 24748
diff changeset
    77
  assumes distrib: "(a + b) * c = a * c + b * c"
25152
bfde2f8c0f63 partially localized
haftmann
parents: 25078
diff changeset
    78
begin
14504
7a3d80e276d4 new type class abelian_group
paulson
parents: 14475
diff changeset
    79
25152
bfde2f8c0f63 partially localized
haftmann
parents: 25078
diff changeset
    80
subclass semiring
bfde2f8c0f63 partially localized
haftmann
parents: 25078
diff changeset
    81
proof unfold_locales
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
    82
  fix a b c :: 'a
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
    83
  show "(a + b) * c = a * c + b * c" by (simp add: distrib)
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
    84
  have "a * (b + c) = (b + c) * a" by (simp add: mult_ac)
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
    85
  also have "... = b * a + c * a" by (simp only: distrib)
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
    86
  also have "... = a * b + a * c" by (simp add: mult_ac)
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
    87
  finally show "a * (b + c) = a * b + a * c" by blast
14504
7a3d80e276d4 new type class abelian_group
paulson
parents: 14475
diff changeset
    88
qed
7a3d80e276d4 new type class abelian_group
paulson
parents: 14475
diff changeset
    89
25152
bfde2f8c0f63 partially localized
haftmann
parents: 25078
diff changeset
    90
end
14504
7a3d80e276d4 new type class abelian_group
paulson
parents: 14475
diff changeset
    91
25152
bfde2f8c0f63 partially localized
haftmann
parents: 25078
diff changeset
    92
class comm_semiring_0 = comm_semiring + comm_monoid_add + mult_zero
bfde2f8c0f63 partially localized
haftmann
parents: 25078
diff changeset
    93
begin
bfde2f8c0f63 partially localized
haftmann
parents: 25078
diff changeset
    94
bfde2f8c0f63 partially localized
haftmann
parents: 25078
diff changeset
    95
subclass semiring_0 by unfold_locales
bfde2f8c0f63 partially localized
haftmann
parents: 25078
diff changeset
    96
bfde2f8c0f63 partially localized
haftmann
parents: 25078
diff changeset
    97
end
14504
7a3d80e276d4 new type class abelian_group
paulson
parents: 14475
diff changeset
    98
22390
378f34b1e380 now using "class"
haftmann
parents: 21328
diff changeset
    99
class comm_semiring_0_cancel = comm_semiring + comm_monoid_add + cancel_ab_semigroup_add
14940
b9ab8babd8b3 Further development of matrix theory
obua
parents: 14770
diff changeset
   100
b9ab8babd8b3 Further development of matrix theory
obua
parents: 14770
diff changeset
   101
instance comm_semiring_0_cancel \<subseteq> semiring_0_cancel ..
b9ab8babd8b3 Further development of matrix theory
obua
parents: 14770
diff changeset
   102
25152
bfde2f8c0f63 partially localized
haftmann
parents: 25078
diff changeset
   103
interpretation comm_semiring_0_cancel \<subseteq> semiring_0_cancel by unfold_locales
21199
2d83f93c3580 * Added annihilation axioms ("x * 0 = 0") to axclass semiring_0.
krauss
parents: 20633
diff changeset
   104
22390
378f34b1e380 now using "class"
haftmann
parents: 21328
diff changeset
   105
class zero_neq_one = zero + one +
25062
af5ef0d4d655 global class syntax
haftmann
parents: 24748
diff changeset
   106
  assumes zero_neq_one [simp]: "0 \<noteq> 1"
14265
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   107
22390
378f34b1e380 now using "class"
haftmann
parents: 21328
diff changeset
   108
class semiring_1 = zero_neq_one + semiring_0 + monoid_mult
14504
7a3d80e276d4 new type class abelian_group
paulson
parents: 14475
diff changeset
   109
22390
378f34b1e380 now using "class"
haftmann
parents: 21328
diff changeset
   110
class comm_semiring_1 = zero_neq_one + comm_semiring_0 + comm_monoid_mult
378f34b1e380 now using "class"
haftmann
parents: 21328
diff changeset
   111
  (*previously almost_semiring*)
25152
bfde2f8c0f63 partially localized
haftmann
parents: 25078
diff changeset
   112
begin
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   113
25152
bfde2f8c0f63 partially localized
haftmann
parents: 25078
diff changeset
   114
subclass semiring_1 by unfold_locales
bfde2f8c0f63 partially localized
haftmann
parents: 25078
diff changeset
   115
bfde2f8c0f63 partially localized
haftmann
parents: 25078
diff changeset
   116
end
14421
ee97b6463cb4 new Ring_and_Field hierarchy, eliminating redundant axioms
paulson
parents: 14398
diff changeset
   117
22390
378f34b1e380 now using "class"
haftmann
parents: 21328
diff changeset
   118
class no_zero_divisors = zero + times +
25062
af5ef0d4d655 global class syntax
haftmann
parents: 24748
diff changeset
   119
  assumes no_zero_divisors: "a \<noteq> 0 \<Longrightarrow> b \<noteq> 0 \<Longrightarrow> a * b \<noteq> 0"
14504
7a3d80e276d4 new type class abelian_group
paulson
parents: 14475
diff changeset
   120
22390
378f34b1e380 now using "class"
haftmann
parents: 21328
diff changeset
   121
class semiring_1_cancel = semiring + comm_monoid_add + zero_neq_one
378f34b1e380 now using "class"
haftmann
parents: 21328
diff changeset
   122
  + cancel_ab_semigroup_add + monoid_mult
14940
b9ab8babd8b3 Further development of matrix theory
obua
parents: 14770
diff changeset
   123
b9ab8babd8b3 Further development of matrix theory
obua
parents: 14770
diff changeset
   124
instance semiring_1_cancel \<subseteq> semiring_0_cancel ..
b9ab8babd8b3 Further development of matrix theory
obua
parents: 14770
diff changeset
   125
25152
bfde2f8c0f63 partially localized
haftmann
parents: 25078
diff changeset
   126
interpretation semiring_1_cancel \<subseteq> semiring_0_cancel by unfold_locales
bfde2f8c0f63 partially localized
haftmann
parents: 25078
diff changeset
   127
bfde2f8c0f63 partially localized
haftmann
parents: 25078
diff changeset
   128
subclass (in semiring_1_cancel) semiring_1 by unfold_locales
21199
2d83f93c3580 * Added annihilation axioms ("x * 0 = 0") to axclass semiring_0.
krauss
parents: 20633
diff changeset
   129
22390
378f34b1e380 now using "class"
haftmann
parents: 21328
diff changeset
   130
class comm_semiring_1_cancel = comm_semiring + comm_monoid_add + comm_monoid_mult
378f34b1e380 now using "class"
haftmann
parents: 21328
diff changeset
   131
  + zero_neq_one + cancel_ab_semigroup_add
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   132
14940
b9ab8babd8b3 Further development of matrix theory
obua
parents: 14770
diff changeset
   133
instance comm_semiring_1_cancel \<subseteq> semiring_1_cancel ..
b9ab8babd8b3 Further development of matrix theory
obua
parents: 14770
diff changeset
   134
25152
bfde2f8c0f63 partially localized
haftmann
parents: 25078
diff changeset
   135
interpretation comm_semiring_1_cancel \<subseteq> semiring_1_cancel by unfold_locales
bfde2f8c0f63 partially localized
haftmann
parents: 25078
diff changeset
   136
bfde2f8c0f63 partially localized
haftmann
parents: 25078
diff changeset
   137
subclass (in comm_semiring_1_cancel) comm_semiring_0_cancel by unfold_locales
14940
b9ab8babd8b3 Further development of matrix theory
obua
parents: 14770
diff changeset
   138
21199
2d83f93c3580 * Added annihilation axioms ("x * 0 = 0") to axclass semiring_0.
krauss
parents: 20633
diff changeset
   139
instance comm_semiring_1_cancel \<subseteq> comm_semiring_1 ..
2d83f93c3580 * Added annihilation axioms ("x * 0 = 0") to axclass semiring_0.
krauss
parents: 20633
diff changeset
   140
25152
bfde2f8c0f63 partially localized
haftmann
parents: 25078
diff changeset
   141
interpretation comm_semiring_1_cancel \<subseteq> comm_semiring_1 by unfold_locales
bfde2f8c0f63 partially localized
haftmann
parents: 25078
diff changeset
   142
22390
378f34b1e380 now using "class"
haftmann
parents: 21328
diff changeset
   143
class ring = semiring + ab_group_add
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   144
14940
b9ab8babd8b3 Further development of matrix theory
obua
parents: 14770
diff changeset
   145
instance ring \<subseteq> semiring_0_cancel ..
14504
7a3d80e276d4 new type class abelian_group
paulson
parents: 14475
diff changeset
   146
25152
bfde2f8c0f63 partially localized
haftmann
parents: 25078
diff changeset
   147
interpretation ring \<subseteq> semiring_0_cancel by unfold_locales
bfde2f8c0f63 partially localized
haftmann
parents: 25078
diff changeset
   148
bfde2f8c0f63 partially localized
haftmann
parents: 25078
diff changeset
   149
context ring
bfde2f8c0f63 partially localized
haftmann
parents: 25078
diff changeset
   150
begin
bfde2f8c0f63 partially localized
haftmann
parents: 25078
diff changeset
   151
bfde2f8c0f63 partially localized
haftmann
parents: 25078
diff changeset
   152
text {* Distribution rules *}
bfde2f8c0f63 partially localized
haftmann
parents: 25078
diff changeset
   153
bfde2f8c0f63 partially localized
haftmann
parents: 25078
diff changeset
   154
lemma minus_mult_left: "- (a * b) = - a * b"
bfde2f8c0f63 partially localized
haftmann
parents: 25078
diff changeset
   155
  by (rule equals_zero_I) (simp add: left_distrib [symmetric]) 
bfde2f8c0f63 partially localized
haftmann
parents: 25078
diff changeset
   156
bfde2f8c0f63 partially localized
haftmann
parents: 25078
diff changeset
   157
lemma minus_mult_right: "- (a * b) = a * - b"
bfde2f8c0f63 partially localized
haftmann
parents: 25078
diff changeset
   158
  by (rule equals_zero_I) (simp add: right_distrib [symmetric]) 
bfde2f8c0f63 partially localized
haftmann
parents: 25078
diff changeset
   159
bfde2f8c0f63 partially localized
haftmann
parents: 25078
diff changeset
   160
lemma minus_mult_minus [simp]: "- a * - b = a * b"
bfde2f8c0f63 partially localized
haftmann
parents: 25078
diff changeset
   161
  by (simp add: minus_mult_left [symmetric] minus_mult_right [symmetric])
bfde2f8c0f63 partially localized
haftmann
parents: 25078
diff changeset
   162
bfde2f8c0f63 partially localized
haftmann
parents: 25078
diff changeset
   163
lemma minus_mult_commute: "- a * b = a * - b"
bfde2f8c0f63 partially localized
haftmann
parents: 25078
diff changeset
   164
  by (simp add: minus_mult_left [symmetric] minus_mult_right [symmetric])
bfde2f8c0f63 partially localized
haftmann
parents: 25078
diff changeset
   165
bfde2f8c0f63 partially localized
haftmann
parents: 25078
diff changeset
   166
lemma right_diff_distrib: "a * (b - c) = a * b - a * c"
bfde2f8c0f63 partially localized
haftmann
parents: 25078
diff changeset
   167
  by (simp add: right_distrib diff_minus 
bfde2f8c0f63 partially localized
haftmann
parents: 25078
diff changeset
   168
    minus_mult_left [symmetric] minus_mult_right [symmetric]) 
bfde2f8c0f63 partially localized
haftmann
parents: 25078
diff changeset
   169
bfde2f8c0f63 partially localized
haftmann
parents: 25078
diff changeset
   170
lemma left_diff_distrib: "(a - b) * c = a * c - b * c"
bfde2f8c0f63 partially localized
haftmann
parents: 25078
diff changeset
   171
  by (simp add: left_distrib diff_minus 
bfde2f8c0f63 partially localized
haftmann
parents: 25078
diff changeset
   172
    minus_mult_left [symmetric] minus_mult_right [symmetric]) 
bfde2f8c0f63 partially localized
haftmann
parents: 25078
diff changeset
   173
bfde2f8c0f63 partially localized
haftmann
parents: 25078
diff changeset
   174
lemmas ring_distribs =
bfde2f8c0f63 partially localized
haftmann
parents: 25078
diff changeset
   175
  right_distrib left_distrib left_diff_distrib right_diff_distrib
bfde2f8c0f63 partially localized
haftmann
parents: 25078
diff changeset
   176
bfde2f8c0f63 partially localized
haftmann
parents: 25078
diff changeset
   177
end
bfde2f8c0f63 partially localized
haftmann
parents: 25078
diff changeset
   178
bfde2f8c0f63 partially localized
haftmann
parents: 25078
diff changeset
   179
lemmas ring_distribs =
bfde2f8c0f63 partially localized
haftmann
parents: 25078
diff changeset
   180
  right_distrib left_distrib left_diff_distrib right_diff_distrib
bfde2f8c0f63 partially localized
haftmann
parents: 25078
diff changeset
   181
bfde2f8c0f63 partially localized
haftmann
parents: 25078
diff changeset
   182
text{*This list of rewrites simplifies ring terms by multiplying
bfde2f8c0f63 partially localized
haftmann
parents: 25078
diff changeset
   183
everything out and bringing sums and products into a canonical form
bfde2f8c0f63 partially localized
haftmann
parents: 25078
diff changeset
   184
(by ordered rewriting). As a result it decides ring equalities but
bfde2f8c0f63 partially localized
haftmann
parents: 25078
diff changeset
   185
also helps with inequalities. *}
bfde2f8c0f63 partially localized
haftmann
parents: 25078
diff changeset
   186
lemmas ring_simps = group_simps ring_distribs
bfde2f8c0f63 partially localized
haftmann
parents: 25078
diff changeset
   187
22390
378f34b1e380 now using "class"
haftmann
parents: 21328
diff changeset
   188
class comm_ring = comm_semiring + ab_group_add
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   189
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   190
instance comm_ring \<subseteq> ring ..
14504
7a3d80e276d4 new type class abelian_group
paulson
parents: 14475
diff changeset
   191
25152
bfde2f8c0f63 partially localized
haftmann
parents: 25078
diff changeset
   192
interpretation comm_ring \<subseteq> ring by unfold_locales
bfde2f8c0f63 partially localized
haftmann
parents: 25078
diff changeset
   193
bfde2f8c0f63 partially localized
haftmann
parents: 25078
diff changeset
   194
instance comm_ring \<subseteq> comm_semiring_0 ..
bfde2f8c0f63 partially localized
haftmann
parents: 25078
diff changeset
   195
bfde2f8c0f63 partially localized
haftmann
parents: 25078
diff changeset
   196
interpretation comm_ring \<subseteq> comm_semiring_0 by unfold_locales
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   197
22390
378f34b1e380 now using "class"
haftmann
parents: 21328
diff changeset
   198
class ring_1 = ring + zero_neq_one + monoid_mult
14265
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   199
14940
b9ab8babd8b3 Further development of matrix theory
obua
parents: 14770
diff changeset
   200
instance ring_1 \<subseteq> semiring_1_cancel ..
b9ab8babd8b3 Further development of matrix theory
obua
parents: 14770
diff changeset
   201
25152
bfde2f8c0f63 partially localized
haftmann
parents: 25078
diff changeset
   202
interpretation ring_1 \<subseteq> semiring_1_cancel by unfold_locales
bfde2f8c0f63 partially localized
haftmann
parents: 25078
diff changeset
   203
22390
378f34b1e380 now using "class"
haftmann
parents: 21328
diff changeset
   204
class comm_ring_1 = comm_ring + zero_neq_one + comm_monoid_mult
378f34b1e380 now using "class"
haftmann
parents: 21328
diff changeset
   205
  (*previously ring*)
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   206
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   207
instance comm_ring_1 \<subseteq> ring_1 ..
14421
ee97b6463cb4 new Ring_and_Field hierarchy, eliminating redundant axioms
paulson
parents: 14398
diff changeset
   208
25152
bfde2f8c0f63 partially localized
haftmann
parents: 25078
diff changeset
   209
interpretation comm_ring_1 \<subseteq> ring_1 by unfold_locales
bfde2f8c0f63 partially localized
haftmann
parents: 25078
diff changeset
   210
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   211
instance comm_ring_1 \<subseteq> comm_semiring_1_cancel ..
14265
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   212
25152
bfde2f8c0f63 partially localized
haftmann
parents: 25078
diff changeset
   213
interpretation comm_ring_1 \<subseteq> comm_semiring_1_cancel by unfold_locales
bfde2f8c0f63 partially localized
haftmann
parents: 25078
diff changeset
   214
22990
775e9de3db48 added classes ring_no_zero_divisors and dom (non-commutative version of idom);
huffman
parents: 22987
diff changeset
   215
class ring_no_zero_divisors = ring + no_zero_divisors
775e9de3db48 added classes ring_no_zero_divisors and dom (non-commutative version of idom);
huffman
parents: 22987
diff changeset
   216
23544
4b4165cb3e0d rename class dom to ring_1_no_zero_divisors
huffman
parents: 23527
diff changeset
   217
class ring_1_no_zero_divisors = ring_1 + ring_no_zero_divisors
22990
775e9de3db48 added classes ring_no_zero_divisors and dom (non-commutative version of idom);
huffman
parents: 22987
diff changeset
   218
22390
378f34b1e380 now using "class"
haftmann
parents: 21328
diff changeset
   219
class idom = comm_ring_1 + no_zero_divisors
14421
ee97b6463cb4 new Ring_and_Field hierarchy, eliminating redundant axioms
paulson
parents: 14398
diff changeset
   220
23544
4b4165cb3e0d rename class dom to ring_1_no_zero_divisors
huffman
parents: 23527
diff changeset
   221
instance idom \<subseteq> ring_1_no_zero_divisors ..
22990
775e9de3db48 added classes ring_no_zero_divisors and dom (non-commutative version of idom);
huffman
parents: 22987
diff changeset
   222
25152
bfde2f8c0f63 partially localized
haftmann
parents: 25078
diff changeset
   223
interpretation idom \<subseteq> ring_1_no_zero_divisors by unfold_locales
bfde2f8c0f63 partially localized
haftmann
parents: 25078
diff changeset
   224
22390
378f34b1e380 now using "class"
haftmann
parents: 21328
diff changeset
   225
class division_ring = ring_1 + inverse +
25062
af5ef0d4d655 global class syntax
haftmann
parents: 24748
diff changeset
   226
  assumes left_inverse [simp]:  "a \<noteq> 0 \<Longrightarrow> inverse a * a = 1"
af5ef0d4d655 global class syntax
haftmann
parents: 24748
diff changeset
   227
  assumes right_inverse [simp]: "a \<noteq> 0 \<Longrightarrow> a * inverse a = 1"
20496
23eb6034c06d added axclass division_ring (like field without commutativity; includes e.g. quaternions) and generalized some theorems from field to division_ring
huffman
parents: 19404
diff changeset
   228
23544
4b4165cb3e0d rename class dom to ring_1_no_zero_divisors
huffman
parents: 23527
diff changeset
   229
instance division_ring \<subseteq> ring_1_no_zero_divisors
22987
550709aa8e66 instance division_ring < no_zero_divisors; clean up field instance proofs
huffman
parents: 22842
diff changeset
   230
proof
550709aa8e66 instance division_ring < no_zero_divisors; clean up field instance proofs
huffman
parents: 22842
diff changeset
   231
  fix a b :: 'a
550709aa8e66 instance division_ring < no_zero_divisors; clean up field instance proofs
huffman
parents: 22842
diff changeset
   232
  assume a: "a \<noteq> 0" and b: "b \<noteq> 0"
550709aa8e66 instance division_ring < no_zero_divisors; clean up field instance proofs
huffman
parents: 22842
diff changeset
   233
  show "a * b \<noteq> 0"
550709aa8e66 instance division_ring < no_zero_divisors; clean up field instance proofs
huffman
parents: 22842
diff changeset
   234
  proof
550709aa8e66 instance division_ring < no_zero_divisors; clean up field instance proofs
huffman
parents: 22842
diff changeset
   235
    assume ab: "a * b = 0"
550709aa8e66 instance division_ring < no_zero_divisors; clean up field instance proofs
huffman
parents: 22842
diff changeset
   236
    hence "0 = inverse a * (a * b) * inverse b"
550709aa8e66 instance division_ring < no_zero_divisors; clean up field instance proofs
huffman
parents: 22842
diff changeset
   237
      by simp
550709aa8e66 instance division_ring < no_zero_divisors; clean up field instance proofs
huffman
parents: 22842
diff changeset
   238
    also have "\<dots> = (inverse a * a) * (b * inverse b)"
550709aa8e66 instance division_ring < no_zero_divisors; clean up field instance proofs
huffman
parents: 22842
diff changeset
   239
      by (simp only: mult_assoc)
550709aa8e66 instance division_ring < no_zero_divisors; clean up field instance proofs
huffman
parents: 22842
diff changeset
   240
    also have "\<dots> = 1"
550709aa8e66 instance division_ring < no_zero_divisors; clean up field instance proofs
huffman
parents: 22842
diff changeset
   241
      using a b by simp
550709aa8e66 instance division_ring < no_zero_divisors; clean up field instance proofs
huffman
parents: 22842
diff changeset
   242
    finally show False
550709aa8e66 instance division_ring < no_zero_divisors; clean up field instance proofs
huffman
parents: 22842
diff changeset
   243
      by simp
550709aa8e66 instance division_ring < no_zero_divisors; clean up field instance proofs
huffman
parents: 22842
diff changeset
   244
  qed
550709aa8e66 instance division_ring < no_zero_divisors; clean up field instance proofs
huffman
parents: 22842
diff changeset
   245
qed
20496
23eb6034c06d added axclass division_ring (like field without commutativity; includes e.g. quaternions) and generalized some theorems from field to division_ring
huffman
parents: 19404
diff changeset
   246
25152
bfde2f8c0f63 partially localized
haftmann
parents: 25078
diff changeset
   247
interpretation division_ring \<subseteq> ring_1_no_zero_divisors
bfde2f8c0f63 partially localized
haftmann
parents: 25078
diff changeset
   248
proof unfold_locales
bfde2f8c0f63 partially localized
haftmann
parents: 25078
diff changeset
   249
  fix a b :: 'a
bfde2f8c0f63 partially localized
haftmann
parents: 25078
diff changeset
   250
  assume a: "a \<noteq> zero" and b: "b \<noteq> zero"
bfde2f8c0f63 partially localized
haftmann
parents: 25078
diff changeset
   251
  show "times a b \<noteq> zero"
bfde2f8c0f63 partially localized
haftmann
parents: 25078
diff changeset
   252
  proof
bfde2f8c0f63 partially localized
haftmann
parents: 25078
diff changeset
   253
    assume ab: "times a b = zero"
bfde2f8c0f63 partially localized
haftmann
parents: 25078
diff changeset
   254
    hence "zero = times (times (inverse a) (times a b)) (inverse b)"
bfde2f8c0f63 partially localized
haftmann
parents: 25078
diff changeset
   255
      by simp
bfde2f8c0f63 partially localized
haftmann
parents: 25078
diff changeset
   256
    also have "\<dots> = times (times (inverse a) a) (times b (inverse b))"
bfde2f8c0f63 partially localized
haftmann
parents: 25078
diff changeset
   257
      by (simp only: mult_assoc)
bfde2f8c0f63 partially localized
haftmann
parents: 25078
diff changeset
   258
    also have "\<dots> = one"
bfde2f8c0f63 partially localized
haftmann
parents: 25078
diff changeset
   259
      using a b by simp
bfde2f8c0f63 partially localized
haftmann
parents: 25078
diff changeset
   260
    finally show False
bfde2f8c0f63 partially localized
haftmann
parents: 25078
diff changeset
   261
      by simp
bfde2f8c0f63 partially localized
haftmann
parents: 25078
diff changeset
   262
  qed
bfde2f8c0f63 partially localized
haftmann
parents: 25078
diff changeset
   263
qed
bfde2f8c0f63 partially localized
haftmann
parents: 25078
diff changeset
   264
22987
550709aa8e66 instance division_ring < no_zero_divisors; clean up field instance proofs
huffman
parents: 22842
diff changeset
   265
class field = comm_ring_1 + inverse +
25062
af5ef0d4d655 global class syntax
haftmann
parents: 24748
diff changeset
   266
  assumes field_inverse:  "a \<noteq> 0 \<Longrightarrow> inverse a * a = 1"
af5ef0d4d655 global class syntax
haftmann
parents: 24748
diff changeset
   267
  assumes divide_inverse: "a / b = a * inverse b"
20496
23eb6034c06d added axclass division_ring (like field without commutativity; includes e.g. quaternions) and generalized some theorems from field to division_ring
huffman
parents: 19404
diff changeset
   268
23eb6034c06d added axclass division_ring (like field without commutativity; includes e.g. quaternions) and generalized some theorems from field to division_ring
huffman
parents: 19404
diff changeset
   269
instance field \<subseteq> division_ring
22987
550709aa8e66 instance division_ring < no_zero_divisors; clean up field instance proofs
huffman
parents: 22842
diff changeset
   270
proof
550709aa8e66 instance division_ring < no_zero_divisors; clean up field instance proofs
huffman
parents: 22842
diff changeset
   271
  fix a :: 'a
550709aa8e66 instance division_ring < no_zero_divisors; clean up field instance proofs
huffman
parents: 22842
diff changeset
   272
  assume "a \<noteq> 0"
550709aa8e66 instance division_ring < no_zero_divisors; clean up field instance proofs
huffman
parents: 22842
diff changeset
   273
  thus "inverse a * a = 1" by (rule field_inverse)
550709aa8e66 instance division_ring < no_zero_divisors; clean up field instance proofs
huffman
parents: 22842
diff changeset
   274
  thus "a * inverse a = 1" by (simp only: mult_commute)
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   275
qed
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   276
25152
bfde2f8c0f63 partially localized
haftmann
parents: 25078
diff changeset
   277
interpretation field \<subseteq> division_ring
bfde2f8c0f63 partially localized
haftmann
parents: 25078
diff changeset
   278
proof unfold_locales
bfde2f8c0f63 partially localized
haftmann
parents: 25078
diff changeset
   279
  fix a :: 'a
bfde2f8c0f63 partially localized
haftmann
parents: 25078
diff changeset
   280
  assume "a \<noteq> zero"
bfde2f8c0f63 partially localized
haftmann
parents: 25078
diff changeset
   281
  thus "times (inverse a) a = one" by (rule field_inverse)
bfde2f8c0f63 partially localized
haftmann
parents: 25078
diff changeset
   282
  thus "times a (inverse a) = one" by (simp only: mult_commute)
bfde2f8c0f63 partially localized
haftmann
parents: 25078
diff changeset
   283
qed
bfde2f8c0f63 partially localized
haftmann
parents: 25078
diff changeset
   284
bfde2f8c0f63 partially localized
haftmann
parents: 25078
diff changeset
   285
subclass (in field) idom by unfold_locales
20496
23eb6034c06d added axclass division_ring (like field without commutativity; includes e.g. quaternions) and generalized some theorems from field to division_ring
huffman
parents: 19404
diff changeset
   286
22390
378f34b1e380 now using "class"
haftmann
parents: 21328
diff changeset
   287
class division_by_zero = zero + inverse +
25062
af5ef0d4d655 global class syntax
haftmann
parents: 24748
diff changeset
   288
  assumes inverse_zero [simp]: "inverse 0 = 0"
14265
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   289
22390
378f34b1e380 now using "class"
haftmann
parents: 21328
diff changeset
   290
class mult_mono = times + zero + ord +
25062
af5ef0d4d655 global class syntax
haftmann
parents: 24748
diff changeset
   291
  assumes mult_left_mono: "a \<le> b \<Longrightarrow> 0 \<le> c \<Longrightarrow> c * a \<le> c * b"
af5ef0d4d655 global class syntax
haftmann
parents: 24748
diff changeset
   292
  assumes mult_right_mono: "a \<le> b \<Longrightarrow> 0 \<le> c \<Longrightarrow> a * c \<le> b * c"
14267
b963e9cee2a0 More refinements to Ring_and_Field and numerics. Conversion of Divides_lemmas
paulson
parents: 14266
diff changeset
   293
22390
378f34b1e380 now using "class"
haftmann
parents: 21328
diff changeset
   294
class pordered_semiring = mult_mono + semiring_0 + pordered_ab_semigroup_add 
21199
2d83f93c3580 * Added annihilation axioms ("x * 0 = 0") to axclass semiring_0.
krauss
parents: 20633
diff changeset
   295
22390
378f34b1e380 now using "class"
haftmann
parents: 21328
diff changeset
   296
class pordered_cancel_semiring = mult_mono + pordered_ab_semigroup_add
22987
550709aa8e66 instance division_ring < no_zero_divisors; clean up field instance proofs
huffman
parents: 22842
diff changeset
   297
  + semiring + comm_monoid_add + cancel_ab_semigroup_add
14268
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   298
14940
b9ab8babd8b3 Further development of matrix theory
obua
parents: 14770
diff changeset
   299
instance pordered_cancel_semiring \<subseteq> semiring_0_cancel ..
b9ab8babd8b3 Further development of matrix theory
obua
parents: 14770
diff changeset
   300
25152
bfde2f8c0f63 partially localized
haftmann
parents: 25078
diff changeset
   301
interpretation pordered_cancel_semiring \<subseteq> semiring_0_cancel by unfold_locales
bfde2f8c0f63 partially localized
haftmann
parents: 25078
diff changeset
   302
bfde2f8c0f63 partially localized
haftmann
parents: 25078
diff changeset
   303
subclass (in pordered_cancel_semiring) pordered_semiring by unfold_locales
21199
2d83f93c3580 * Added annihilation axioms ("x * 0 = 0") to axclass semiring_0.
krauss
parents: 20633
diff changeset
   304
23521
195fe3fe2831 added ordered_ring and ordered_semiring
obua
parents: 23496
diff changeset
   305
class ordered_semiring = semiring + comm_monoid_add + ordered_cancel_ab_semigroup_add + mult_mono
195fe3fe2831 added ordered_ring and ordered_semiring
obua
parents: 23496
diff changeset
   306
195fe3fe2831 added ordered_ring and ordered_semiring
obua
parents: 23496
diff changeset
   307
instance ordered_semiring \<subseteq> pordered_cancel_semiring ..
195fe3fe2831 added ordered_ring and ordered_semiring
obua
parents: 23496
diff changeset
   308
25152
bfde2f8c0f63 partially localized
haftmann
parents: 25078
diff changeset
   309
interpretation ordered_semiring \<subseteq> pordered_cancel_semiring by unfold_locales
bfde2f8c0f63 partially localized
haftmann
parents: 25078
diff changeset
   310
22390
378f34b1e380 now using "class"
haftmann
parents: 21328
diff changeset
   311
class ordered_semiring_strict = semiring + comm_monoid_add + ordered_cancel_ab_semigroup_add +
25062
af5ef0d4d655 global class syntax
haftmann
parents: 24748
diff changeset
   312
  assumes mult_strict_left_mono: "a < b \<Longrightarrow> 0 < c \<Longrightarrow> c * a < c * b"
af5ef0d4d655 global class syntax
haftmann
parents: 24748
diff changeset
   313
  assumes mult_strict_right_mono: "a < b \<Longrightarrow> 0 < c \<Longrightarrow> a * c < b * c"
14341
a09441bd4f1e Ring_and_Field now requires axiom add_left_imp_eq for semirings.
paulson
parents: 14334
diff changeset
   314
14940
b9ab8babd8b3 Further development of matrix theory
obua
parents: 14770
diff changeset
   315
instance ordered_semiring_strict \<subseteq> semiring_0_cancel ..
b9ab8babd8b3 Further development of matrix theory
obua
parents: 14770
diff changeset
   316
25152
bfde2f8c0f63 partially localized
haftmann
parents: 25078
diff changeset
   317
interpretation ordered_semiring_strict \<subseteq> semiring_0_cancel by unfold_locales
bfde2f8c0f63 partially localized
haftmann
parents: 25078
diff changeset
   318
23521
195fe3fe2831 added ordered_ring and ordered_semiring
obua
parents: 23496
diff changeset
   319
instance ordered_semiring_strict \<subseteq> ordered_semiring
23550
d4f1d6ef119c convert instance proofs to Isar style
huffman
parents: 23544
diff changeset
   320
proof
d4f1d6ef119c convert instance proofs to Isar style
huffman
parents: 23544
diff changeset
   321
  fix a b c :: 'a
d4f1d6ef119c convert instance proofs to Isar style
huffman
parents: 23544
diff changeset
   322
  assume A: "a \<le> b" "0 \<le> c"
d4f1d6ef119c convert instance proofs to Isar style
huffman
parents: 23544
diff changeset
   323
  from A show "c * a \<le> c * b"
d4f1d6ef119c convert instance proofs to Isar style
huffman
parents: 23544
diff changeset
   324
    unfolding order_le_less
d4f1d6ef119c convert instance proofs to Isar style
huffman
parents: 23544
diff changeset
   325
    using mult_strict_left_mono by auto
d4f1d6ef119c convert instance proofs to Isar style
huffman
parents: 23544
diff changeset
   326
  from A show "a * c \<le> b * c"
d4f1d6ef119c convert instance proofs to Isar style
huffman
parents: 23544
diff changeset
   327
    unfolding order_le_less
d4f1d6ef119c convert instance proofs to Isar style
huffman
parents: 23544
diff changeset
   328
    using mult_strict_right_mono by auto
d4f1d6ef119c convert instance proofs to Isar style
huffman
parents: 23544
diff changeset
   329
qed
14270
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   330
25152
bfde2f8c0f63 partially localized
haftmann
parents: 25078
diff changeset
   331
interpretation ordered_semiring_strict \<subseteq> ordered_semiring
bfde2f8c0f63 partially localized
haftmann
parents: 25078
diff changeset
   332
proof unfold_locales
bfde2f8c0f63 partially localized
haftmann
parents: 25078
diff changeset
   333
  fix a b c :: 'a
bfde2f8c0f63 partially localized
haftmann
parents: 25078
diff changeset
   334
  assume A: "less_eq a b" "less_eq zero c"
bfde2f8c0f63 partially localized
haftmann
parents: 25078
diff changeset
   335
  from A show "less_eq (times c a) (times c b)"
bfde2f8c0f63 partially localized
haftmann
parents: 25078
diff changeset
   336
    unfolding le_less  
bfde2f8c0f63 partially localized
haftmann
parents: 25078
diff changeset
   337
    using mult_strict_left_mono by (cases "c = zero") auto
bfde2f8c0f63 partially localized
haftmann
parents: 25078
diff changeset
   338
  from A show "less_eq (times a c) (times b c)"
bfde2f8c0f63 partially localized
haftmann
parents: 25078
diff changeset
   339
    unfolding le_less
bfde2f8c0f63 partially localized
haftmann
parents: 25078
diff changeset
   340
    using mult_strict_right_mono by (cases "c = zero") auto
bfde2f8c0f63 partially localized
haftmann
parents: 25078
diff changeset
   341
qed
bfde2f8c0f63 partially localized
haftmann
parents: 25078
diff changeset
   342
22390
378f34b1e380 now using "class"
haftmann
parents: 21328
diff changeset
   343
class mult_mono1 = times + zero + ord +
25062
af5ef0d4d655 global class syntax
haftmann
parents: 24748
diff changeset
   344
  assumes mult_mono: "a \<le> b \<Longrightarrow> 0 \<le> c \<Longrightarrow> c * a \<le> c * b"
14270
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   345
22390
378f34b1e380 now using "class"
haftmann
parents: 21328
diff changeset
   346
class pordered_comm_semiring = comm_semiring_0
378f34b1e380 now using "class"
haftmann
parents: 21328
diff changeset
   347
  + pordered_ab_semigroup_add + mult_mono1
14270
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   348
22390
378f34b1e380 now using "class"
haftmann
parents: 21328
diff changeset
   349
class pordered_cancel_comm_semiring = comm_semiring_0_cancel
378f34b1e380 now using "class"
haftmann
parents: 21328
diff changeset
   350
  + pordered_ab_semigroup_add + mult_mono1
25152
bfde2f8c0f63 partially localized
haftmann
parents: 25078
diff changeset
   351
bfde2f8c0f63 partially localized
haftmann
parents: 25078
diff changeset
   352
-- {*FIXME: continue localization here*}
bfde2f8c0f63 partially localized
haftmann
parents: 25078
diff changeset
   353
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   354
instance pordered_cancel_comm_semiring \<subseteq> pordered_comm_semiring ..
14270
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   355
22390
378f34b1e380 now using "class"
haftmann
parents: 21328
diff changeset
   356
class ordered_comm_semiring_strict = comm_semiring_0 + ordered_cancel_ab_semigroup_add +
25062
af5ef0d4d655 global class syntax
haftmann
parents: 24748
diff changeset
   357
  assumes mult_strict_mono: "a < b \<Longrightarrow> 0 < c \<Longrightarrow> c * a < c * b"
14265
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   358
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   359
instance pordered_comm_semiring \<subseteq> pordered_semiring
21199
2d83f93c3580 * Added annihilation axioms ("x * 0 = 0") to axclass semiring_0.
krauss
parents: 20633
diff changeset
   360
proof
2d83f93c3580 * Added annihilation axioms ("x * 0 = 0") to axclass semiring_0.
krauss
parents: 20633
diff changeset
   361
  fix a b c :: 'a
23550
d4f1d6ef119c convert instance proofs to Isar style
huffman
parents: 23544
diff changeset
   362
  assume "a \<le> b" "0 \<le> c"
d4f1d6ef119c convert instance proofs to Isar style
huffman
parents: 23544
diff changeset
   363
  thus "c * a \<le> c * b" by (rule mult_mono)
d4f1d6ef119c convert instance proofs to Isar style
huffman
parents: 23544
diff changeset
   364
  thus "a * c \<le> b * c" by (simp only: mult_commute)
21199
2d83f93c3580 * Added annihilation axioms ("x * 0 = 0") to axclass semiring_0.
krauss
parents: 20633
diff changeset
   365
qed
14265
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   366
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   367
instance pordered_cancel_comm_semiring \<subseteq> pordered_cancel_semiring ..
14265
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   368
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   369
instance ordered_comm_semiring_strict \<subseteq> ordered_semiring_strict
23550
d4f1d6ef119c convert instance proofs to Isar style
huffman
parents: 23544
diff changeset
   370
proof
d4f1d6ef119c convert instance proofs to Isar style
huffman
parents: 23544
diff changeset
   371
  fix a b c :: 'a
d4f1d6ef119c convert instance proofs to Isar style
huffman
parents: 23544
diff changeset
   372
  assume "a < b" "0 < c"
d4f1d6ef119c convert instance proofs to Isar style
huffman
parents: 23544
diff changeset
   373
  thus "c * a < c * b" by (rule mult_strict_mono)
d4f1d6ef119c convert instance proofs to Isar style
huffman
parents: 23544
diff changeset
   374
  thus "a * c < b * c" by (simp only: mult_commute)
d4f1d6ef119c convert instance proofs to Isar style
huffman
parents: 23544
diff changeset
   375
qed
14272
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
   376
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   377
instance ordered_comm_semiring_strict \<subseteq> pordered_cancel_comm_semiring
23550
d4f1d6ef119c convert instance proofs to Isar style
huffman
parents: 23544
diff changeset
   378
proof
d4f1d6ef119c convert instance proofs to Isar style
huffman
parents: 23544
diff changeset
   379
  fix a b c :: 'a
d4f1d6ef119c convert instance proofs to Isar style
huffman
parents: 23544
diff changeset
   380
  assume "a \<le> b" "0 \<le> c"
d4f1d6ef119c convert instance proofs to Isar style
huffman
parents: 23544
diff changeset
   381
  thus "c * a \<le> c * b"
d4f1d6ef119c convert instance proofs to Isar style
huffman
parents: 23544
diff changeset
   382
    unfolding order_le_less
d4f1d6ef119c convert instance proofs to Isar style
huffman
parents: 23544
diff changeset
   383
    using mult_strict_mono by auto
d4f1d6ef119c convert instance proofs to Isar style
huffman
parents: 23544
diff changeset
   384
qed
14272
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
   385
22390
378f34b1e380 now using "class"
haftmann
parents: 21328
diff changeset
   386
class pordered_ring = ring + pordered_cancel_semiring 
14270
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   387
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   388
instance pordered_ring \<subseteq> pordered_ab_group_add ..
14270
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   389
22452
8a86fd2a1bf0 adjusted to new lattice theory developement in Lattices.thy / FixedPoint.thy
haftmann
parents: 22422
diff changeset
   390
class lordered_ring = pordered_ring + lordered_ab_group_abs
14270
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   391
14940
b9ab8babd8b3 Further development of matrix theory
obua
parents: 14770
diff changeset
   392
instance lordered_ring \<subseteq> lordered_ab_group_meet ..
b9ab8babd8b3 Further development of matrix theory
obua
parents: 14770
diff changeset
   393
b9ab8babd8b3 Further development of matrix theory
obua
parents: 14770
diff changeset
   394
instance lordered_ring \<subseteq> lordered_ab_group_join ..
b9ab8babd8b3 Further development of matrix theory
obua
parents: 14770
diff changeset
   395
23879
4776af8be741 split class abs from class minus
haftmann
parents: 23550
diff changeset
   396
class abs_if = minus + ord + zero + abs +
25062
af5ef0d4d655 global class syntax
haftmann
parents: 24748
diff changeset
   397
  assumes abs_if: "abs a = (if a < 0 then (uminus a) else a)"
14270
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   398
24506
020db6ec334a final(?) iteration of sgn saga.
nipkow
parents: 24491
diff changeset
   399
class sgn_if = sgn + zero + one + minus + ord +
25062
af5ef0d4d655 global class syntax
haftmann
parents: 24748
diff changeset
   400
  assumes sgn_if: "sgn x = (if x = 0 then 0 else if 0 < x then 1 else uminus 1)"
24506
020db6ec334a final(?) iteration of sgn saga.
nipkow
parents: 24491
diff changeset
   401
23521
195fe3fe2831 added ordered_ring and ordered_semiring
obua
parents: 23496
diff changeset
   402
(* The "strict" suffix can be seen as describing the combination of ordered_ring and no_zero_divisors.
195fe3fe2831 added ordered_ring and ordered_semiring
obua
parents: 23496
diff changeset
   403
   Basically, ordered_ring + no_zero_divisors = ordered_ring_strict.
195fe3fe2831 added ordered_ring and ordered_semiring
obua
parents: 23496
diff changeset
   404
 *)
195fe3fe2831 added ordered_ring and ordered_semiring
obua
parents: 23496
diff changeset
   405
class ordered_ring = ring + ordered_semiring + lordered_ab_group + abs_if
14270
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   406
23550
d4f1d6ef119c convert instance proofs to Isar style
huffman
parents: 23544
diff changeset
   407
instance ordered_ring \<subseteq> lordered_ring
d4f1d6ef119c convert instance proofs to Isar style
huffman
parents: 23544
diff changeset
   408
proof
d4f1d6ef119c convert instance proofs to Isar style
huffman
parents: 23544
diff changeset
   409
  fix x :: 'a
d4f1d6ef119c convert instance proofs to Isar style
huffman
parents: 23544
diff changeset
   410
  show "\<bar>x\<bar> = sup x (- x)"
d4f1d6ef119c convert instance proofs to Isar style
huffman
parents: 23544
diff changeset
   411
    by (simp only: abs_if sup_eq_if)
d4f1d6ef119c convert instance proofs to Isar style
huffman
parents: 23544
diff changeset
   412
qed
23521
195fe3fe2831 added ordered_ring and ordered_semiring
obua
parents: 23496
diff changeset
   413
24506
020db6ec334a final(?) iteration of sgn saga.
nipkow
parents: 24491
diff changeset
   414
class ordered_ring_strict =
020db6ec334a final(?) iteration of sgn saga.
nipkow
parents: 24491
diff changeset
   415
  ring + ordered_semiring_strict + lordered_ab_group + abs_if
23521
195fe3fe2831 added ordered_ring and ordered_semiring
obua
parents: 23496
diff changeset
   416
195fe3fe2831 added ordered_ring and ordered_semiring
obua
parents: 23496
diff changeset
   417
instance ordered_ring_strict \<subseteq> ordered_ring ..
14270
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   418
22390
378f34b1e380 now using "class"
haftmann
parents: 21328
diff changeset
   419
class pordered_comm_ring = comm_ring + pordered_comm_semiring
14270
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   420
23527
c1d2fa4b76df instance pordered_comm_ring < pordered_ring
huffman
parents: 23521
diff changeset
   421
instance pordered_comm_ring \<subseteq> pordered_ring ..
c1d2fa4b76df instance pordered_comm_ring < pordered_ring
huffman
parents: 23521
diff changeset
   422
23073
d810dc04b96d add missing instance declarations
huffman
parents: 22993
diff changeset
   423
instance pordered_comm_ring \<subseteq> pordered_cancel_comm_semiring ..
d810dc04b96d add missing instance declarations
huffman
parents: 22993
diff changeset
   424
22390
378f34b1e380 now using "class"
haftmann
parents: 21328
diff changeset
   425
class ordered_semidom = comm_semiring_1_cancel + ordered_comm_semiring_strict +
378f34b1e380 now using "class"
haftmann
parents: 21328
diff changeset
   426
  (*previously ordered_semiring*)
25062
af5ef0d4d655 global class syntax
haftmann
parents: 24748
diff changeset
   427
  assumes zero_less_one [simp]: "0 < 1"
14270
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   428
24422
c0b5ff9e9e4d moved class dense_linear_order to Orderings.thy
haftmann
parents: 24286
diff changeset
   429
lemma pos_add_strict:
c0b5ff9e9e4d moved class dense_linear_order to Orderings.thy
haftmann
parents: 24286
diff changeset
   430
  fixes a b c :: "'a\<Colon>ordered_semidom"
c0b5ff9e9e4d moved class dense_linear_order to Orderings.thy
haftmann
parents: 24286
diff changeset
   431
  shows "0 < a \<Longrightarrow> b < c \<Longrightarrow> b < a + c"
c0b5ff9e9e4d moved class dense_linear_order to Orderings.thy
haftmann
parents: 24286
diff changeset
   432
  using add_strict_mono [of 0 a b c] by simp
c0b5ff9e9e4d moved class dense_linear_order to Orderings.thy
haftmann
parents: 24286
diff changeset
   433
24506
020db6ec334a final(?) iteration of sgn saga.
nipkow
parents: 24491
diff changeset
   434
class ordered_idom =
020db6ec334a final(?) iteration of sgn saga.
nipkow
parents: 24491
diff changeset
   435
  comm_ring_1 +
020db6ec334a final(?) iteration of sgn saga.
nipkow
parents: 24491
diff changeset
   436
  ordered_comm_semiring_strict +
020db6ec334a final(?) iteration of sgn saga.
nipkow
parents: 24491
diff changeset
   437
  lordered_ab_group +
020db6ec334a final(?) iteration of sgn saga.
nipkow
parents: 24491
diff changeset
   438
  abs_if + sgn_if
22390
378f34b1e380 now using "class"
haftmann
parents: 21328
diff changeset
   439
  (*previously ordered_ring*)
14270
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   440
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   441
instance ordered_idom \<subseteq> ordered_ring_strict ..
14272
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
   442
23073
d810dc04b96d add missing instance declarations
huffman
parents: 22993
diff changeset
   443
instance ordered_idom \<subseteq> pordered_comm_ring ..
d810dc04b96d add missing instance declarations
huffman
parents: 22993
diff changeset
   444
22390
378f34b1e380 now using "class"
haftmann
parents: 21328
diff changeset
   445
class ordered_field = field + ordered_idom
14272
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
   446
24515
d4dc5dc2db98 linorder_neqE_ordered_idom: proper proof, avoid illegal schematic type vars;
wenzelm
parents: 24506
diff changeset
   447
lemma linorder_neqE_ordered_idom:
d4dc5dc2db98 linorder_neqE_ordered_idom: proper proof, avoid illegal schematic type vars;
wenzelm
parents: 24506
diff changeset
   448
  fixes x y :: "'a :: ordered_idom"
d4dc5dc2db98 linorder_neqE_ordered_idom: proper proof, avoid illegal schematic type vars;
wenzelm
parents: 24506
diff changeset
   449
  assumes "x \<noteq> y" obtains "x < y" | "y < x"
d4dc5dc2db98 linorder_neqE_ordered_idom: proper proof, avoid illegal schematic type vars;
wenzelm
parents: 24506
diff changeset
   450
  using assms by (rule linorder_neqE)
15923
01d5d0c1c078 fixed lin.arith
nipkow
parents: 15769
diff changeset
   451
14272
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
   452
lemma eq_add_iff1:
23477
f4b83f03cac9 tuned and renamed group_eq_simps and ring_eq_simps
nipkow
parents: 23413
diff changeset
   453
  "(a*e + c = b*e + d) = ((a-b)*e + c = (d::'a::ring))"
f4b83f03cac9 tuned and renamed group_eq_simps and ring_eq_simps
nipkow
parents: 23413
diff changeset
   454
by (simp add: ring_simps)
14272
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
   455
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
   456
lemma eq_add_iff2:
23477
f4b83f03cac9 tuned and renamed group_eq_simps and ring_eq_simps
nipkow
parents: 23413
diff changeset
   457
  "(a*e + c = b*e + d) = (c = (b-a)*e + (d::'a::ring))"
f4b83f03cac9 tuned and renamed group_eq_simps and ring_eq_simps
nipkow
parents: 23413
diff changeset
   458
by (simp add: ring_simps)
14272
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
   459
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
   460
lemma less_add_iff1:
23477
f4b83f03cac9 tuned and renamed group_eq_simps and ring_eq_simps
nipkow
parents: 23413
diff changeset
   461
  "(a*e + c < b*e + d) = ((a-b)*e + c < (d::'a::pordered_ring))"
f4b83f03cac9 tuned and renamed group_eq_simps and ring_eq_simps
nipkow
parents: 23413
diff changeset
   462
by (simp add: ring_simps)
14272
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
   463
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
   464
lemma less_add_iff2:
23477
f4b83f03cac9 tuned and renamed group_eq_simps and ring_eq_simps
nipkow
parents: 23413
diff changeset
   465
  "(a*e + c < b*e + d) = (c < (b-a)*e + (d::'a::pordered_ring))"
f4b83f03cac9 tuned and renamed group_eq_simps and ring_eq_simps
nipkow
parents: 23413
diff changeset
   466
by (simp add: ring_simps)
14272
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
   467
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
   468
lemma le_add_iff1:
23477
f4b83f03cac9 tuned and renamed group_eq_simps and ring_eq_simps
nipkow
parents: 23413
diff changeset
   469
  "(a*e + c \<le> b*e + d) = ((a-b)*e + c \<le> (d::'a::pordered_ring))"
f4b83f03cac9 tuned and renamed group_eq_simps and ring_eq_simps
nipkow
parents: 23413
diff changeset
   470
by (simp add: ring_simps)
14272
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
   471
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
   472
lemma le_add_iff2:
23477
f4b83f03cac9 tuned and renamed group_eq_simps and ring_eq_simps
nipkow
parents: 23413
diff changeset
   473
  "(a*e + c \<le> b*e + d) = (c \<le> (b-a)*e + (d::'a::pordered_ring))"
f4b83f03cac9 tuned and renamed group_eq_simps and ring_eq_simps
nipkow
parents: 23413
diff changeset
   474
by (simp add: ring_simps)
14272
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
   475
23389
aaca6a8e5414 tuned proofs: avoid implicit prems;
wenzelm
parents: 23326
diff changeset
   476
14270
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   477
subsection {* Ordering Rules for Multiplication *}
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   478
14348
744c868ee0b7 Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents: 14341
diff changeset
   479
lemma mult_left_le_imp_le:
23477
f4b83f03cac9 tuned and renamed group_eq_simps and ring_eq_simps
nipkow
parents: 23413
diff changeset
   480
  "[|c*a \<le> c*b; 0 < c|] ==> a \<le> (b::'a::ordered_semiring_strict)"
f4b83f03cac9 tuned and renamed group_eq_simps and ring_eq_simps
nipkow
parents: 23413
diff changeset
   481
by (force simp add: mult_strict_left_mono linorder_not_less [symmetric])
14348
744c868ee0b7 Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents: 14341
diff changeset
   482
 
744c868ee0b7 Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents: 14341
diff changeset
   483
lemma mult_right_le_imp_le:
23477
f4b83f03cac9 tuned and renamed group_eq_simps and ring_eq_simps
nipkow
parents: 23413
diff changeset
   484
  "[|a*c \<le> b*c; 0 < c|] ==> a \<le> (b::'a::ordered_semiring_strict)"
f4b83f03cac9 tuned and renamed group_eq_simps and ring_eq_simps
nipkow
parents: 23413
diff changeset
   485
by (force simp add: mult_strict_right_mono linorder_not_less [symmetric])
14348
744c868ee0b7 Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents: 14341
diff changeset
   486
744c868ee0b7 Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents: 14341
diff changeset
   487
lemma mult_left_less_imp_less:
23521
195fe3fe2831 added ordered_ring and ordered_semiring
obua
parents: 23496
diff changeset
   488
  "[|c*a < c*b; 0 \<le> c|] ==> a < (b::'a::ordered_semiring)"
23477
f4b83f03cac9 tuned and renamed group_eq_simps and ring_eq_simps
nipkow
parents: 23413
diff changeset
   489
by (force simp add: mult_left_mono linorder_not_le [symmetric])
14348
744c868ee0b7 Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents: 14341
diff changeset
   490
 
744c868ee0b7 Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents: 14341
diff changeset
   491
lemma mult_right_less_imp_less:
23521
195fe3fe2831 added ordered_ring and ordered_semiring
obua
parents: 23496
diff changeset
   492
  "[|a*c < b*c; 0 \<le> c|] ==> a < (b::'a::ordered_semiring)"
23477
f4b83f03cac9 tuned and renamed group_eq_simps and ring_eq_simps
nipkow
parents: 23413
diff changeset
   493
by (force simp add: mult_right_mono linorder_not_le [symmetric])
14348
744c868ee0b7 Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents: 14341
diff changeset
   494
14265
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   495
lemma mult_strict_left_mono_neg:
23477
f4b83f03cac9 tuned and renamed group_eq_simps and ring_eq_simps
nipkow
parents: 23413
diff changeset
   496
  "[|b < a; c < 0|] ==> c * a < c * (b::'a::ordered_ring_strict)"
14265
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   497
apply (drule mult_strict_left_mono [of _ _ "-c"])
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   498
apply (simp_all add: minus_mult_left [symmetric]) 
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   499
done
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   500
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   501
lemma mult_left_mono_neg:
23477
f4b83f03cac9 tuned and renamed group_eq_simps and ring_eq_simps
nipkow
parents: 23413
diff changeset
   502
  "[|b \<le> a; c \<le> 0|] ==> c * a \<le>  c * (b::'a::pordered_ring)"
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   503
apply (drule mult_left_mono [of _ _ "-c"])
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   504
apply (simp_all add: minus_mult_left [symmetric]) 
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   505
done
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   506
14265
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   507
lemma mult_strict_right_mono_neg:
23477
f4b83f03cac9 tuned and renamed group_eq_simps and ring_eq_simps
nipkow
parents: 23413
diff changeset
   508
  "[|b < a; c < 0|] ==> a * c < b * (c::'a::ordered_ring_strict)"
14265
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   509
apply (drule mult_strict_right_mono [of _ _ "-c"])
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   510
apply (simp_all add: minus_mult_right [symmetric]) 
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   511
done
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   512
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   513
lemma mult_right_mono_neg:
23477
f4b83f03cac9 tuned and renamed group_eq_simps and ring_eq_simps
nipkow
parents: 23413
diff changeset
   514
  "[|b \<le> a; c \<le> 0|] ==> a * c \<le>  (b::'a::pordered_ring) * c"
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   515
apply (drule mult_right_mono [of _ _ "-c"])
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   516
apply (simp)
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   517
apply (simp_all add: minus_mult_right [symmetric]) 
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   518
done
14265
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   519
23389
aaca6a8e5414 tuned proofs: avoid implicit prems;
wenzelm
parents: 23326
diff changeset
   520
14265
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   521
subsection{* Products of Signs *}
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   522
16775
c1b87ef4a1c3 added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
avigad
parents: 16568
diff changeset
   523
lemma mult_pos_pos: "[| (0::'a::ordered_semiring_strict) < a; 0 < b |] ==> 0 < a*b"
14265
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   524
by (drule mult_strict_left_mono [of 0 b], auto)
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   525
16775
c1b87ef4a1c3 added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
avigad
parents: 16568
diff changeset
   526
lemma mult_nonneg_nonneg: "[| (0::'a::pordered_cancel_semiring) \<le> a; 0 \<le> b |] ==> 0 \<le> a*b"
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   527
by (drule mult_left_mono [of 0 b], auto)
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   528
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   529
lemma mult_pos_neg: "[| (0::'a::ordered_semiring_strict) < a; b < 0 |] ==> a*b < 0"
14265
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   530
by (drule mult_strict_left_mono [of b 0], auto)
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   531
16775
c1b87ef4a1c3 added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
avigad
parents: 16568
diff changeset
   532
lemma mult_nonneg_nonpos: "[| (0::'a::pordered_cancel_semiring) \<le> a; b \<le> 0 |] ==> a*b \<le> 0"
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   533
by (drule mult_left_mono [of b 0], auto)
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   534
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   535
lemma mult_pos_neg2: "[| (0::'a::ordered_semiring_strict) < a; b < 0 |] ==> b*a < 0" 
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   536
by (drule mult_strict_right_mono[of b 0], auto)
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   537
16775
c1b87ef4a1c3 added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
avigad
parents: 16568
diff changeset
   538
lemma mult_nonneg_nonpos2: "[| (0::'a::pordered_cancel_semiring) \<le> a; b \<le> 0 |] ==> b*a \<le> 0" 
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   539
by (drule mult_right_mono[of b 0], auto)
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   540
16775
c1b87ef4a1c3 added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
avigad
parents: 16568
diff changeset
   541
lemma mult_neg_neg: "[| a < (0::'a::ordered_ring_strict); b < 0 |] ==> 0 < a*b"
14265
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   542
by (drule mult_strict_right_mono_neg, auto)
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   543
16775
c1b87ef4a1c3 added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
avigad
parents: 16568
diff changeset
   544
lemma mult_nonpos_nonpos: "[| a \<le> (0::'a::pordered_ring); b \<le> 0 |] ==> 0 \<le> a*b"
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   545
by (drule mult_right_mono_neg[of a 0 b ], auto)
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   546
14341
a09441bd4f1e Ring_and_Field now requires axiom add_left_imp_eq for semirings.
paulson
parents: 14334
diff changeset
   547
lemma zero_less_mult_pos:
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   548
     "[| 0 < a*b; 0 < a|] ==> 0 < (b::'a::ordered_semiring_strict)"
21328
73bb86d0f483 dropped Inductive dependency
haftmann
parents: 21258
diff changeset
   549
apply (cases "b\<le>0") 
14265
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   550
 apply (auto simp add: order_le_less linorder_not_less)
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   551
apply (drule_tac mult_pos_neg [of a b]) 
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   552
 apply (auto dest: order_less_not_sym)
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   553
done
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   554
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   555
lemma zero_less_mult_pos2:
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   556
     "[| 0 < b*a; 0 < a|] ==> 0 < (b::'a::ordered_semiring_strict)"
21328
73bb86d0f483 dropped Inductive dependency
haftmann
parents: 21258
diff changeset
   557
apply (cases "b\<le>0") 
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   558
 apply (auto simp add: order_le_less linorder_not_less)
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   559
apply (drule_tac mult_pos_neg2 [of a b]) 
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   560
 apply (auto dest: order_less_not_sym)
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   561
done
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   562
14265
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   563
lemma zero_less_mult_iff:
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   564
     "((0::'a::ordered_ring_strict) < a*b) = (0 < a & 0 < b | a < 0 & b < 0)"
16775
c1b87ef4a1c3 added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
avigad
parents: 16568
diff changeset
   565
apply (auto simp add: order_le_less linorder_not_less mult_pos_pos 
c1b87ef4a1c3 added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
avigad
parents: 16568
diff changeset
   566
  mult_neg_neg)
14265
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   567
apply (blast dest: zero_less_mult_pos) 
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   568
apply (blast dest: zero_less_mult_pos2)
14265
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   569
done
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   570
22990
775e9de3db48 added classes ring_no_zero_divisors and dom (non-commutative version of idom);
huffman
parents: 22987
diff changeset
   571
lemma mult_eq_0_iff [simp]:
775e9de3db48 added classes ring_no_zero_divisors and dom (non-commutative version of idom);
huffman
parents: 22987
diff changeset
   572
  fixes a b :: "'a::ring_no_zero_divisors"
775e9de3db48 added classes ring_no_zero_divisors and dom (non-commutative version of idom);
huffman
parents: 22987
diff changeset
   573
  shows "(a * b = 0) = (a = 0 \<or> b = 0)"
775e9de3db48 added classes ring_no_zero_divisors and dom (non-commutative version of idom);
huffman
parents: 22987
diff changeset
   574
by (cases "a = 0 \<or> b = 0", auto dest: no_zero_divisors)
775e9de3db48 added classes ring_no_zero_divisors and dom (non-commutative version of idom);
huffman
parents: 22987
diff changeset
   575
775e9de3db48 added classes ring_no_zero_divisors and dom (non-commutative version of idom);
huffman
parents: 22987
diff changeset
   576
instance ordered_ring_strict \<subseteq> ring_no_zero_divisors
775e9de3db48 added classes ring_no_zero_divisors and dom (non-commutative version of idom);
huffman
parents: 22987
diff changeset
   577
apply intro_classes
14265
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   578
apply (auto simp add: linorder_not_less order_le_less linorder_neq_iff)
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   579
apply (force dest: mult_strict_right_mono_neg mult_strict_right_mono)+
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   580
done
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   581
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   582
lemma zero_le_mult_iff:
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   583
     "((0::'a::ordered_ring_strict) \<le> a*b) = (0 \<le> a & 0 \<le> b | a \<le> 0 & b \<le> 0)"
14265
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   584
by (auto simp add: eq_commute [of 0] order_le_less linorder_not_less
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   585
                   zero_less_mult_iff)
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   586
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   587
lemma mult_less_0_iff:
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   588
     "(a*b < (0::'a::ordered_ring_strict)) = (0 < a & b < 0 | a < 0 & 0 < b)"
14265
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   589
apply (insert zero_less_mult_iff [of "-a" b]) 
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   590
apply (force simp add: minus_mult_left[symmetric]) 
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   591
done
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   592
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   593
lemma mult_le_0_iff:
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   594
     "(a*b \<le> (0::'a::ordered_ring_strict)) = (0 \<le> a & b \<le> 0 | a \<le> 0 & 0 \<le> b)"
14265
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   595
apply (insert zero_le_mult_iff [of "-a" b]) 
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   596
apply (force simp add: minus_mult_left[symmetric]) 
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   597
done
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   598
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   599
lemma split_mult_pos_le: "(0 \<le> a & 0 \<le> b) | (a \<le> 0 & b \<le> 0) \<Longrightarrow> 0 \<le> a * (b::_::pordered_ring)"
16775
c1b87ef4a1c3 added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
avigad
parents: 16568
diff changeset
   600
by (auto simp add: mult_nonneg_nonneg mult_nonpos_nonpos)
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   601
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   602
lemma split_mult_neg_le: "(0 \<le> a & b \<le> 0) | (a \<le> 0 & 0 \<le> b) \<Longrightarrow> a * b \<le> (0::_::pordered_cancel_semiring)" 
16775
c1b87ef4a1c3 added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
avigad
parents: 16568
diff changeset
   603
by (auto simp add: mult_nonneg_nonpos mult_nonneg_nonpos2)
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   604
23095
45f10b70e891 Squared things out.
obua
parents: 23073
diff changeset
   605
lemma zero_le_square[simp]: "(0::'a::ordered_ring_strict) \<le> a*a"
45f10b70e891 Squared things out.
obua
parents: 23073
diff changeset
   606
by (simp add: zero_le_mult_iff linorder_linear)
45f10b70e891 Squared things out.
obua
parents: 23073
diff changeset
   607
45f10b70e891 Squared things out.
obua
parents: 23073
diff changeset
   608
lemma not_square_less_zero[simp]: "\<not> (a * a < (0::'a::ordered_ring_strict))"
45f10b70e891 Squared things out.
obua
parents: 23073
diff changeset
   609
by (simp add: not_less)
14265
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   610
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   611
text{*Proving axiom @{text zero_less_one} makes all @{text ordered_semidom}
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   612
      theorems available to members of @{term ordered_idom} *}
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   613
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   614
instance ordered_idom \<subseteq> ordered_semidom
14421
ee97b6463cb4 new Ring_and_Field hierarchy, eliminating redundant axioms
paulson
parents: 14398
diff changeset
   615
proof
ee97b6463cb4 new Ring_and_Field hierarchy, eliminating redundant axioms
paulson
parents: 14398
diff changeset
   616
  have "(0::'a) \<le> 1*1" by (rule zero_le_square)
14430
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14421
diff changeset
   617
  thus "(0::'a) < 1" by (simp add: order_le_less) 
14421
ee97b6463cb4 new Ring_and_Field hierarchy, eliminating redundant axioms
paulson
parents: 14398
diff changeset
   618
qed
ee97b6463cb4 new Ring_and_Field hierarchy, eliminating redundant axioms
paulson
parents: 14398
diff changeset
   619
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   620
instance ordered_idom \<subseteq> idom ..
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   621
14387
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14377
diff changeset
   622
text{*All three types of comparision involving 0 and 1 are covered.*}
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14377
diff changeset
   623
17085
5b57f995a179 more simprules now have names
paulson
parents: 16775
diff changeset
   624
lemmas one_neq_zero = zero_neq_one [THEN not_sym]
5b57f995a179 more simprules now have names
paulson
parents: 16775
diff changeset
   625
declare one_neq_zero [simp]
14387
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14377
diff changeset
   626
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   627
lemma zero_le_one [simp]: "(0::'a::ordered_semidom) \<le> 1"
14268
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   628
  by (rule zero_less_one [THEN order_less_imp_le]) 
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   629
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   630
lemma not_one_le_zero [simp]: "~ (1::'a::ordered_semidom) \<le> 0"
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   631
by (simp add: linorder_not_le) 
14387
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14377
diff changeset
   632
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   633
lemma not_one_less_zero [simp]: "~ (1::'a::ordered_semidom) < 0"
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   634
by (simp add: linorder_not_less) 
14268
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   635
23389
aaca6a8e5414 tuned proofs: avoid implicit prems;
wenzelm
parents: 23326
diff changeset
   636
14268
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   637
subsection{*More Monotonicity*}
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   638
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   639
text{*Strict monotonicity in both arguments*}
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   640
lemma mult_strict_mono:
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   641
     "[|a<b; c<d; 0<b; 0\<le>c|] ==> a * c < b * (d::'a::ordered_semiring_strict)"
21328
73bb86d0f483 dropped Inductive dependency
haftmann
parents: 21258
diff changeset
   642
apply (cases "c=0")
16775
c1b87ef4a1c3 added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
avigad
parents: 16568
diff changeset
   643
 apply (simp add: mult_pos_pos) 
14268
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   644
apply (erule mult_strict_right_mono [THEN order_less_trans])
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   645
 apply (force simp add: order_le_less) 
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   646
apply (erule mult_strict_left_mono, assumption)
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   647
done
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   648
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   649
text{*This weaker variant has more natural premises*}
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   650
lemma mult_strict_mono':
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   651
     "[| a<b; c<d; 0 \<le> a; 0 \<le> c|] ==> a * c < b * (d::'a::ordered_semiring_strict)"
14268
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   652
apply (rule mult_strict_mono)
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   653
apply (blast intro: order_le_less_trans)+
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   654
done
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   655
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   656
lemma mult_mono:
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   657
     "[|a \<le> b; c \<le> d; 0 \<le> b; 0 \<le> c|] 
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   658
      ==> a * c  \<le>  b * (d::'a::pordered_semiring)"
14268
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   659
apply (erule mult_right_mono [THEN order_trans], assumption)
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   660
apply (erule mult_left_mono, assumption)
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   661
done
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   662
21258
62f25a96f0c1 added lemma mult_mono'
huffman
parents: 21199
diff changeset
   663
lemma mult_mono':
62f25a96f0c1 added lemma mult_mono'
huffman
parents: 21199
diff changeset
   664
     "[|a \<le> b; c \<le> d; 0 \<le> a; 0 \<le> c|] 
62f25a96f0c1 added lemma mult_mono'
huffman
parents: 21199
diff changeset
   665
      ==> a * c  \<le>  b * (d::'a::pordered_semiring)"
62f25a96f0c1 added lemma mult_mono'
huffman
parents: 21199
diff changeset
   666
apply (rule mult_mono)
62f25a96f0c1 added lemma mult_mono'
huffman
parents: 21199
diff changeset
   667
apply (fast intro: order_trans)+
62f25a96f0c1 added lemma mult_mono'
huffman
parents: 21199
diff changeset
   668
done
62f25a96f0c1 added lemma mult_mono'
huffman
parents: 21199
diff changeset
   669
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   670
lemma less_1_mult: "[| 1 < m; 1 < n |] ==> 1 < m*(n::'a::ordered_semidom)"
14387
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14377
diff changeset
   671
apply (insert mult_strict_mono [of 1 m 1 n]) 
14430
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14421
diff changeset
   672
apply (simp add:  order_less_trans [OF zero_less_one]) 
14387
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14377
diff changeset
   673
done
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14377
diff changeset
   674
16775
c1b87ef4a1c3 added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
avigad
parents: 16568
diff changeset
   675
lemma mult_less_le_imp_less: "(a::'a::ordered_semiring_strict) < b ==>
c1b87ef4a1c3 added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
avigad
parents: 16568
diff changeset
   676
    c <= d ==> 0 <= a ==> 0 < c ==> a * c < b * d"
c1b87ef4a1c3 added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
avigad
parents: 16568
diff changeset
   677
  apply (subgoal_tac "a * c < b * c")
c1b87ef4a1c3 added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
avigad
parents: 16568
diff changeset
   678
  apply (erule order_less_le_trans)
c1b87ef4a1c3 added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
avigad
parents: 16568
diff changeset
   679
  apply (erule mult_left_mono)
c1b87ef4a1c3 added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
avigad
parents: 16568
diff changeset
   680
  apply simp
c1b87ef4a1c3 added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
avigad
parents: 16568
diff changeset
   681
  apply (erule mult_strict_right_mono)
c1b87ef4a1c3 added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
avigad
parents: 16568
diff changeset
   682
  apply assumption
c1b87ef4a1c3 added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
avigad
parents: 16568
diff changeset
   683
done
c1b87ef4a1c3 added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
avigad
parents: 16568
diff changeset
   684
c1b87ef4a1c3 added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
avigad
parents: 16568
diff changeset
   685
lemma mult_le_less_imp_less: "(a::'a::ordered_semiring_strict) <= b ==>
c1b87ef4a1c3 added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
avigad
parents: 16568
diff changeset
   686
    c < d ==> 0 < a ==> 0 <= c ==> a * c < b * d"
c1b87ef4a1c3 added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
avigad
parents: 16568
diff changeset
   687
  apply (subgoal_tac "a * c <= b * c")
c1b87ef4a1c3 added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
avigad
parents: 16568
diff changeset
   688
  apply (erule order_le_less_trans)
c1b87ef4a1c3 added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
avigad
parents: 16568
diff changeset
   689
  apply (erule mult_strict_left_mono)
c1b87ef4a1c3 added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
avigad
parents: 16568
diff changeset
   690
  apply simp
c1b87ef4a1c3 added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
avigad
parents: 16568
diff changeset
   691
  apply (erule mult_right_mono)
c1b87ef4a1c3 added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
avigad
parents: 16568
diff changeset
   692
  apply simp
c1b87ef4a1c3 added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
avigad
parents: 16568
diff changeset
   693
done
c1b87ef4a1c3 added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
avigad
parents: 16568
diff changeset
   694
23389
aaca6a8e5414 tuned proofs: avoid implicit prems;
wenzelm
parents: 23326
diff changeset
   695
14268
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   696
subsection{*Cancellation Laws for Relationships With a Common Factor*}
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   697
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   698
text{*Cancellation laws for @{term "c*a < c*b"} and @{term "a*c < b*c"},
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   699
   also with the relations @{text "\<le>"} and equality.*}
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   700
15234
ec91a90c604e simplification tweaks for better arithmetic reasoning
paulson
parents: 15229
diff changeset
   701
text{*These ``disjunction'' versions produce two cases when the comparison is
ec91a90c604e simplification tweaks for better arithmetic reasoning
paulson
parents: 15229
diff changeset
   702
 an assumption, but effectively four when the comparison is a goal.*}
ec91a90c604e simplification tweaks for better arithmetic reasoning
paulson
parents: 15229
diff changeset
   703
ec91a90c604e simplification tweaks for better arithmetic reasoning
paulson
parents: 15229
diff changeset
   704
lemma mult_less_cancel_right_disj:
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   705
    "(a*c < b*c) = ((0 < c & a < b) | (c < 0 & b < (a::'a::ordered_ring_strict)))"
21328
73bb86d0f483 dropped Inductive dependency
haftmann
parents: 21258
diff changeset
   706
apply (cases "c = 0")
14268
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   707
apply (auto simp add: linorder_neq_iff mult_strict_right_mono 
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   708
                      mult_strict_right_mono_neg)
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   709
apply (auto simp add: linorder_not_less 
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   710
                      linorder_not_le [symmetric, of "a*c"]
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   711
                      linorder_not_le [symmetric, of a])
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   712
apply (erule_tac [!] notE)
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   713
apply (auto simp add: order_less_imp_le mult_right_mono 
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   714
                      mult_right_mono_neg)
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   715
done
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   716
15234
ec91a90c604e simplification tweaks for better arithmetic reasoning
paulson
parents: 15229
diff changeset
   717
lemma mult_less_cancel_left_disj:
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   718
    "(c*a < c*b) = ((0 < c & a < b) | (c < 0 & b < (a::'a::ordered_ring_strict)))"
21328
73bb86d0f483 dropped Inductive dependency
haftmann
parents: 21258
diff changeset
   719
apply (cases "c = 0")
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   720
apply (auto simp add: linorder_neq_iff mult_strict_left_mono 
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   721
                      mult_strict_left_mono_neg)
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   722
apply (auto simp add: linorder_not_less 
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   723
                      linorder_not_le [symmetric, of "c*a"]
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   724
                      linorder_not_le [symmetric, of a])
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   725
apply (erule_tac [!] notE)
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   726
apply (auto simp add: order_less_imp_le mult_left_mono 
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   727
                      mult_left_mono_neg)
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   728
done
14268
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   729
15234
ec91a90c604e simplification tweaks for better arithmetic reasoning
paulson
parents: 15229
diff changeset
   730
ec91a90c604e simplification tweaks for better arithmetic reasoning
paulson
parents: 15229
diff changeset
   731
text{*The ``conjunction of implication'' lemmas produce two cases when the
ec91a90c604e simplification tweaks for better arithmetic reasoning
paulson
parents: 15229
diff changeset
   732
comparison is a goal, but give four when the comparison is an assumption.*}
ec91a90c604e simplification tweaks for better arithmetic reasoning
paulson
parents: 15229
diff changeset
   733
ec91a90c604e simplification tweaks for better arithmetic reasoning
paulson
parents: 15229
diff changeset
   734
lemma mult_less_cancel_right:
ec91a90c604e simplification tweaks for better arithmetic reasoning
paulson
parents: 15229
diff changeset
   735
  fixes c :: "'a :: ordered_ring_strict"
ec91a90c604e simplification tweaks for better arithmetic reasoning
paulson
parents: 15229
diff changeset
   736
  shows      "(a*c < b*c) = ((0 \<le> c --> a < b) & (c \<le> 0 --> b < a))"
ec91a90c604e simplification tweaks for better arithmetic reasoning
paulson
parents: 15229
diff changeset
   737
by (insert mult_less_cancel_right_disj [of a c b], auto)
ec91a90c604e simplification tweaks for better arithmetic reasoning
paulson
parents: 15229
diff changeset
   738
ec91a90c604e simplification tweaks for better arithmetic reasoning
paulson
parents: 15229
diff changeset
   739
lemma mult_less_cancel_left:
ec91a90c604e simplification tweaks for better arithmetic reasoning
paulson
parents: 15229
diff changeset
   740
  fixes c :: "'a :: ordered_ring_strict"
ec91a90c604e simplification tweaks for better arithmetic reasoning
paulson
parents: 15229
diff changeset
   741
  shows      "(c*a < c*b) = ((0 \<le> c --> a < b) & (c \<le> 0 --> b < a))"
ec91a90c604e simplification tweaks for better arithmetic reasoning
paulson
parents: 15229
diff changeset
   742
by (insert mult_less_cancel_left_disj [of c a b], auto)
ec91a90c604e simplification tweaks for better arithmetic reasoning
paulson
parents: 15229
diff changeset
   743
14268
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   744
lemma mult_le_cancel_right:
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   745
     "(a*c \<le> b*c) = ((0<c --> a\<le>b) & (c<0 --> b \<le> (a::'a::ordered_ring_strict)))"
15234
ec91a90c604e simplification tweaks for better arithmetic reasoning
paulson
parents: 15229
diff changeset
   746
by (simp add: linorder_not_less [symmetric] mult_less_cancel_right_disj)
14268
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   747
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   748
lemma mult_le_cancel_left:
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   749
     "(c*a \<le> c*b) = ((0<c --> a\<le>b) & (c<0 --> b \<le> (a::'a::ordered_ring_strict)))"
15234
ec91a90c604e simplification tweaks for better arithmetic reasoning
paulson
parents: 15229
diff changeset
   750
by (simp add: linorder_not_less [symmetric] mult_less_cancel_left_disj)
14268
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   751
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   752
lemma mult_less_imp_less_left:
14341
a09441bd4f1e Ring_and_Field now requires axiom add_left_imp_eq for semirings.
paulson
parents: 14334
diff changeset
   753
      assumes less: "c*a < c*b" and nonneg: "0 \<le> c"
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   754
      shows "a < (b::'a::ordered_semiring_strict)"
14377
f454b3004f8f tidying up, especially the Complex numbers
paulson
parents: 14370
diff changeset
   755
proof (rule ccontr)
f454b3004f8f tidying up, especially the Complex numbers
paulson
parents: 14370
diff changeset
   756
  assume "~ a < b"
f454b3004f8f tidying up, especially the Complex numbers
paulson
parents: 14370
diff changeset
   757
  hence "b \<le> a" by (simp add: linorder_not_less)
23389
aaca6a8e5414 tuned proofs: avoid implicit prems;
wenzelm
parents: 23326
diff changeset
   758
  hence "c*b \<le> c*a" using nonneg by (rule mult_left_mono)
14377
f454b3004f8f tidying up, especially the Complex numbers
paulson
parents: 14370
diff changeset
   759
  with this and less show False 
f454b3004f8f tidying up, especially the Complex numbers
paulson
parents: 14370
diff changeset
   760
    by (simp add: linorder_not_less [symmetric])
f454b3004f8f tidying up, especially the Complex numbers
paulson
parents: 14370
diff changeset
   761
qed
14268
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   762
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   763
lemma mult_less_imp_less_right:
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   764
  assumes less: "a*c < b*c" and nonneg: "0 <= c"
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   765
  shows "a < (b::'a::ordered_semiring_strict)"
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   766
proof (rule ccontr)
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   767
  assume "~ a < b"
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   768
  hence "b \<le> a" by (simp add: linorder_not_less)
23389
aaca6a8e5414 tuned proofs: avoid implicit prems;
wenzelm
parents: 23326
diff changeset
   769
  hence "b*c \<le> a*c" using nonneg by (rule mult_right_mono)
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   770
  with this and less show False 
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   771
    by (simp add: linorder_not_less [symmetric])
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   772
qed  
14268
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   773
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   774
text{*Cancellation of equalities with a common factor*}
24286
7619080e49f0 ATP blacklisting is now in theory data, attribute noatp
paulson
parents: 23879
diff changeset
   775
lemma mult_cancel_right [simp,noatp]:
22990
775e9de3db48 added classes ring_no_zero_divisors and dom (non-commutative version of idom);
huffman
parents: 22987
diff changeset
   776
  fixes a b c :: "'a::ring_no_zero_divisors"
775e9de3db48 added classes ring_no_zero_divisors and dom (non-commutative version of idom);
huffman
parents: 22987
diff changeset
   777
  shows "(a * c = b * c) = (c = 0 \<or> a = b)"
775e9de3db48 added classes ring_no_zero_divisors and dom (non-commutative version of idom);
huffman
parents: 22987
diff changeset
   778
proof -
775e9de3db48 added classes ring_no_zero_divisors and dom (non-commutative version of idom);
huffman
parents: 22987
diff changeset
   779
  have "(a * c = b * c) = ((a - b) * c = 0)"
23477
f4b83f03cac9 tuned and renamed group_eq_simps and ring_eq_simps
nipkow
parents: 23413
diff changeset
   780
    by (simp add: ring_distribs)
22990
775e9de3db48 added classes ring_no_zero_divisors and dom (non-commutative version of idom);
huffman
parents: 22987
diff changeset
   781
  thus ?thesis
775e9de3db48 added classes ring_no_zero_divisors and dom (non-commutative version of idom);
huffman
parents: 22987
diff changeset
   782
    by (simp add: disj_commute)
775e9de3db48 added classes ring_no_zero_divisors and dom (non-commutative version of idom);
huffman
parents: 22987
diff changeset
   783
qed
14268
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   784
24286
7619080e49f0 ATP blacklisting is now in theory data, attribute noatp
paulson
parents: 23879
diff changeset
   785
lemma mult_cancel_left [simp,noatp]:
22990
775e9de3db48 added classes ring_no_zero_divisors and dom (non-commutative version of idom);
huffman
parents: 22987
diff changeset
   786
  fixes a b c :: "'a::ring_no_zero_divisors"
775e9de3db48 added classes ring_no_zero_divisors and dom (non-commutative version of idom);
huffman
parents: 22987
diff changeset
   787
  shows "(c * a = c * b) = (c = 0 \<or> a = b)"
775e9de3db48 added classes ring_no_zero_divisors and dom (non-commutative version of idom);
huffman
parents: 22987
diff changeset
   788
proof -
775e9de3db48 added classes ring_no_zero_divisors and dom (non-commutative version of idom);
huffman
parents: 22987
diff changeset
   789
  have "(c * a = c * b) = (c * (a - b) = 0)"
23477
f4b83f03cac9 tuned and renamed group_eq_simps and ring_eq_simps
nipkow
parents: 23413
diff changeset
   790
    by (simp add: ring_distribs)
22990
775e9de3db48 added classes ring_no_zero_divisors and dom (non-commutative version of idom);
huffman
parents: 22987
diff changeset
   791
  thus ?thesis
775e9de3db48 added classes ring_no_zero_divisors and dom (non-commutative version of idom);
huffman
parents: 22987
diff changeset
   792
    by simp
775e9de3db48 added classes ring_no_zero_divisors and dom (non-commutative version of idom);
huffman
parents: 22987
diff changeset
   793
qed
14268
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   794
15234
ec91a90c604e simplification tweaks for better arithmetic reasoning
paulson
parents: 15229
diff changeset
   795
ec91a90c604e simplification tweaks for better arithmetic reasoning
paulson
parents: 15229
diff changeset
   796
subsubsection{*Special Cancellation Simprules for Multiplication*}
ec91a90c604e simplification tweaks for better arithmetic reasoning
paulson
parents: 15229
diff changeset
   797
ec91a90c604e simplification tweaks for better arithmetic reasoning
paulson
parents: 15229
diff changeset
   798
text{*These also produce two cases when the comparison is a goal.*}
ec91a90c604e simplification tweaks for better arithmetic reasoning
paulson
parents: 15229
diff changeset
   799
ec91a90c604e simplification tweaks for better arithmetic reasoning
paulson
parents: 15229
diff changeset
   800
lemma mult_le_cancel_right1:
ec91a90c604e simplification tweaks for better arithmetic reasoning
paulson
parents: 15229
diff changeset
   801
  fixes c :: "'a :: ordered_idom"
ec91a90c604e simplification tweaks for better arithmetic reasoning
paulson
parents: 15229
diff changeset
   802
  shows "(c \<le> b*c) = ((0<c --> 1\<le>b) & (c<0 --> b \<le> 1))"
ec91a90c604e simplification tweaks for better arithmetic reasoning
paulson
parents: 15229
diff changeset
   803
by (insert mult_le_cancel_right [of 1 c b], simp)
ec91a90c604e simplification tweaks for better arithmetic reasoning
paulson
parents: 15229
diff changeset
   804
ec91a90c604e simplification tweaks for better arithmetic reasoning
paulson
parents: 15229
diff changeset
   805
lemma mult_le_cancel_right2:
ec91a90c604e simplification tweaks for better arithmetic reasoning
paulson
parents: 15229
diff changeset
   806
  fixes c :: "'a :: ordered_idom"
ec91a90c604e simplification tweaks for better arithmetic reasoning
paulson
parents: 15229
diff changeset
   807
  shows "(a*c \<le> c) = ((0<c --> a\<le>1) & (c<0 --> 1 \<le> a))"
ec91a90c604e simplification tweaks for better arithmetic reasoning
paulson
parents: 15229
diff changeset
   808
by (insert mult_le_cancel_right [of a c 1], simp)
ec91a90c604e simplification tweaks for better arithmetic reasoning
paulson
parents: 15229
diff changeset
   809
ec91a90c604e simplification tweaks for better arithmetic reasoning
paulson
parents: 15229
diff changeset
   810
lemma mult_le_cancel_left1:
ec91a90c604e simplification tweaks for better arithmetic reasoning
paulson
parents: 15229
diff changeset
   811
  fixes c :: "'a :: ordered_idom"
ec91a90c604e simplification tweaks for better arithmetic reasoning
paulson
parents: 15229
diff changeset
   812
  shows "(c \<le> c*b) = ((0<c --> 1\<le>b) & (c<0 --> b \<le> 1))"
ec91a90c604e simplification tweaks for better arithmetic reasoning
paulson
parents: 15229
diff changeset
   813
by (insert mult_le_cancel_left [of c 1 b], simp)
ec91a90c604e simplification tweaks for better arithmetic reasoning
paulson
parents: 15229
diff changeset
   814
ec91a90c604e simplification tweaks for better arithmetic reasoning
paulson
parents: 15229
diff changeset
   815
lemma mult_le_cancel_left2:
ec91a90c604e simplification tweaks for better arithmetic reasoning
paulson
parents: 15229
diff changeset
   816
  fixes c :: "'a :: ordered_idom"
ec91a90c604e simplification tweaks for better arithmetic reasoning
paulson
parents: 15229
diff changeset
   817
  shows "(c*a \<le> c) = ((0<c --> a\<le>1) & (c<0 --> 1 \<le> a))"
ec91a90c604e simplification tweaks for better arithmetic reasoning
paulson
parents: 15229
diff changeset
   818
by (insert mult_le_cancel_left [of c a 1], simp)
ec91a90c604e simplification tweaks for better arithmetic reasoning
paulson
parents: 15229
diff changeset
   819
ec91a90c604e simplification tweaks for better arithmetic reasoning
paulson
parents: 15229
diff changeset
   820
lemma mult_less_cancel_right1:
ec91a90c604e simplification tweaks for better arithmetic reasoning
paulson
parents: 15229
diff changeset
   821
  fixes c :: "'a :: ordered_idom"
ec91a90c604e simplification tweaks for better arithmetic reasoning
paulson
parents: 15229
diff changeset
   822
  shows "(c < b*c) = ((0 \<le> c --> 1<b) & (c \<le> 0 --> b < 1))"
ec91a90c604e simplification tweaks for better arithmetic reasoning
paulson
parents: 15229
diff changeset
   823
by (insert mult_less_cancel_right [of 1 c b], simp)
ec91a90c604e simplification tweaks for better arithmetic reasoning
paulson
parents: 15229
diff changeset
   824
ec91a90c604e simplification tweaks for better arithmetic reasoning
paulson
parents: 15229
diff changeset
   825
lemma mult_less_cancel_right2:
ec91a90c604e simplification tweaks for better arithmetic reasoning
paulson
parents: 15229
diff changeset
   826
  fixes c :: "'a :: ordered_idom"
ec91a90c604e simplification tweaks for better arithmetic reasoning
paulson
parents: 15229
diff changeset
   827
  shows "(a*c < c) = ((0 \<le> c --> a<1) & (c \<le> 0 --> 1 < a))"
ec91a90c604e simplification tweaks for better arithmetic reasoning
paulson
parents: 15229
diff changeset
   828
by (insert mult_less_cancel_right [of a c 1], simp)
ec91a90c604e simplification tweaks for better arithmetic reasoning
paulson
parents: 15229
diff changeset
   829
ec91a90c604e simplification tweaks for better arithmetic reasoning
paulson
parents: 15229
diff changeset
   830
lemma mult_less_cancel_left1:
ec91a90c604e simplification tweaks for better arithmetic reasoning
paulson
parents: 15229
diff changeset
   831
  fixes c :: "'a :: ordered_idom"
ec91a90c604e simplification tweaks for better arithmetic reasoning
paulson
parents: 15229
diff changeset
   832
  shows "(c < c*b) = ((0 \<le> c --> 1<b) & (c \<le> 0 --> b < 1))"
ec91a90c604e simplification tweaks for better arithmetic reasoning
paulson
parents: 15229
diff changeset
   833
by (insert mult_less_cancel_left [of c 1 b], simp)
ec91a90c604e simplification tweaks for better arithmetic reasoning
paulson
parents: 15229
diff changeset
   834
ec91a90c604e simplification tweaks for better arithmetic reasoning
paulson
parents: 15229
diff changeset
   835
lemma mult_less_cancel_left2:
ec91a90c604e simplification tweaks for better arithmetic reasoning
paulson
parents: 15229
diff changeset
   836
  fixes c :: "'a :: ordered_idom"
ec91a90c604e simplification tweaks for better arithmetic reasoning
paulson
parents: 15229
diff changeset
   837
  shows "(c*a < c) = ((0 \<le> c --> a<1) & (c \<le> 0 --> 1 < a))"
ec91a90c604e simplification tweaks for better arithmetic reasoning
paulson
parents: 15229
diff changeset
   838
by (insert mult_less_cancel_left [of c a 1], simp)
ec91a90c604e simplification tweaks for better arithmetic reasoning
paulson
parents: 15229
diff changeset
   839
ec91a90c604e simplification tweaks for better arithmetic reasoning
paulson
parents: 15229
diff changeset
   840
lemma mult_cancel_right1 [simp]:
23544
4b4165cb3e0d rename class dom to ring_1_no_zero_divisors
huffman
parents: 23527
diff changeset
   841
  fixes c :: "'a :: ring_1_no_zero_divisors"
15234
ec91a90c604e simplification tweaks for better arithmetic reasoning
paulson
parents: 15229
diff changeset
   842
  shows "(c = b*c) = (c = 0 | b=1)"
ec91a90c604e simplification tweaks for better arithmetic reasoning
paulson
parents: 15229
diff changeset
   843
by (insert mult_cancel_right [of 1 c b], force)
ec91a90c604e simplification tweaks for better arithmetic reasoning
paulson
parents: 15229
diff changeset
   844
ec91a90c604e simplification tweaks for better arithmetic reasoning
paulson
parents: 15229
diff changeset
   845
lemma mult_cancel_right2 [simp]:
23544
4b4165cb3e0d rename class dom to ring_1_no_zero_divisors
huffman
parents: 23527
diff changeset
   846
  fixes c :: "'a :: ring_1_no_zero_divisors"
15234
ec91a90c604e simplification tweaks for better arithmetic reasoning
paulson
parents: 15229
diff changeset
   847
  shows "(a*c = c) = (c = 0 | a=1)"
ec91a90c604e simplification tweaks for better arithmetic reasoning
paulson
parents: 15229
diff changeset
   848
by (insert mult_cancel_right [of a c 1], simp)
ec91a90c604e simplification tweaks for better arithmetic reasoning
paulson
parents: 15229
diff changeset
   849
 
ec91a90c604e simplification tweaks for better arithmetic reasoning
paulson
parents: 15229
diff changeset
   850
lemma mult_cancel_left1 [simp]:
23544
4b4165cb3e0d rename class dom to ring_1_no_zero_divisors
huffman
parents: 23527
diff changeset
   851
  fixes c :: "'a :: ring_1_no_zero_divisors"
15234
ec91a90c604e simplification tweaks for better arithmetic reasoning
paulson
parents: 15229
diff changeset
   852
  shows "(c = c*b) = (c = 0 | b=1)"
ec91a90c604e simplification tweaks for better arithmetic reasoning
paulson
parents: 15229
diff changeset
   853
by (insert mult_cancel_left [of c 1 b], force)
ec91a90c604e simplification tweaks for better arithmetic reasoning
paulson
parents: 15229
diff changeset
   854
ec91a90c604e simplification tweaks for better arithmetic reasoning
paulson
parents: 15229
diff changeset
   855
lemma mult_cancel_left2 [simp]:
23544
4b4165cb3e0d rename class dom to ring_1_no_zero_divisors
huffman
parents: 23527
diff changeset
   856
  fixes c :: "'a :: ring_1_no_zero_divisors"
15234
ec91a90c604e simplification tweaks for better arithmetic reasoning
paulson
parents: 15229
diff changeset
   857
  shows "(c*a = c) = (c = 0 | a=1)"
ec91a90c604e simplification tweaks for better arithmetic reasoning
paulson
parents: 15229
diff changeset
   858
by (insert mult_cancel_left [of c a 1], simp)
ec91a90c604e simplification tweaks for better arithmetic reasoning
paulson
parents: 15229
diff changeset
   859
ec91a90c604e simplification tweaks for better arithmetic reasoning
paulson
parents: 15229
diff changeset
   860
ec91a90c604e simplification tweaks for better arithmetic reasoning
paulson
parents: 15229
diff changeset
   861
text{*Simprules for comparisons where common factors can be cancelled.*}
ec91a90c604e simplification tweaks for better arithmetic reasoning
paulson
parents: 15229
diff changeset
   862
lemmas mult_compare_simps =
ec91a90c604e simplification tweaks for better arithmetic reasoning
paulson
parents: 15229
diff changeset
   863
    mult_le_cancel_right mult_le_cancel_left
ec91a90c604e simplification tweaks for better arithmetic reasoning
paulson
parents: 15229
diff changeset
   864
    mult_le_cancel_right1 mult_le_cancel_right2
ec91a90c604e simplification tweaks for better arithmetic reasoning
paulson
parents: 15229
diff changeset
   865
    mult_le_cancel_left1 mult_le_cancel_left2
ec91a90c604e simplification tweaks for better arithmetic reasoning
paulson
parents: 15229
diff changeset
   866
    mult_less_cancel_right mult_less_cancel_left
ec91a90c604e simplification tweaks for better arithmetic reasoning
paulson
parents: 15229
diff changeset
   867
    mult_less_cancel_right1 mult_less_cancel_right2
ec91a90c604e simplification tweaks for better arithmetic reasoning
paulson
parents: 15229
diff changeset
   868
    mult_less_cancel_left1 mult_less_cancel_left2
ec91a90c604e simplification tweaks for better arithmetic reasoning
paulson
parents: 15229
diff changeset
   869
    mult_cancel_right mult_cancel_left
ec91a90c604e simplification tweaks for better arithmetic reasoning
paulson
parents: 15229
diff changeset
   870
    mult_cancel_right1 mult_cancel_right2
ec91a90c604e simplification tweaks for better arithmetic reasoning
paulson
parents: 15229
diff changeset
   871
    mult_cancel_left1 mult_cancel_left2
ec91a90c604e simplification tweaks for better arithmetic reasoning
paulson
parents: 15229
diff changeset
   872
ec91a90c604e simplification tweaks for better arithmetic reasoning
paulson
parents: 15229
diff changeset
   873
14265
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   874
subsection {* Fields *}
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   875
14288
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
   876
lemma right_inverse_eq: "b \<noteq> 0 ==> (a / b = 1) = (a = (b::'a::field))"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
   877
proof
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
   878
  assume neq: "b \<noteq> 0"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
   879
  {
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
   880
    hence "a = (a / b) * b" by (simp add: divide_inverse mult_ac)
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
   881
    also assume "a / b = 1"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
   882
    finally show "a = b" by simp
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
   883
  next
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
   884
    assume "a = b"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
   885
    with neq show "a / b = 1" by (simp add: divide_inverse)
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
   886
  }
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
   887
qed
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
   888
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
   889
lemma nonzero_inverse_eq_divide: "a \<noteq> 0 ==> inverse (a::'a::field) = 1/a"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
   890
by (simp add: divide_inverse)
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
   891
23398
0b5a400c7595 made divide_self a simp rule
nipkow
parents: 23389
diff changeset
   892
lemma divide_self[simp]: "a \<noteq> 0 ==> a / (a::'a::field) = 1"
14288
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
   893
  by (simp add: divide_inverse)
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
   894
14430
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14421
diff changeset
   895
lemma divide_zero [simp]: "a / 0 = (0::'a::{field,division_by_zero})"
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14421
diff changeset
   896
by (simp add: divide_inverse)
14277
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
   897
15228
4d332d10fa3d revised simprules for division
paulson
parents: 15197
diff changeset
   898
lemma divide_self_if [simp]:
4d332d10fa3d revised simprules for division
paulson
parents: 15197
diff changeset
   899
     "a / (a::'a::{field,division_by_zero}) = (if a=0 then 0 else 1)"
4d332d10fa3d revised simprules for division
paulson
parents: 15197
diff changeset
   900
  by (simp add: divide_self)
4d332d10fa3d revised simprules for division
paulson
parents: 15197
diff changeset
   901
14430
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14421
diff changeset
   902
lemma divide_zero_left [simp]: "0/a = (0::'a::field)"
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14421
diff changeset
   903
by (simp add: divide_inverse)
14277
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
   904
14430
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14421
diff changeset
   905
lemma inverse_eq_divide: "inverse (a::'a::field) = 1/a"
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14421
diff changeset
   906
by (simp add: divide_inverse)
14277
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
   907
14430
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14421
diff changeset
   908
lemma add_divide_distrib: "(a+b)/(c::'a::field) = a/c + b/c"
23477
f4b83f03cac9 tuned and renamed group_eq_simps and ring_eq_simps
nipkow
parents: 23413
diff changeset
   909
by (simp add: divide_inverse ring_distribs) 
14293
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
   910
23482
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
   911
(* what ordering?? this is a straight instance of mult_eq_0_iff
14270
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   912
text{*Compared with @{text mult_eq_0_iff}, this version removes the requirement
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   913
      of an ordering.*}
20496
23eb6034c06d added axclass division_ring (like field without commutativity; includes e.g. quaternions) and generalized some theorems from field to division_ring
huffman
parents: 19404
diff changeset
   914
lemma field_mult_eq_0_iff [simp]:
23eb6034c06d added axclass division_ring (like field without commutativity; includes e.g. quaternions) and generalized some theorems from field to division_ring
huffman
parents: 19404
diff changeset
   915
  "(a*b = (0::'a::division_ring)) = (a = 0 | b = 0)"
22990
775e9de3db48 added classes ring_no_zero_divisors and dom (non-commutative version of idom);
huffman
parents: 22987
diff changeset
   916
by simp
23482
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
   917
*)
23496
84e9216a6d0e removed redundant lemmas
nipkow
parents: 23483
diff changeset
   918
(* subsumed by mult_cancel lemmas on ring_no_zero_divisors
14268
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   919
text{*Cancellation of equalities with a common factor*}
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   920
lemma field_mult_cancel_right_lemma:
20496
23eb6034c06d added axclass division_ring (like field without commutativity; includes e.g. quaternions) and generalized some theorems from field to division_ring
huffman
parents: 19404
diff changeset
   921
      assumes cnz: "c \<noteq> (0::'a::division_ring)"
23eb6034c06d added axclass division_ring (like field without commutativity; includes e.g. quaternions) and generalized some theorems from field to division_ring
huffman
parents: 19404
diff changeset
   922
         and eq:  "a*c = b*c"
23eb6034c06d added axclass division_ring (like field without commutativity; includes e.g. quaternions) and generalized some theorems from field to division_ring
huffman
parents: 19404
diff changeset
   923
        shows "a=b"
14377
f454b3004f8f tidying up, especially the Complex numbers
paulson
parents: 14370
diff changeset
   924
proof -
14268
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   925
  have "(a * c) * inverse c = (b * c) * inverse c"
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   926
    by (simp add: eq)
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   927
  thus "a=b"
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   928
    by (simp add: mult_assoc cnz)
14377
f454b3004f8f tidying up, especially the Complex numbers
paulson
parents: 14370
diff changeset
   929
qed
14268
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   930
14348
744c868ee0b7 Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents: 14341
diff changeset
   931
lemma field_mult_cancel_right [simp]:
20496
23eb6034c06d added axclass division_ring (like field without commutativity; includes e.g. quaternions) and generalized some theorems from field to division_ring
huffman
parents: 19404
diff changeset
   932
     "(a*c = b*c) = (c = (0::'a::division_ring) | a=b)"
22990
775e9de3db48 added classes ring_no_zero_divisors and dom (non-commutative version of idom);
huffman
parents: 22987
diff changeset
   933
by simp
14268
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   934
14348
744c868ee0b7 Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents: 14341
diff changeset
   935
lemma field_mult_cancel_left [simp]:
20496
23eb6034c06d added axclass division_ring (like field without commutativity; includes e.g. quaternions) and generalized some theorems from field to division_ring
huffman
parents: 19404
diff changeset
   936
     "(c*a = c*b) = (c = (0::'a::division_ring) | a=b)"
22990
775e9de3db48 added classes ring_no_zero_divisors and dom (non-commutative version of idom);
huffman
parents: 22987
diff changeset
   937
by simp
23496
84e9216a6d0e removed redundant lemmas
nipkow
parents: 23483
diff changeset
   938
*)
20496
23eb6034c06d added axclass division_ring (like field without commutativity; includes e.g. quaternions) and generalized some theorems from field to division_ring
huffman
parents: 19404
diff changeset
   939
lemma nonzero_imp_inverse_nonzero:
23eb6034c06d added axclass division_ring (like field without commutativity; includes e.g. quaternions) and generalized some theorems from field to division_ring
huffman
parents: 19404
diff changeset
   940
  "a \<noteq> 0 ==> inverse a \<noteq> (0::'a::division_ring)"
14377
f454b3004f8f tidying up, especially the Complex numbers
paulson
parents: 14370
diff changeset
   941
proof
14268
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   942
  assume ianz: "inverse a = 0"
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   943
  assume "a \<noteq> 0"
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   944
  hence "1 = a * inverse a" by simp
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   945
  also have "... = 0" by (simp add: ianz)
20496
23eb6034c06d added axclass division_ring (like field without commutativity; includes e.g. quaternions) and generalized some theorems from field to division_ring
huffman
parents: 19404
diff changeset
   946
  finally have "1 = (0::'a::division_ring)" .
14268
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   947
  thus False by (simp add: eq_commute)
14377
f454b3004f8f tidying up, especially the Complex numbers
paulson
parents: 14370
diff changeset
   948
qed
14268
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   949
14277
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
   950
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
   951
subsection{*Basic Properties of @{term inverse}*}
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
   952
20496
23eb6034c06d added axclass division_ring (like field without commutativity; includes e.g. quaternions) and generalized some theorems from field to division_ring
huffman
parents: 19404
diff changeset
   953
lemma inverse_zero_imp_zero: "inverse a = 0 ==> a = (0::'a::division_ring)"
14268
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   954
apply (rule ccontr) 
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   955
apply (blast dest: nonzero_imp_inverse_nonzero) 
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   956
done
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   957
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   958
lemma inverse_nonzero_imp_nonzero:
20496
23eb6034c06d added axclass division_ring (like field without commutativity; includes e.g. quaternions) and generalized some theorems from field to division_ring
huffman
parents: 19404
diff changeset
   959
   "inverse a = 0 ==> a = (0::'a::division_ring)"
14268
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   960
apply (rule ccontr) 
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   961
apply (blast dest: nonzero_imp_inverse_nonzero) 
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   962
done
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   963
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   964
lemma inverse_nonzero_iff_nonzero [simp]:
20496
23eb6034c06d added axclass division_ring (like field without commutativity; includes e.g. quaternions) and generalized some theorems from field to division_ring
huffman
parents: 19404
diff changeset
   965
   "(inverse a = 0) = (a = (0::'a::{division_ring,division_by_zero}))"
14268
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   966
by (force dest: inverse_nonzero_imp_nonzero) 
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   967
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   968
lemma nonzero_inverse_minus_eq:
20496
23eb6034c06d added axclass division_ring (like field without commutativity; includes e.g. quaternions) and generalized some theorems from field to division_ring
huffman
parents: 19404
diff changeset
   969
      assumes [simp]: "a\<noteq>0"
23eb6034c06d added axclass division_ring (like field without commutativity; includes e.g. quaternions) and generalized some theorems from field to division_ring
huffman
parents: 19404
diff changeset
   970
      shows "inverse(-a) = -inverse(a::'a::division_ring)"
14377
f454b3004f8f tidying up, especially the Complex numbers
paulson
parents: 14370
diff changeset
   971
proof -
f454b3004f8f tidying up, especially the Complex numbers
paulson
parents: 14370
diff changeset
   972
  have "-a * inverse (- a) = -a * - inverse a"
f454b3004f8f tidying up, especially the Complex numbers
paulson
parents: 14370
diff changeset
   973
    by simp
f454b3004f8f tidying up, especially the Complex numbers
paulson
parents: 14370
diff changeset
   974
  thus ?thesis 
23496
84e9216a6d0e removed redundant lemmas
nipkow
parents: 23483
diff changeset
   975
    by (simp only: mult_cancel_left, simp)
14377
f454b3004f8f tidying up, especially the Complex numbers
paulson
parents: 14370
diff changeset
   976
qed
14268
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   977
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   978
lemma inverse_minus_eq [simp]:
20496
23eb6034c06d added axclass division_ring (like field without commutativity; includes e.g. quaternions) and generalized some theorems from field to division_ring
huffman
parents: 19404
diff changeset
   979
   "inverse(-a) = -inverse(a::'a::{division_ring,division_by_zero})"
14377
f454b3004f8f tidying up, especially the Complex numbers
paulson
parents: 14370
diff changeset
   980
proof cases
f454b3004f8f tidying up, especially the Complex numbers
paulson
parents: 14370
diff changeset
   981
  assume "a=0" thus ?thesis by (simp add: inverse_zero)
f454b3004f8f tidying up, especially the Complex numbers
paulson
parents: 14370
diff changeset
   982
next
f454b3004f8f tidying up, especially the Complex numbers
paulson
parents: 14370
diff changeset
   983
  assume "a\<noteq>0" 
f454b3004f8f tidying up, especially the Complex numbers
paulson
parents: 14370
diff changeset
   984
  thus ?thesis by (simp add: nonzero_inverse_minus_eq)
f454b3004f8f tidying up, especially the Complex numbers
paulson
parents: 14370
diff changeset
   985
qed
14268
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   986
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   987
lemma nonzero_inverse_eq_imp_eq:
14269
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 14268
diff changeset
   988
      assumes inveq: "inverse a = inverse b"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 14268
diff changeset
   989
	  and anz:  "a \<noteq> 0"
502a7c95de73 conversion of some Real theories to Isar scripts
paulson
parents: 14268
diff changeset
   990
	  and bnz:  "b \<noteq> 0"
20496
23eb6034c06d added axclass division_ring (like field without commutativity; includes e.g. quaternions) and generalized some theorems from field to division_ring
huffman
parents: 19404
diff changeset
   991
	 shows "a = (b::'a::division_ring)"
14377
f454b3004f8f tidying up, especially the Complex numbers
paulson
parents: 14370
diff changeset
   992
proof -
14268
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   993
  have "a * inverse b = a * inverse a"
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   994
    by (simp add: inveq)
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   995
  hence "(a * inverse b) * b = (a * inverse a) * b"
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   996
    by simp
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   997
  thus "a = b"
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   998
    by (simp add: mult_assoc anz bnz)
14377
f454b3004f8f tidying up, especially the Complex numbers
paulson
parents: 14370
diff changeset
   999
qed
14268
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
  1000
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
  1001
lemma inverse_eq_imp_eq:
20496
23eb6034c06d added axclass division_ring (like field without commutativity; includes e.g. quaternions) and generalized some theorems from field to division_ring
huffman
parents: 19404
diff changeset
  1002
  "inverse a = inverse b ==> a = (b::'a::{division_ring,division_by_zero})"
21328
73bb86d0f483 dropped Inductive dependency
haftmann
parents: 21258
diff changeset
  1003
apply (cases "a=0 | b=0") 
14268
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
  1004
 apply (force dest!: inverse_zero_imp_zero
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
  1005
              simp add: eq_commute [of "0::'a"])
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
  1006
apply (force dest!: nonzero_inverse_eq_imp_eq) 
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
  1007
done
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
  1008
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
  1009
lemma inverse_eq_iff_eq [simp]:
20496
23eb6034c06d added axclass division_ring (like field without commutativity; includes e.g. quaternions) and generalized some theorems from field to division_ring
huffman
parents: 19404
diff changeset
  1010
  "(inverse a = inverse b) = (a = (b::'a::{division_ring,division_by_zero}))"
23eb6034c06d added axclass division_ring (like field without commutativity; includes e.g. quaternions) and generalized some theorems from field to division_ring
huffman
parents: 19404
diff changeset
  1011
by (force dest!: inverse_eq_imp_eq)
14268
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
  1012
14270
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
  1013
lemma nonzero_inverse_inverse_eq:
20496
23eb6034c06d added axclass division_ring (like field without commutativity; includes e.g. quaternions) and generalized some theorems from field to division_ring
huffman
parents: 19404
diff changeset
  1014
      assumes [simp]: "a \<noteq> 0"
23eb6034c06d added axclass division_ring (like field without commutativity; includes e.g. quaternions) and generalized some theorems from field to division_ring
huffman
parents: 19404
diff changeset
  1015
      shows "inverse(inverse (a::'a::division_ring)) = a"
14270
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
  1016
  proof -
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
  1017
  have "(inverse (inverse a) * inverse a) * a = a" 
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
  1018
    by (simp add: nonzero_imp_inverse_nonzero)
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
  1019
  thus ?thesis
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
  1020
    by (simp add: mult_assoc)
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
  1021
  qed
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
  1022
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
  1023
lemma inverse_inverse_eq [simp]:
20496
23eb6034c06d added axclass division_ring (like field without commutativity; includes e.g. quaternions) and generalized some theorems from field to division_ring
huffman
parents: 19404
diff changeset
  1024
     "inverse(inverse (a::'a::{division_ring,division_by_zero})) = a"
14270
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
  1025
  proof cases
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
  1026
    assume "a=0" thus ?thesis by simp
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
  1027
  next
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
  1028
    assume "a\<noteq>0" 
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
  1029
    thus ?thesis by (simp add: nonzero_inverse_inverse_eq)
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
  1030
  qed
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
  1031
20496
23eb6034c06d added axclass division_ring (like field without commutativity; includes e.g. quaternions) and generalized some theorems from field to division_ring
huffman
parents: 19404
diff changeset
  1032
lemma inverse_1 [simp]: "inverse 1 = (1::'a::division_ring)"
14270
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
  1033
  proof -
20496
23eb6034c06d added axclass division_ring (like field without commutativity; includes e.g. quaternions) and generalized some theorems from field to division_ring
huffman
parents: 19404
diff changeset
  1034
  have "inverse 1 * 1 = (1::'a::division_ring)" 
14270
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
  1035
    by (rule left_inverse [OF zero_neq_one [symmetric]])
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
  1036
  thus ?thesis  by simp
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
  1037
  qed
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
  1038
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15010
diff changeset
  1039
lemma inverse_unique: 
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15010
diff changeset
  1040
  assumes ab: "a*b = 1"
20496
23eb6034c06d added axclass division_ring (like field without commutativity; includes e.g. quaternions) and generalized some theorems from field to division_ring
huffman
parents: 19404
diff changeset
  1041
  shows "inverse a = (b::'a::division_ring)"
15077
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15010
diff changeset
  1042
proof -
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15010
diff changeset
  1043
  have "a \<noteq> 0" using ab by auto
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15010
diff changeset
  1044
  moreover have "inverse a * (a * b) = inverse a" by (simp add: ab) 
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15010
diff changeset
  1045
  ultimately show ?thesis by (simp add: mult_assoc [symmetric]) 
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15010
diff changeset
  1046
qed
89840837108e converting Hyperreal/Transcendental to Isar script
paulson
parents: 15010
diff changeset
  1047
14270
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
  1048
lemma nonzero_inverse_mult_distrib: 
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
  1049
      assumes anz: "a \<noteq> 0"
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
  1050
          and bnz: "b \<noteq> 0"
20496
23eb6034c06d added axclass division_ring (like field without commutativity; includes e.g. quaternions) and generalized some theorems from field to division_ring
huffman
parents: 19404
diff changeset
  1051
      shows "inverse(a*b) = inverse(b) * inverse(a::'a::division_ring)"
14270
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
  1052
  proof -
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
  1053
  have "inverse(a*b) * (a * b) * inverse(b) = inverse(b)" 
23482
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
  1054
    by (simp add: anz bnz)
14270
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
  1055
  hence "inverse(a*b) * a = inverse(b)" 
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
  1056
    by (simp add: mult_assoc bnz)
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
  1057
  hence "inverse(a*b) * a * inverse(a) = inverse(b) * inverse(a)" 
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
  1058
    by simp
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
  1059
  thus ?thesis
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
  1060
    by (simp add: mult_assoc anz)
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
  1061
  qed
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
  1062
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
  1063
text{*This version builds in division by zero while also re-orienting
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
  1064
      the right-hand side.*}
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
  1065
lemma inverse_mult_distrib [simp]:
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
  1066
     "inverse(a*b) = inverse(a) * inverse(b::'a::{field,division_by_zero})"
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
  1067
  proof cases
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
  1068
    assume "a \<noteq> 0 & b \<noteq> 0" 
22993
haftmann
parents: 22990
diff changeset
  1069
    thus ?thesis
haftmann
parents: 22990
diff changeset
  1070
      by (simp add: nonzero_inverse_mult_distrib mult_commute)
14270
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
  1071
  next
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
  1072
    assume "~ (a \<noteq> 0 & b \<noteq> 0)" 
22993
haftmann
parents: 22990
diff changeset
  1073
    thus ?thesis
haftmann
parents: 22990
diff changeset
  1074
      by force
14270
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
  1075
  qed
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
  1076
20496
23eb6034c06d added axclass division_ring (like field without commutativity; includes e.g. quaternions) and generalized some theorems from field to division_ring
huffman
parents: 19404
diff changeset
  1077
lemma division_ring_inverse_add:
23eb6034c06d added axclass division_ring (like field without commutativity; includes e.g. quaternions) and generalized some theorems from field to division_ring
huffman
parents: 19404
diff changeset
  1078
  "[|(a::'a::division_ring) \<noteq> 0; b \<noteq> 0|]
23eb6034c06d added axclass division_ring (like field without commutativity; includes e.g. quaternions) and generalized some theorems from field to division_ring
huffman
parents: 19404
diff changeset
  1079
   ==> inverse a + inverse b = inverse a * (a+b) * inverse b"
23477
f4b83f03cac9 tuned and renamed group_eq_simps and ring_eq_simps
nipkow
parents: 23413
diff changeset
  1080
by (simp add: ring_simps)
20496
23eb6034c06d added axclass division_ring (like field without commutativity; includes e.g. quaternions) and generalized some theorems from field to division_ring
huffman
parents: 19404
diff changeset
  1081
23eb6034c06d added axclass division_ring (like field without commutativity; includes e.g. quaternions) and generalized some theorems from field to division_ring
huffman
parents: 19404
diff changeset
  1082
lemma division_ring_inverse_diff:
23eb6034c06d added axclass division_ring (like field without commutativity; includes e.g. quaternions) and generalized some theorems from field to division_ring
huffman
parents: 19404
diff changeset
  1083
  "[|(a::'a::division_ring) \<noteq> 0; b \<noteq> 0|]
23eb6034c06d added axclass division_ring (like field without commutativity; includes e.g. quaternions) and generalized some theorems from field to division_ring
huffman
parents: 19404
diff changeset
  1084
   ==> inverse a - inverse b = inverse a * (b-a) * inverse b"
23477
f4b83f03cac9 tuned and renamed group_eq_simps and ring_eq_simps
nipkow
parents: 23413
diff changeset
  1085
by (simp add: ring_simps)
20496
23eb6034c06d added axclass division_ring (like field without commutativity; includes e.g. quaternions) and generalized some theorems from field to division_ring
huffman
parents: 19404
diff changeset
  1086
14270
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
  1087
text{*There is no slick version using division by zero.*}
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
  1088
lemma inverse_add:
23477
f4b83f03cac9 tuned and renamed group_eq_simps and ring_eq_simps
nipkow
parents: 23413
diff changeset
  1089
  "[|a \<noteq> 0;  b \<noteq> 0|]
f4b83f03cac9 tuned and renamed group_eq_simps and ring_eq_simps
nipkow
parents: 23413
diff changeset
  1090
   ==> inverse a + inverse b = (a+b) * inverse a * inverse (b::'a::field)"
20496
23eb6034c06d added axclass division_ring (like field without commutativity; includes e.g. quaternions) and generalized some theorems from field to division_ring
huffman
parents: 19404
diff changeset
  1091
by (simp add: division_ring_inverse_add mult_ac)
14270
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
  1092
14365
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14353
diff changeset
  1093
lemma inverse_divide [simp]:
23477
f4b83f03cac9 tuned and renamed group_eq_simps and ring_eq_simps
nipkow
parents: 23413
diff changeset
  1094
  "inverse (a/b) = b / (a::'a::{field,division_by_zero})"
f4b83f03cac9 tuned and renamed group_eq_simps and ring_eq_simps
nipkow
parents: 23413
diff changeset
  1095
by (simp add: divide_inverse mult_commute)
14365
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14353
diff changeset
  1096
23389
aaca6a8e5414 tuned proofs: avoid implicit prems;
wenzelm
parents: 23326
diff changeset
  1097
16775
c1b87ef4a1c3 added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
avigad
parents: 16568
diff changeset
  1098
subsection {* Calculations with fractions *}
c1b87ef4a1c3 added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
avigad
parents: 16568
diff changeset
  1099
23413
5caa2710dd5b tuned laws for cancellation in divisions for fields.
nipkow
parents: 23406
diff changeset
  1100
text{* There is a whole bunch of simp-rules just for class @{text
5caa2710dd5b tuned laws for cancellation in divisions for fields.
nipkow
parents: 23406
diff changeset
  1101
field} but none for class @{text field} and @{text nonzero_divides}
5caa2710dd5b tuned laws for cancellation in divisions for fields.
nipkow
parents: 23406
diff changeset
  1102
because the latter are covered by a simproc. *}
5caa2710dd5b tuned laws for cancellation in divisions for fields.
nipkow
parents: 23406
diff changeset
  1103
24427
bc5cf3b09ff3 revised blacklisting for ATP linkup
paulson
parents: 24422
diff changeset
  1104
lemma nonzero_mult_divide_mult_cancel_left[simp,noatp]:
23477
f4b83f03cac9 tuned and renamed group_eq_simps and ring_eq_simps
nipkow
parents: 23413
diff changeset
  1105
assumes [simp]: "b\<noteq>0" and [simp]: "c\<noteq>0" shows "(c*a)/(c*b) = a/(b::'a::field)"
14277
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
  1106
proof -
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
  1107
  have "(c*a)/(c*b) = c * a * (inverse b * inverse c)"
23482
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
  1108
    by (simp add: divide_inverse nonzero_inverse_mult_distrib)
14277
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
  1109
  also have "... =  a * inverse b * (inverse c * c)"
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
  1110
    by (simp only: mult_ac)
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
  1111
  also have "... =  a * inverse b"
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
  1112
    by simp
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
  1113
    finally show ?thesis 
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
  1114
    by (simp add: divide_inverse)
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
  1115
qed
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
  1116
23413
5caa2710dd5b tuned laws for cancellation in divisions for fields.
nipkow
parents: 23406
diff changeset
  1117
lemma mult_divide_mult_cancel_left:
23477
f4b83f03cac9 tuned and renamed group_eq_simps and ring_eq_simps
nipkow
parents: 23413
diff changeset
  1118
  "c\<noteq>0 ==> (c*a) / (c*b) = a / (b::'a::{field,division_by_zero})"
21328
73bb86d0f483 dropped Inductive dependency
haftmann
parents: 21258
diff changeset
  1119
apply (cases "b = 0")
23413
5caa2710dd5b tuned laws for cancellation in divisions for fields.
nipkow
parents: 23406
diff changeset
  1120
apply (simp_all add: nonzero_mult_divide_mult_cancel_left)
14277
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
  1121
done
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
  1122
24427
bc5cf3b09ff3 revised blacklisting for ATP linkup
paulson
parents: 24422
diff changeset
  1123
lemma nonzero_mult_divide_mult_cancel_right [noatp]:
23477
f4b83f03cac9 tuned and renamed group_eq_simps and ring_eq_simps
nipkow
parents: 23413
diff changeset
  1124
  "[|b\<noteq>0; c\<noteq>0|] ==> (a*c) / (b*c) = a/(b::'a::field)"
23413
5caa2710dd5b tuned laws for cancellation in divisions for fields.
nipkow
parents: 23406
diff changeset
  1125
by (simp add: mult_commute [of _ c] nonzero_mult_divide_mult_cancel_left) 
14321
55c688d2eefa new theorems
paulson
parents: 14305
diff changeset
  1126
23413
5caa2710dd5b tuned laws for cancellation in divisions for fields.
nipkow
parents: 23406
diff changeset
  1127
lemma mult_divide_mult_cancel_right:
23477
f4b83f03cac9 tuned and renamed group_eq_simps and ring_eq_simps
nipkow
parents: 23413
diff changeset
  1128
  "c\<noteq>0 ==> (a*c) / (b*c) = a / (b::'a::{field,division_by_zero})"
21328
73bb86d0f483 dropped Inductive dependency
haftmann
parents: 21258
diff changeset
  1129
apply (cases "b = 0")
23413
5caa2710dd5b tuned laws for cancellation in divisions for fields.
nipkow
parents: 23406
diff changeset
  1130
apply (simp_all add: nonzero_mult_divide_mult_cancel_right)
14321
55c688d2eefa new theorems
paulson
parents: 14305
diff changeset
  1131
done
23413
5caa2710dd5b tuned laws for cancellation in divisions for fields.
nipkow
parents: 23406
diff changeset
  1132
14284
f1abe67c448a re-organisation of Real/RealArith0.ML; more `Isar scripts
paulson
parents: 14277
diff changeset
  1133
lemma divide_1 [simp]: "a/1 = (a::'a::field)"
23477
f4b83f03cac9 tuned and renamed group_eq_simps and ring_eq_simps
nipkow
parents: 23413
diff changeset
  1134
by (simp add: divide_inverse)
14284
f1abe67c448a re-organisation of Real/RealArith0.ML; more `Isar scripts
paulson
parents: 14277
diff changeset
  1135
15234
ec91a90c604e simplification tweaks for better arithmetic reasoning
paulson
parents: 15229
diff changeset
  1136
lemma times_divide_eq_right: "a * (b/c) = (a*b) / (c::'a::field)"
14430
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14421
diff changeset
  1137
by (simp add: divide_inverse mult_assoc)
14288
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1138
14430
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14421
diff changeset
  1139
lemma times_divide_eq_left: "(b/c) * a = (b*a) / (c::'a::field)"
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14421
diff changeset
  1140
by (simp add: divide_inverse mult_ac)
14288
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1141
23482
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
  1142
lemmas times_divide_eq = times_divide_eq_right times_divide_eq_left
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
  1143
24286
7619080e49f0 ATP blacklisting is now in theory data, attribute noatp
paulson
parents: 23879
diff changeset
  1144
lemma divide_divide_eq_right [simp,noatp]:
23477
f4b83f03cac9 tuned and renamed group_eq_simps and ring_eq_simps
nipkow
parents: 23413
diff changeset
  1145
  "a / (b/c) = (a*c) / (b::'a::{field,division_by_zero})"
14430
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14421
diff changeset
  1146
by (simp add: divide_inverse mult_ac)
14288
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1147
24286
7619080e49f0 ATP blacklisting is now in theory data, attribute noatp
paulson
parents: 23879
diff changeset
  1148
lemma divide_divide_eq_left [simp,noatp]:
23477
f4b83f03cac9 tuned and renamed group_eq_simps and ring_eq_simps
nipkow
parents: 23413
diff changeset
  1149
  "(a / b) / (c::'a::{field,division_by_zero}) = a / (b*c)"
14430
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14421
diff changeset
  1150
by (simp add: divide_inverse mult_assoc)
14288
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1151
16775
c1b87ef4a1c3 added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
avigad
parents: 16568
diff changeset
  1152
lemma add_frac_eq: "(y::'a::field) ~= 0 ==> z ~= 0 ==>
c1b87ef4a1c3 added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
avigad
parents: 16568
diff changeset
  1153
    x / y + w / z = (x * z + w * y) / (y * z)"
23477
f4b83f03cac9 tuned and renamed group_eq_simps and ring_eq_simps
nipkow
parents: 23413
diff changeset
  1154
apply (subgoal_tac "x / y = (x * z) / (y * z)")
f4b83f03cac9 tuned and renamed group_eq_simps and ring_eq_simps
nipkow
parents: 23413
diff changeset
  1155
apply (erule ssubst)
f4b83f03cac9 tuned and renamed group_eq_simps and ring_eq_simps
nipkow
parents: 23413
diff changeset
  1156
apply (subgoal_tac "w / z = (w * y) / (y * z)")
f4b83f03cac9 tuned and renamed group_eq_simps and ring_eq_simps
nipkow
parents: 23413
diff changeset
  1157
apply (erule ssubst)
f4b83f03cac9 tuned and renamed group_eq_simps and ring_eq_simps
nipkow
parents: 23413
diff changeset
  1158
apply (rule add_divide_distrib [THEN sym])
f4b83f03cac9 tuned and renamed group_eq_simps and ring_eq_simps
nipkow
parents: 23413
diff changeset
  1159
apply (subst mult_commute)
f4b83f03cac9 tuned and renamed group_eq_simps and ring_eq_simps
nipkow
parents: 23413
diff changeset
  1160
apply (erule nonzero_mult_divide_mult_cancel_left [THEN sym])
f4b83f03cac9 tuned and renamed group_eq_simps and ring_eq_simps
nipkow
parents: 23413
diff changeset
  1161
apply assumption
f4b83f03cac9 tuned and renamed group_eq_simps and ring_eq_simps
nipkow
parents: 23413
diff changeset
  1162
apply (erule nonzero_mult_divide_mult_cancel_right [THEN sym])
f4b83f03cac9 tuned and renamed group_eq_simps and ring_eq_simps
nipkow
parents: 23413
diff changeset
  1163
apply assumption
16775
c1b87ef4a1c3 added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
avigad
parents: 16568
diff changeset
  1164
done
14268
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
  1165
23389
aaca6a8e5414 tuned proofs: avoid implicit prems;
wenzelm
parents: 23326
diff changeset
  1166
15234
ec91a90c604e simplification tweaks for better arithmetic reasoning
paulson
parents: 15229
diff changeset
  1167
subsubsection{*Special Cancellation Simprules for Division*}
ec91a90c604e simplification tweaks for better arithmetic reasoning
paulson
parents: 15229
diff changeset
  1168
24427
bc5cf3b09ff3 revised blacklisting for ATP linkup
paulson
parents: 24422
diff changeset
  1169
lemma mult_divide_mult_cancel_left_if[simp,noatp]:
23477
f4b83f03cac9 tuned and renamed group_eq_simps and ring_eq_simps
nipkow
parents: 23413
diff changeset
  1170
fixes c :: "'a :: {field,division_by_zero}"
f4b83f03cac9 tuned and renamed group_eq_simps and ring_eq_simps
nipkow
parents: 23413
diff changeset
  1171
shows "(c*a) / (c*b) = (if c=0 then 0 else a/b)"
23413
5caa2710dd5b tuned laws for cancellation in divisions for fields.
nipkow
parents: 23406
diff changeset
  1172
by (simp add: mult_divide_mult_cancel_left)
5caa2710dd5b tuned laws for cancellation in divisions for fields.
nipkow
parents: 23406
diff changeset
  1173
24427
bc5cf3b09ff3 revised blacklisting for ATP linkup
paulson
parents: 24422
diff changeset
  1174
lemma nonzero_mult_divide_cancel_right[simp,noatp]:
23413
5caa2710dd5b tuned laws for cancellation in divisions for fields.
nipkow
parents: 23406
diff changeset
  1175
  "b \<noteq> 0 \<Longrightarrow> a * b / b = (a::'a::field)"
5caa2710dd5b tuned laws for cancellation in divisions for fields.
nipkow
parents: 23406
diff changeset
  1176
using nonzero_mult_divide_mult_cancel_right[of 1 b a] by simp
5caa2710dd5b tuned laws for cancellation in divisions for fields.
nipkow
parents: 23406
diff changeset
  1177
24427
bc5cf3b09ff3 revised blacklisting for ATP linkup
paulson
parents: 24422
diff changeset
  1178
lemma nonzero_mult_divide_cancel_left[simp,noatp]:
23413
5caa2710dd5b tuned laws for cancellation in divisions for fields.
nipkow
parents: 23406
diff changeset
  1179
  "a \<noteq> 0 \<Longrightarrow> a * b / a = (b::'a::field)"
5caa2710dd5b tuned laws for cancellation in divisions for fields.
nipkow
parents: 23406
diff changeset
  1180
using nonzero_mult_divide_mult_cancel_left[of 1 a b] by simp
5caa2710dd5b tuned laws for cancellation in divisions for fields.
nipkow
parents: 23406
diff changeset
  1181
5caa2710dd5b tuned laws for cancellation in divisions for fields.
nipkow
parents: 23406
diff changeset
  1182
24427
bc5cf3b09ff3 revised blacklisting for ATP linkup
paulson
parents: 24422
diff changeset
  1183
lemma nonzero_divide_mult_cancel_right[simp,noatp]:
23413
5caa2710dd5b tuned laws for cancellation in divisions for fields.
nipkow
parents: 23406
diff changeset
  1184
  "\<lbrakk> a\<noteq>0; b\<noteq>0 \<rbrakk> \<Longrightarrow> b / (a * b) = 1/(a::'a::field)"
5caa2710dd5b tuned laws for cancellation in divisions for fields.
nipkow
parents: 23406
diff changeset
  1185
using nonzero_mult_divide_mult_cancel_right[of a b 1] by simp
5caa2710dd5b tuned laws for cancellation in divisions for fields.
nipkow
parents: 23406
diff changeset
  1186
24427
bc5cf3b09ff3 revised blacklisting for ATP linkup
paulson
parents: 24422
diff changeset
  1187
lemma nonzero_divide_mult_cancel_left[simp,noatp]:
23413
5caa2710dd5b tuned laws for cancellation in divisions for fields.
nipkow
parents: 23406
diff changeset
  1188
  "\<lbrakk> a\<noteq>0; b\<noteq>0 \<rbrakk> \<Longrightarrow> a / (a * b) = 1/(b::'a::field)"
5caa2710dd5b tuned laws for cancellation in divisions for fields.
nipkow
parents: 23406
diff changeset
  1189
using nonzero_mult_divide_mult_cancel_left[of b a 1] by simp
5caa2710dd5b tuned laws for cancellation in divisions for fields.
nipkow
parents: 23406
diff changeset
  1190
5caa2710dd5b tuned laws for cancellation in divisions for fields.
nipkow
parents: 23406
diff changeset
  1191
24427
bc5cf3b09ff3 revised blacklisting for ATP linkup
paulson
parents: 24422
diff changeset
  1192
lemma nonzero_mult_divide_mult_cancel_left2[simp,noatp]:
23477
f4b83f03cac9 tuned and renamed group_eq_simps and ring_eq_simps
nipkow
parents: 23413
diff changeset
  1193
  "[|b\<noteq>0; c\<noteq>0|] ==> (c*a) / (b*c) = a/(b::'a::field)"
23413
5caa2710dd5b tuned laws for cancellation in divisions for fields.
nipkow
parents: 23406
diff changeset
  1194
using nonzero_mult_divide_mult_cancel_left[of b c a] by(simp add:mult_ac)
5caa2710dd5b tuned laws for cancellation in divisions for fields.
nipkow
parents: 23406
diff changeset
  1195
24427
bc5cf3b09ff3 revised blacklisting for ATP linkup
paulson
parents: 24422
diff changeset
  1196
lemma nonzero_mult_divide_mult_cancel_right2[simp,noatp]:
23477
f4b83f03cac9 tuned and renamed group_eq_simps and ring_eq_simps
nipkow
parents: 23413
diff changeset
  1197
  "[|b\<noteq>0; c\<noteq>0|] ==> (a*c) / (c*b) = a/(b::'a::field)"
23413
5caa2710dd5b tuned laws for cancellation in divisions for fields.
nipkow
parents: 23406
diff changeset
  1198
using nonzero_mult_divide_mult_cancel_right[of b c a] by(simp add:mult_ac)
5caa2710dd5b tuned laws for cancellation in divisions for fields.
nipkow
parents: 23406
diff changeset
  1199
15234
ec91a90c604e simplification tweaks for better arithmetic reasoning
paulson
parents: 15229
diff changeset
  1200
14293
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1201
subsection {* Division and Unary Minus *}
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1202
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1203
lemma nonzero_minus_divide_left: "b \<noteq> 0 ==> - (a/b) = (-a) / (b::'a::field)"
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1204
by (simp add: divide_inverse minus_mult_left)
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1205
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1206
lemma nonzero_minus_divide_right: "b \<noteq> 0 ==> - (a/b) = a / -(b::'a::field)"
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1207
by (simp add: divide_inverse nonzero_inverse_minus_eq minus_mult_right)
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1208
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1209
lemma nonzero_minus_divide_divide: "b \<noteq> 0 ==> (-a)/(-b) = a / (b::'a::field)"
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1210
by (simp add: divide_inverse nonzero_inverse_minus_eq)
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1211
14430
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14421
diff changeset
  1212
lemma minus_divide_left: "- (a/b) = (-a) / (b::'a::field)"
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14421
diff changeset
  1213
by (simp add: divide_inverse minus_mult_left [symmetric])
14293
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1214
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1215
lemma minus_divide_right: "- (a/b) = a / -(b::'a::{field,division_by_zero})"
14430
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14421
diff changeset
  1216
by (simp add: divide_inverse minus_mult_right [symmetric])
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14421
diff changeset
  1217
14293
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1218
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1219
text{*The effect is to extract signs from divisions*}
17085
5b57f995a179 more simprules now have names
paulson
parents: 16775
diff changeset
  1220
lemmas divide_minus_left = minus_divide_left [symmetric]
5b57f995a179 more simprules now have names
paulson
parents: 16775
diff changeset
  1221
lemmas divide_minus_right = minus_divide_right [symmetric]
5b57f995a179 more simprules now have names
paulson
parents: 16775
diff changeset
  1222
declare divide_minus_left [simp]   divide_minus_right [simp]
14293
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1223
14387
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14377
diff changeset
  1224
text{*Also, extract signs from products*}
17085
5b57f995a179 more simprules now have names
paulson
parents: 16775
diff changeset
  1225
lemmas mult_minus_left = minus_mult_left [symmetric]
5b57f995a179 more simprules now have names
paulson
parents: 16775
diff changeset
  1226
lemmas mult_minus_right = minus_mult_right [symmetric]
5b57f995a179 more simprules now have names
paulson
parents: 16775
diff changeset
  1227
declare mult_minus_left [simp]   mult_minus_right [simp]
14387
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14377
diff changeset
  1228
14293
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1229
lemma minus_divide_divide [simp]:
23477
f4b83f03cac9 tuned and renamed group_eq_simps and ring_eq_simps
nipkow
parents: 23413
diff changeset
  1230
  "(-a)/(-b) = a / (b::'a::{field,division_by_zero})"
21328
73bb86d0f483 dropped Inductive dependency
haftmann
parents: 21258
diff changeset
  1231
apply (cases "b=0", simp) 
14293
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1232
apply (simp add: nonzero_minus_divide_divide) 
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1233
done
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1234
14430
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14421
diff changeset
  1235
lemma diff_divide_distrib: "(a-b)/(c::'a::field) = a/c - b/c"
14387
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14377
diff changeset
  1236
by (simp add: diff_minus add_divide_distrib) 
e96d5c42c4b0 Polymorphic treatment of binary arithmetic using axclasses
paulson
parents: 14377
diff changeset
  1237
23482
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
  1238
lemma add_divide_eq_iff:
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
  1239
  "(z::'a::field) \<noteq> 0 \<Longrightarrow> x + y/z = (z*x + y)/z"
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
  1240
by(simp add:add_divide_distrib nonzero_mult_divide_cancel_left)
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
  1241
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
  1242
lemma divide_add_eq_iff:
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
  1243
  "(z::'a::field) \<noteq> 0 \<Longrightarrow> x/z + y = (x + z*y)/z"
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
  1244
by(simp add:add_divide_distrib nonzero_mult_divide_cancel_left)
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
  1245
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
  1246
lemma diff_divide_eq_iff:
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
  1247
  "(z::'a::field) \<noteq> 0 \<Longrightarrow> x - y/z = (z*x - y)/z"
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
  1248
by(simp add:diff_divide_distrib nonzero_mult_divide_cancel_left)
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
  1249
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
  1250
lemma divide_diff_eq_iff:
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
  1251
  "(z::'a::field) \<noteq> 0 \<Longrightarrow> x/z - y = (x - z*y)/z"
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
  1252
by(simp add:diff_divide_distrib nonzero_mult_divide_cancel_left)
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
  1253
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
  1254
lemma nonzero_eq_divide_eq: "c\<noteq>0 ==> ((a::'a::field) = b/c) = (a*c = b)"
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
  1255
proof -
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
  1256
  assume [simp]: "c\<noteq>0"
23496
84e9216a6d0e removed redundant lemmas
nipkow
parents: 23483
diff changeset
  1257
  have "(a = b/c) = (a*c = (b/c)*c)" by simp
84e9216a6d0e removed redundant lemmas
nipkow
parents: 23483
diff changeset
  1258
  also have "... = (a*c = b)" by (simp add: divide_inverse mult_assoc)
23482
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
  1259
  finally show ?thesis .
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
  1260
qed
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
  1261
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
  1262
lemma nonzero_divide_eq_eq: "c\<noteq>0 ==> (b/c = (a::'a::field)) = (b = a*c)"
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
  1263
proof -
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
  1264
  assume [simp]: "c\<noteq>0"
23496
84e9216a6d0e removed redundant lemmas
nipkow
parents: 23483
diff changeset
  1265
  have "(b/c = a) = ((b/c)*c = a*c)"  by simp
84e9216a6d0e removed redundant lemmas
nipkow
parents: 23483
diff changeset
  1266
  also have "... = (b = a*c)"  by (simp add: divide_inverse mult_assoc) 
23482
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
  1267
  finally show ?thesis .
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
  1268
qed
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
  1269
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
  1270
lemma eq_divide_eq:
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
  1271
  "((a::'a::{field,division_by_zero}) = b/c) = (if c\<noteq>0 then a*c = b else a=0)"
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
  1272
by (simp add: nonzero_eq_divide_eq) 
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
  1273
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
  1274
lemma divide_eq_eq:
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
  1275
  "(b/c = (a::'a::{field,division_by_zero})) = (if c\<noteq>0 then b = a*c else a=0)"
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
  1276
by (force simp add: nonzero_divide_eq_eq) 
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
  1277
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
  1278
lemma divide_eq_imp: "(c::'a::{division_by_zero,field}) ~= 0 ==>
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
  1279
    b = a * c ==> b / c = a"
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
  1280
  by (subst divide_eq_eq, simp)
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
  1281
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
  1282
lemma eq_divide_imp: "(c::'a::{division_by_zero,field}) ~= 0 ==>
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
  1283
    a * c = b ==> a = b / c"
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
  1284
  by (subst eq_divide_eq, simp)
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
  1285
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
  1286
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
  1287
lemmas field_eq_simps = ring_simps
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
  1288
  (* pull / out*)
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
  1289
  add_divide_eq_iff divide_add_eq_iff
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
  1290
  diff_divide_eq_iff divide_diff_eq_iff
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
  1291
  (* multiply eqn *)
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
  1292
  nonzero_eq_divide_eq nonzero_divide_eq_eq
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
  1293
(* is added later:
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
  1294
  times_divide_eq_left times_divide_eq_right
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
  1295
*)
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
  1296
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
  1297
text{*An example:*}
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
  1298
lemma fixes a b c d e f :: "'a::field"
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
  1299
shows "\<lbrakk>a\<noteq>b; c\<noteq>d; e\<noteq>f \<rbrakk> \<Longrightarrow> ((a-b)*(c-d)*(e-f))/((c-d)*(e-f)*(a-b)) = 1"
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
  1300
apply(subgoal_tac "(c-d)*(e-f)*(a-b) \<noteq> 0")
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
  1301
 apply(simp add:field_eq_simps)
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
  1302
apply(simp)
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
  1303
done
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
  1304
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
  1305
16775
c1b87ef4a1c3 added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
avigad
parents: 16568
diff changeset
  1306
lemma diff_frac_eq: "(y::'a::field) ~= 0 ==> z ~= 0 ==>
c1b87ef4a1c3 added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
avigad
parents: 16568
diff changeset
  1307
    x / y - w / z = (x * z - w * y) / (y * z)"
23482
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
  1308
by (simp add:field_eq_simps times_divide_eq)
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
  1309
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
  1310
lemma frac_eq_eq: "(y::'a::field) ~= 0 ==> z ~= 0 ==>
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
  1311
    (x / y = w / z) = (x * z = w * y)"
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
  1312
by (simp add:field_eq_simps times_divide_eq)
14293
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1313
23389
aaca6a8e5414 tuned proofs: avoid implicit prems;
wenzelm
parents: 23326
diff changeset
  1314
14268
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
  1315
subsection {* Ordered Fields *}
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
  1316
14277
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
  1317
lemma positive_imp_inverse_positive: 
23482
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
  1318
assumes a_gt_0: "0 < a"  shows "0 < inverse (a::'a::ordered_field)"
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
  1319
proof -
14268
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
  1320
  have "0 < a * inverse a" 
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
  1321
    by (simp add: a_gt_0 [THEN order_less_imp_not_eq2] zero_less_one)
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
  1322
  thus "0 < inverse a" 
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
  1323
    by (simp add: a_gt_0 [THEN order_less_not_sym] zero_less_mult_iff)
23482
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
  1324
qed
14268
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
  1325
14277
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
  1326
lemma negative_imp_inverse_negative:
23482
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
  1327
  "a < 0 ==> inverse a < (0::'a::ordered_field)"
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
  1328
by (insert positive_imp_inverse_positive [of "-a"], 
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
  1329
    simp add: nonzero_inverse_minus_eq order_less_imp_not_eq)
14268
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
  1330
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
  1331
lemma inverse_le_imp_le:
23482
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
  1332
assumes invle: "inverse a \<le> inverse b" and apos:  "0 < a"
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
  1333
shows "b \<le> (a::'a::ordered_field)"
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
  1334
proof (rule classical)
14268
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
  1335
  assume "~ b \<le> a"
23482
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
  1336
  hence "a < b"  by (simp add: linorder_not_le)
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
  1337
  hence bpos: "0 < b"  by (blast intro: apos order_less_trans)
14268
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
  1338
  hence "a * inverse a \<le> a * inverse b"
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
  1339
    by (simp add: apos invle order_less_imp_le mult_left_mono)
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
  1340
  hence "(a * inverse a) * b \<le> (a * inverse b) * b"
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
  1341
    by (simp add: bpos order_less_imp_le mult_right_mono)
23482
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
  1342
  thus "b \<le> a"  by (simp add: mult_assoc apos bpos order_less_imp_not_eq2)
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
  1343
qed
14268
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
  1344
14277
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
  1345
lemma inverse_positive_imp_positive:
23482
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
  1346
assumes inv_gt_0: "0 < inverse a" and nz: "a \<noteq> 0"
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
  1347
shows "0 < (a::'a::ordered_field)"
23389
aaca6a8e5414 tuned proofs: avoid implicit prems;
wenzelm
parents: 23326
diff changeset
  1348
proof -
14277
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
  1349
  have "0 < inverse (inverse a)"
23389
aaca6a8e5414 tuned proofs: avoid implicit prems;
wenzelm
parents: 23326
diff changeset
  1350
    using inv_gt_0 by (rule positive_imp_inverse_positive)
14277
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
  1351
  thus "0 < a"
23389
aaca6a8e5414 tuned proofs: avoid implicit prems;
wenzelm
parents: 23326
diff changeset
  1352
    using nz by (simp add: nonzero_inverse_inverse_eq)
aaca6a8e5414 tuned proofs: avoid implicit prems;
wenzelm
parents: 23326
diff changeset
  1353
qed
14277
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
  1354
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
  1355
lemma inverse_positive_iff_positive [simp]:
23482
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
  1356
  "(0 < inverse a) = (0 < (a::'a::{ordered_field,division_by_zero}))"
21328
73bb86d0f483 dropped Inductive dependency
haftmann
parents: 21258
diff changeset
  1357
apply (cases "a = 0", simp)
14277
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
  1358
apply (blast intro: inverse_positive_imp_positive positive_imp_inverse_positive)
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
  1359
done
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
  1360
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
  1361
lemma inverse_negative_imp_negative:
23482
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
  1362
assumes inv_less_0: "inverse a < 0" and nz:  "a \<noteq> 0"
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
  1363
shows "a < (0::'a::ordered_field)"
23389
aaca6a8e5414 tuned proofs: avoid implicit prems;
wenzelm
parents: 23326
diff changeset
  1364
proof -
14277
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
  1365
  have "inverse (inverse a) < 0"
23389
aaca6a8e5414 tuned proofs: avoid implicit prems;
wenzelm
parents: 23326
diff changeset
  1366
    using inv_less_0 by (rule negative_imp_inverse_negative)
23482
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
  1367
  thus "a < 0" using nz by (simp add: nonzero_inverse_inverse_eq)
23389
aaca6a8e5414 tuned proofs: avoid implicit prems;
wenzelm
parents: 23326
diff changeset
  1368
qed
14277
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
  1369
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
  1370
lemma inverse_negative_iff_negative [simp]:
23482
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
  1371
  "(inverse a < 0) = (a < (0::'a::{ordered_field,division_by_zero}))"
21328
73bb86d0f483 dropped Inductive dependency
haftmann
parents: 21258
diff changeset
  1372
apply (cases "a = 0", simp)
14277
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
  1373
apply (blast intro: inverse_negative_imp_negative negative_imp_inverse_negative)
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
  1374
done
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
  1375
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
  1376
lemma inverse_nonnegative_iff_nonnegative [simp]:
23482
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
  1377
  "(0 \<le> inverse a) = (0 \<le> (a::'a::{ordered_field,division_by_zero}))"
14277
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
  1378
by (simp add: linorder_not_less [symmetric])
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
  1379
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
  1380
lemma inverse_nonpositive_iff_nonpositive [simp]:
23482
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
  1381
  "(inverse a \<le> 0) = (a \<le> (0::'a::{ordered_field,division_by_zero}))"
14277
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
  1382
by (simp add: linorder_not_less [symmetric])
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
  1383
23406
167b53019d6f added theorems nonzero_mult_divide_cancel_right' nonzero_mult_divide_cancel_left' ordered_field_no_lb ordered_field_no_ub
chaieb
parents: 23400
diff changeset
  1384
lemma ordered_field_no_lb: "\<forall> x. \<exists>y. y < (x::'a::ordered_field)"
167b53019d6f added theorems nonzero_mult_divide_cancel_right' nonzero_mult_divide_cancel_left' ordered_field_no_lb ordered_field_no_ub
chaieb
parents: 23400
diff changeset
  1385
proof
167b53019d6f added theorems nonzero_mult_divide_cancel_right' nonzero_mult_divide_cancel_left' ordered_field_no_lb ordered_field_no_ub
chaieb
parents: 23400
diff changeset
  1386
  fix x::'a
167b53019d6f added theorems nonzero_mult_divide_cancel_right' nonzero_mult_divide_cancel_left' ordered_field_no_lb ordered_field_no_ub
chaieb
parents: 23400
diff changeset
  1387
  have m1: "- (1::'a) < 0" by simp
167b53019d6f added theorems nonzero_mult_divide_cancel_right' nonzero_mult_divide_cancel_left' ordered_field_no_lb ordered_field_no_ub
chaieb
parents: 23400
diff changeset
  1388
  from add_strict_right_mono[OF m1, where c=x] 
167b53019d6f added theorems nonzero_mult_divide_cancel_right' nonzero_mult_divide_cancel_left' ordered_field_no_lb ordered_field_no_ub
chaieb
parents: 23400
diff changeset
  1389
  have "(- 1) + x < x" by simp
167b53019d6f added theorems nonzero_mult_divide_cancel_right' nonzero_mult_divide_cancel_left' ordered_field_no_lb ordered_field_no_ub
chaieb
parents: 23400
diff changeset
  1390
  thus "\<exists>y. y < x" by blast
167b53019d6f added theorems nonzero_mult_divide_cancel_right' nonzero_mult_divide_cancel_left' ordered_field_no_lb ordered_field_no_ub
chaieb
parents: 23400
diff changeset
  1391
qed
167b53019d6f added theorems nonzero_mult_divide_cancel_right' nonzero_mult_divide_cancel_left' ordered_field_no_lb ordered_field_no_ub
chaieb
parents: 23400
diff changeset
  1392
167b53019d6f added theorems nonzero_mult_divide_cancel_right' nonzero_mult_divide_cancel_left' ordered_field_no_lb ordered_field_no_ub
chaieb
parents: 23400
diff changeset
  1393
lemma ordered_field_no_ub: "\<forall> x. \<exists>y. y > (x::'a::ordered_field)"
167b53019d6f added theorems nonzero_mult_divide_cancel_right' nonzero_mult_divide_cancel_left' ordered_field_no_lb ordered_field_no_ub
chaieb
parents: 23400
diff changeset
  1394
proof
167b53019d6f added theorems nonzero_mult_divide_cancel_right' nonzero_mult_divide_cancel_left' ordered_field_no_lb ordered_field_no_ub
chaieb
parents: 23400
diff changeset
  1395
  fix x::'a
167b53019d6f added theorems nonzero_mult_divide_cancel_right' nonzero_mult_divide_cancel_left' ordered_field_no_lb ordered_field_no_ub
chaieb
parents: 23400
diff changeset
  1396
  have m1: " (1::'a) > 0" by simp
167b53019d6f added theorems nonzero_mult_divide_cancel_right' nonzero_mult_divide_cancel_left' ordered_field_no_lb ordered_field_no_ub
chaieb
parents: 23400
diff changeset
  1397
  from add_strict_right_mono[OF m1, where c=x] 
167b53019d6f added theorems nonzero_mult_divide_cancel_right' nonzero_mult_divide_cancel_left' ordered_field_no_lb ordered_field_no_ub
chaieb
parents: 23400
diff changeset
  1398
  have "1 + x > x" by simp
167b53019d6f added theorems nonzero_mult_divide_cancel_right' nonzero_mult_divide_cancel_left' ordered_field_no_lb ordered_field_no_ub
chaieb
parents: 23400
diff changeset
  1399
  thus "\<exists>y. y > x" by blast
167b53019d6f added theorems nonzero_mult_divide_cancel_right' nonzero_mult_divide_cancel_left' ordered_field_no_lb ordered_field_no_ub
chaieb
parents: 23400
diff changeset
  1400
qed
14277
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
  1401
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
  1402
subsection{*Anti-Monotonicity of @{term inverse}*}
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
  1403
14268
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
  1404
lemma less_imp_inverse_less:
23482
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
  1405
assumes less: "a < b" and apos:  "0 < a"
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
  1406
shows "inverse b < inverse (a::'a::ordered_field)"
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
  1407
proof (rule ccontr)
14268
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
  1408
  assume "~ inverse b < inverse a"
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
  1409
  hence "inverse a \<le> inverse b"
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
  1410
    by (simp add: linorder_not_less)
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
  1411
  hence "~ (a < b)"
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
  1412
    by (simp add: linorder_not_less inverse_le_imp_le [OF _ apos])
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
  1413
  thus False
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
  1414
    by (rule notE [OF _ less])
23482
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
  1415
qed
14268
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
  1416
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
  1417
lemma inverse_less_imp_less:
23482
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
  1418
  "[|inverse a < inverse b; 0 < a|] ==> b < (a::'a::ordered_field)"
14268
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
  1419
apply (simp add: order_less_le [of "inverse a"] order_less_le [of "b"])
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
  1420
apply (force dest!: inverse_le_imp_le nonzero_inverse_eq_imp_eq) 
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
  1421
done
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
  1422
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
  1423
text{*Both premises are essential. Consider -1 and 1.*}
24286
7619080e49f0 ATP blacklisting is now in theory data, attribute noatp
paulson
parents: 23879
diff changeset
  1424
lemma inverse_less_iff_less [simp,noatp]:
23482
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
  1425
  "[|0 < a; 0 < b|] ==> (inverse a < inverse b) = (b < (a::'a::ordered_field))"
14268
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
  1426
by (blast intro: less_imp_inverse_less dest: inverse_less_imp_less) 
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
  1427
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
  1428
lemma le_imp_inverse_le:
23482
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
  1429
  "[|a \<le> b; 0 < a|] ==> inverse b \<le> inverse (a::'a::ordered_field)"
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
  1430
by (force simp add: order_le_less less_imp_inverse_less)
14268
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
  1431
24286
7619080e49f0 ATP blacklisting is now in theory data, attribute noatp
paulson
parents: 23879
diff changeset
  1432
lemma inverse_le_iff_le [simp,noatp]:
23482
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
  1433
 "[|0 < a; 0 < b|] ==> (inverse a \<le> inverse b) = (b \<le> (a::'a::ordered_field))"
14268
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
  1434
by (blast intro: le_imp_inverse_le dest: inverse_le_imp_le) 
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
  1435
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
  1436
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
  1437
text{*These results refer to both operands being negative.  The opposite-sign
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
  1438
case is trivial, since inverse preserves signs.*}
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
  1439
lemma inverse_le_imp_le_neg:
23482
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
  1440
  "[|inverse a \<le> inverse b; b < 0|] ==> b \<le> (a::'a::ordered_field)"
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
  1441
apply (rule classical) 
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
  1442
apply (subgoal_tac "a < 0") 
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
  1443
 prefer 2 apply (force simp add: linorder_not_le intro: order_less_trans) 
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
  1444
apply (insert inverse_le_imp_le [of "-b" "-a"])
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
  1445
apply (simp add: order_less_imp_not_eq nonzero_inverse_minus_eq) 
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
  1446
done
14268
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
  1447
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
  1448
lemma less_imp_inverse_less_neg:
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
  1449
   "[|a < b; b < 0|] ==> inverse b < inverse (a::'a::ordered_field)"
23482
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
  1450
apply (subgoal_tac "a < 0") 
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
  1451
 prefer 2 apply (blast intro: order_less_trans) 
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
  1452
apply (insert less_imp_inverse_less [of "-b" "-a"])
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
  1453
apply (simp add: order_less_imp_not_eq nonzero_inverse_minus_eq) 
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
  1454
done
14268
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
  1455
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
  1456
lemma inverse_less_imp_less_neg:
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
  1457
   "[|inverse a < inverse b; b < 0|] ==> b < (a::'a::ordered_field)"
23482
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
  1458
apply (rule classical) 
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
  1459
apply (subgoal_tac "a < 0") 
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
  1460
 prefer 2
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
  1461
 apply (force simp add: linorder_not_less intro: order_le_less_trans) 
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
  1462
apply (insert inverse_less_imp_less [of "-b" "-a"])
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
  1463
apply (simp add: order_less_imp_not_eq nonzero_inverse_minus_eq) 
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
  1464
done
14268
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
  1465
24286
7619080e49f0 ATP blacklisting is now in theory data, attribute noatp
paulson
parents: 23879
diff changeset
  1466
lemma inverse_less_iff_less_neg [simp,noatp]:
23482
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
  1467
  "[|a < 0; b < 0|] ==> (inverse a < inverse b) = (b < (a::'a::ordered_field))"
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
  1468
apply (insert inverse_less_iff_less [of "-b" "-a"])
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
  1469
apply (simp del: inverse_less_iff_less 
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
  1470
            add: order_less_imp_not_eq nonzero_inverse_minus_eq)
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
  1471
done
14268
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
  1472
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
  1473
lemma le_imp_inverse_le_neg:
23482
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
  1474
  "[|a \<le> b; b < 0|] ==> inverse b \<le> inverse (a::'a::ordered_field)"
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
  1475
by (force simp add: order_le_less less_imp_inverse_less_neg)
14268
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
  1476
24286
7619080e49f0 ATP blacklisting is now in theory data, attribute noatp
paulson
parents: 23879
diff changeset
  1477
lemma inverse_le_iff_le_neg [simp,noatp]:
23482
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
  1478
 "[|a < 0; b < 0|] ==> (inverse a \<le> inverse b) = (b \<le> (a::'a::ordered_field))"
14268
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
  1479
by (blast intro: le_imp_inverse_le_neg dest: inverse_le_imp_le_neg) 
14265
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
  1480
14277
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
  1481
14365
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14353
diff changeset
  1482
subsection{*Inverses and the Number One*}
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14353
diff changeset
  1483
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14353
diff changeset
  1484
lemma one_less_inverse_iff:
23482
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
  1485
  "(1 < inverse x) = (0 < x & x < (1::'a::{ordered_field,division_by_zero}))"
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
  1486
proof cases
14365
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14353
diff changeset
  1487
  assume "0 < x"
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14353
diff changeset
  1488
    with inverse_less_iff_less [OF zero_less_one, of x]
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14353
diff changeset
  1489
    show ?thesis by simp
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14353
diff changeset
  1490
next
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14353
diff changeset
  1491
  assume notless: "~ (0 < x)"
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14353
diff changeset
  1492
  have "~ (1 < inverse x)"
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14353
diff changeset
  1493
  proof
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14353
diff changeset
  1494
    assume "1 < inverse x"
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14353
diff changeset
  1495
    also with notless have "... \<le> 0" by (simp add: linorder_not_less)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14353
diff changeset
  1496
    also have "... < 1" by (rule zero_less_one) 
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14353
diff changeset
  1497
    finally show False by auto
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14353
diff changeset
  1498
  qed
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14353
diff changeset
  1499
  with notless show ?thesis by simp
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14353
diff changeset
  1500
qed
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14353
diff changeset
  1501
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14353
diff changeset
  1502
lemma inverse_eq_1_iff [simp]:
23482
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
  1503
  "(inverse x = 1) = (x = (1::'a::{field,division_by_zero}))"
14365
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14353
diff changeset
  1504
by (insert inverse_eq_iff_eq [of x 1], simp) 
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14353
diff changeset
  1505
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14353
diff changeset
  1506
lemma one_le_inverse_iff:
23482
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
  1507
  "(1 \<le> inverse x) = (0 < x & x \<le> (1::'a::{ordered_field,division_by_zero}))"
14365
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14353
diff changeset
  1508
by (force simp add: order_le_less one_less_inverse_iff zero_less_one 
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14353
diff changeset
  1509
                    eq_commute [of 1]) 
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14353
diff changeset
  1510
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14353
diff changeset
  1511
lemma inverse_less_1_iff:
23482
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
  1512
  "(inverse x < 1) = (x \<le> 0 | 1 < (x::'a::{ordered_field,division_by_zero}))"
14365
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14353
diff changeset
  1513
by (simp add: linorder_not_le [symmetric] one_le_inverse_iff) 
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14353
diff changeset
  1514
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14353
diff changeset
  1515
lemma inverse_le_1_iff:
23482
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
  1516
  "(inverse x \<le> 1) = (x \<le> 0 | 1 \<le> (x::'a::{ordered_field,division_by_zero}))"
14365
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14353
diff changeset
  1517
by (simp add: linorder_not_less [symmetric] one_less_inverse_iff) 
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14353
diff changeset
  1518
23389
aaca6a8e5414 tuned proofs: avoid implicit prems;
wenzelm
parents: 23326
diff changeset
  1519
14288
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1520
subsection{*Simplification of Inequalities Involving Literal Divisors*}
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1521
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1522
lemma pos_le_divide_eq: "0 < (c::'a::ordered_field) ==> (a \<le> b/c) = (a*c \<le> b)"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1523
proof -
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1524
  assume less: "0<c"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1525
  hence "(a \<le> b/c) = (a*c \<le> (b/c)*c)"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1526
    by (simp add: mult_le_cancel_right order_less_not_sym [OF less])
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1527
  also have "... = (a*c \<le> b)"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1528
    by (simp add: order_less_imp_not_eq2 [OF less] divide_inverse mult_assoc) 
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1529
  finally show ?thesis .
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1530
qed
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1531
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1532
lemma neg_le_divide_eq: "c < (0::'a::ordered_field) ==> (a \<le> b/c) = (b \<le> a*c)"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1533
proof -
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1534
  assume less: "c<0"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1535
  hence "(a \<le> b/c) = ((b/c)*c \<le> a*c)"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1536
    by (simp add: mult_le_cancel_right order_less_not_sym [OF less])
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1537
  also have "... = (b \<le> a*c)"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1538
    by (simp add: order_less_imp_not_eq [OF less] divide_inverse mult_assoc) 
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1539
  finally show ?thesis .
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1540
qed
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1541
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1542
lemma le_divide_eq:
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1543
  "(a \<le> b/c) = 
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1544
   (if 0 < c then a*c \<le> b
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1545
             else if c < 0 then b \<le> a*c
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1546
             else  a \<le> (0::'a::{ordered_field,division_by_zero}))"
21328
73bb86d0f483 dropped Inductive dependency
haftmann
parents: 21258
diff changeset
  1547
apply (cases "c=0", simp) 
14288
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1548
apply (force simp add: pos_le_divide_eq neg_le_divide_eq linorder_neq_iff) 
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1549
done
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1550
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1551
lemma pos_divide_le_eq: "0 < (c::'a::ordered_field) ==> (b/c \<le> a) = (b \<le> a*c)"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1552
proof -
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1553
  assume less: "0<c"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1554
  hence "(b/c \<le> a) = ((b/c)*c \<le> a*c)"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1555
    by (simp add: mult_le_cancel_right order_less_not_sym [OF less])
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1556
  also have "... = (b \<le> a*c)"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1557
    by (simp add: order_less_imp_not_eq2 [OF less] divide_inverse mult_assoc) 
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1558
  finally show ?thesis .
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1559
qed
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1560
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1561
lemma neg_divide_le_eq: "c < (0::'a::ordered_field) ==> (b/c \<le> a) = (a*c \<le> b)"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1562
proof -
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1563
  assume less: "c<0"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1564
  hence "(b/c \<le> a) = (a*c \<le> (b/c)*c)"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1565
    by (simp add: mult_le_cancel_right order_less_not_sym [OF less])
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1566
  also have "... = (a*c \<le> b)"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1567
    by (simp add: order_less_imp_not_eq [OF less] divide_inverse mult_assoc) 
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1568
  finally show ?thesis .
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1569
qed
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1570
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1571
lemma divide_le_eq:
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1572
  "(b/c \<le> a) = 
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1573
   (if 0 < c then b \<le> a*c
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1574
             else if c < 0 then a*c \<le> b
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1575
             else 0 \<le> (a::'a::{ordered_field,division_by_zero}))"
21328
73bb86d0f483 dropped Inductive dependency
haftmann
parents: 21258
diff changeset
  1576
apply (cases "c=0", simp) 
14288
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1577
apply (force simp add: pos_divide_le_eq neg_divide_le_eq linorder_neq_iff) 
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1578
done
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1579
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1580
lemma pos_less_divide_eq:
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1581
     "0 < (c::'a::ordered_field) ==> (a < b/c) = (a*c < b)"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1582
proof -
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1583
  assume less: "0<c"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1584
  hence "(a < b/c) = (a*c < (b/c)*c)"
15234
ec91a90c604e simplification tweaks for better arithmetic reasoning
paulson
parents: 15229
diff changeset
  1585
    by (simp add: mult_less_cancel_right_disj order_less_not_sym [OF less])
14288
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1586
  also have "... = (a*c < b)"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1587
    by (simp add: order_less_imp_not_eq2 [OF less] divide_inverse mult_assoc) 
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1588
  finally show ?thesis .
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1589
qed
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1590
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1591
lemma neg_less_divide_eq:
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1592
 "c < (0::'a::ordered_field) ==> (a < b/c) = (b < a*c)"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1593
proof -
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1594
  assume less: "c<0"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1595
  hence "(a < b/c) = ((b/c)*c < a*c)"
15234
ec91a90c604e simplification tweaks for better arithmetic reasoning
paulson
parents: 15229
diff changeset
  1596
    by (simp add: mult_less_cancel_right_disj order_less_not_sym [OF less])
14288
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1597
  also have "... = (b < a*c)"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1598
    by (simp add: order_less_imp_not_eq [OF less] divide_inverse mult_assoc) 
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1599
  finally show ?thesis .
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1600
qed
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1601
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1602
lemma less_divide_eq:
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1603
  "(a < b/c) = 
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1604
   (if 0 < c then a*c < b
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1605
             else if c < 0 then b < a*c
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1606
             else  a < (0::'a::{ordered_field,division_by_zero}))"
21328
73bb86d0f483 dropped Inductive dependency
haftmann
parents: 21258
diff changeset
  1607
apply (cases "c=0", simp) 
14288
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1608
apply (force simp add: pos_less_divide_eq neg_less_divide_eq linorder_neq_iff) 
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1609
done
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1610
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1611
lemma pos_divide_less_eq:
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1612
     "0 < (c::'a::ordered_field) ==> (b/c < a) = (b < a*c)"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1613
proof -
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1614
  assume less: "0<c"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1615
  hence "(b/c < a) = ((b/c)*c < a*c)"
15234
ec91a90c604e simplification tweaks for better arithmetic reasoning
paulson
parents: 15229
diff changeset
  1616
    by (simp add: mult_less_cancel_right_disj order_less_not_sym [OF less])
14288
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1617
  also have "... = (b < a*c)"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1618
    by (simp add: order_less_imp_not_eq2 [OF less] divide_inverse mult_assoc) 
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1619
  finally show ?thesis .
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1620
qed
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1621
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1622
lemma neg_divide_less_eq:
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1623
 "c < (0::'a::ordered_field) ==> (b/c < a) = (a*c < b)"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1624
proof -
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1625
  assume less: "c<0"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1626
  hence "(b/c < a) = (a*c < (b/c)*c)"
15234
ec91a90c604e simplification tweaks for better arithmetic reasoning
paulson
parents: 15229
diff changeset
  1627
    by (simp add: mult_less_cancel_right_disj order_less_not_sym [OF less])
14288
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1628
  also have "... = (a*c < b)"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1629
    by (simp add: order_less_imp_not_eq [OF less] divide_inverse mult_assoc) 
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1630
  finally show ?thesis .
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1631
qed
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1632
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1633
lemma divide_less_eq:
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1634
  "(b/c < a) = 
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1635
   (if 0 < c then b < a*c
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1636
             else if c < 0 then a*c < b
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1637
             else 0 < (a::'a::{ordered_field,division_by_zero}))"
21328
73bb86d0f483 dropped Inductive dependency
haftmann
parents: 21258
diff changeset
  1638
apply (cases "c=0", simp) 
14288
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1639
apply (force simp add: pos_divide_less_eq neg_divide_less_eq linorder_neq_iff) 
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1640
done
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1641
23482
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
  1642
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
  1643
subsection{*Field simplification*}
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
  1644
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
  1645
text{* Lemmas @{text field_simps} multiply with denominators in
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
  1646
in(equations) if they can be proved to be non-zero (for equations) or
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
  1647
positive/negative (for inequations). *}
14288
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1648
23482
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
  1649
lemmas field_simps = field_eq_simps
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
  1650
  (* multiply ineqn *)
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
  1651
  pos_divide_less_eq neg_divide_less_eq
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
  1652
  pos_less_divide_eq neg_less_divide_eq
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
  1653
  pos_divide_le_eq neg_divide_le_eq
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
  1654
  pos_le_divide_eq neg_le_divide_eq
14288
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1655
23482
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
  1656
text{* Lemmas @{text sign_simps} is a first attempt to automate proofs
23483
a9356b40fbd3 tex problem fixed
nipkow
parents: 23482
diff changeset
  1657
of positivity/negativity needed for @{text field_simps}. Have not added @{text
23482
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
  1658
sign_simps} to @{text field_simps} because the former can lead to case
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
  1659
explosions. *}
14288
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1660
23482
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
  1661
lemmas sign_simps = group_simps
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
  1662
  zero_less_mult_iff  mult_less_0_iff
14288
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1663
23482
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
  1664
(* Only works once linear arithmetic is installed:
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
  1665
text{*An example:*}
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
  1666
lemma fixes a b c d e f :: "'a::ordered_field"
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
  1667
shows "\<lbrakk>a>b; c<d; e<f; 0 < u \<rbrakk> \<Longrightarrow>
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
  1668
 ((a-b)*(c-d)*(e-f))/((c-d)*(e-f)*(a-b)) <
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
  1669
 ((e-f)*(a-b)*(c-d))/((e-f)*(a-b)*(c-d)) + u"
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
  1670
apply(subgoal_tac "(c-d)*(e-f)*(a-b) > 0")
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
  1671
 prefer 2 apply(simp add:sign_simps)
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
  1672
apply(subgoal_tac "(c-d)*(e-f)*(a-b)*u > 0")
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
  1673
 prefer 2 apply(simp add:sign_simps)
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
  1674
apply(simp add:field_simps)
16775
c1b87ef4a1c3 added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
avigad
parents: 16568
diff changeset
  1675
done
23482
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
  1676
*)
16775
c1b87ef4a1c3 added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
avigad
parents: 16568
diff changeset
  1677
23389
aaca6a8e5414 tuned proofs: avoid implicit prems;
wenzelm
parents: 23326
diff changeset
  1678
16775
c1b87ef4a1c3 added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
avigad
parents: 16568
diff changeset
  1679
subsection{*Division and Signs*}
c1b87ef4a1c3 added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
avigad
parents: 16568
diff changeset
  1680
c1b87ef4a1c3 added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
avigad
parents: 16568
diff changeset
  1681
lemma zero_less_divide_iff:
c1b87ef4a1c3 added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
avigad
parents: 16568
diff changeset
  1682
     "((0::'a::{ordered_field,division_by_zero}) < a/b) = (0 < a & 0 < b | a < 0 & b < 0)"
c1b87ef4a1c3 added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
avigad
parents: 16568
diff changeset
  1683
by (simp add: divide_inverse zero_less_mult_iff)
c1b87ef4a1c3 added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
avigad
parents: 16568
diff changeset
  1684
c1b87ef4a1c3 added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
avigad
parents: 16568
diff changeset
  1685
lemma divide_less_0_iff:
c1b87ef4a1c3 added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
avigad
parents: 16568
diff changeset
  1686
     "(a/b < (0::'a::{ordered_field,division_by_zero})) = 
c1b87ef4a1c3 added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
avigad
parents: 16568
diff changeset
  1687
      (0 < a & b < 0 | a < 0 & 0 < b)"
c1b87ef4a1c3 added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
avigad
parents: 16568
diff changeset
  1688
by (simp add: divide_inverse mult_less_0_iff)
c1b87ef4a1c3 added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
avigad
parents: 16568
diff changeset
  1689
c1b87ef4a1c3 added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
avigad
parents: 16568
diff changeset
  1690
lemma zero_le_divide_iff:
c1b87ef4a1c3 added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
avigad
parents: 16568
diff changeset
  1691
     "((0::'a::{ordered_field,division_by_zero}) \<le> a/b) =
c1b87ef4a1c3 added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
avigad
parents: 16568
diff changeset
  1692
      (0 \<le> a & 0 \<le> b | a \<le> 0 & b \<le> 0)"
c1b87ef4a1c3 added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
avigad
parents: 16568
diff changeset
  1693
by (simp add: divide_inverse zero_le_mult_iff)
c1b87ef4a1c3 added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
avigad
parents: 16568
diff changeset
  1694
c1b87ef4a1c3 added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
avigad
parents: 16568
diff changeset
  1695
lemma divide_le_0_iff:
c1b87ef4a1c3 added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
avigad
parents: 16568
diff changeset
  1696
     "(a/b \<le> (0::'a::{ordered_field,division_by_zero})) =
c1b87ef4a1c3 added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
avigad
parents: 16568
diff changeset
  1697
      (0 \<le> a & b \<le> 0 | a \<le> 0 & 0 \<le> b)"
c1b87ef4a1c3 added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
avigad
parents: 16568
diff changeset
  1698
by (simp add: divide_inverse mult_le_0_iff)
c1b87ef4a1c3 added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
avigad
parents: 16568
diff changeset
  1699
24286
7619080e49f0 ATP blacklisting is now in theory data, attribute noatp
paulson
parents: 23879
diff changeset
  1700
lemma divide_eq_0_iff [simp,noatp]:
16775
c1b87ef4a1c3 added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
avigad
parents: 16568
diff changeset
  1701
     "(a/b = 0) = (a=0 | b=(0::'a::{field,division_by_zero}))"
23482
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
  1702
by (simp add: divide_inverse)
16775
c1b87ef4a1c3 added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
avigad
parents: 16568
diff changeset
  1703
23482
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
  1704
lemma divide_pos_pos:
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
  1705
  "0 < (x::'a::ordered_field) ==> 0 < y ==> 0 < x / y"
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
  1706
by(simp add:field_simps)
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
  1707
16775
c1b87ef4a1c3 added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
avigad
parents: 16568
diff changeset
  1708
23482
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
  1709
lemma divide_nonneg_pos:
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
  1710
  "0 <= (x::'a::ordered_field) ==> 0 < y ==> 0 <= x / y"
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
  1711
by(simp add:field_simps)
16775
c1b87ef4a1c3 added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
avigad
parents: 16568
diff changeset
  1712
23482
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
  1713
lemma divide_neg_pos:
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
  1714
  "(x::'a::ordered_field) < 0 ==> 0 < y ==> x / y < 0"
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
  1715
by(simp add:field_simps)
16775
c1b87ef4a1c3 added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
avigad
parents: 16568
diff changeset
  1716
23482
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
  1717
lemma divide_nonpos_pos:
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
  1718
  "(x::'a::ordered_field) <= 0 ==> 0 < y ==> x / y <= 0"
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
  1719
by(simp add:field_simps)
16775
c1b87ef4a1c3 added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
avigad
parents: 16568
diff changeset
  1720
23482
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
  1721
lemma divide_pos_neg:
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
  1722
  "0 < (x::'a::ordered_field) ==> y < 0 ==> x / y < 0"
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
  1723
by(simp add:field_simps)
16775
c1b87ef4a1c3 added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
avigad
parents: 16568
diff changeset
  1724
23482
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
  1725
lemma divide_nonneg_neg:
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
  1726
  "0 <= (x::'a::ordered_field) ==> y < 0 ==> x / y <= 0" 
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
  1727
by(simp add:field_simps)
16775
c1b87ef4a1c3 added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
avigad
parents: 16568
diff changeset
  1728
23482
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
  1729
lemma divide_neg_neg:
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
  1730
  "(x::'a::ordered_field) < 0 ==> y < 0 ==> 0 < x / y"
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
  1731
by(simp add:field_simps)
16775
c1b87ef4a1c3 added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
avigad
parents: 16568
diff changeset
  1732
23482
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
  1733
lemma divide_nonpos_neg:
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
  1734
  "(x::'a::ordered_field) <= 0 ==> y < 0 ==> 0 <= x / y"
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
  1735
by(simp add:field_simps)
15234
ec91a90c604e simplification tweaks for better arithmetic reasoning
paulson
parents: 15229
diff changeset
  1736
23389
aaca6a8e5414 tuned proofs: avoid implicit prems;
wenzelm
parents: 23326
diff changeset
  1737
14288
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1738
subsection{*Cancellation Laws for Division*}
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1739
24286
7619080e49f0 ATP blacklisting is now in theory data, attribute noatp
paulson
parents: 23879
diff changeset
  1740
lemma divide_cancel_right [simp,noatp]:
14288
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1741
     "(a/c = b/c) = (c = 0 | a = (b::'a::{field,division_by_zero}))"
23482
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
  1742
apply (cases "c=0", simp)
23496
84e9216a6d0e removed redundant lemmas
nipkow
parents: 23483
diff changeset
  1743
apply (simp add: divide_inverse)
14288
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1744
done
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1745
24286
7619080e49f0 ATP blacklisting is now in theory data, attribute noatp
paulson
parents: 23879
diff changeset
  1746
lemma divide_cancel_left [simp,noatp]:
14288
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1747
     "(c/a = c/b) = (c = 0 | a = (b::'a::{field,division_by_zero}))" 
23482
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
  1748
apply (cases "c=0", simp)
23496
84e9216a6d0e removed redundant lemmas
nipkow
parents: 23483
diff changeset
  1749
apply (simp add: divide_inverse)
14288
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1750
done
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
  1751
23389
aaca6a8e5414 tuned proofs: avoid implicit prems;
wenzelm
parents: 23326
diff changeset
  1752
14353
79f9fbef9106 Added lemmas to Ring_and_Field with slightly modified simplification rules
paulson
parents: 14348
diff changeset
  1753
subsection {* Division and the Number One *}
79f9fbef9106 Added lemmas to Ring_and_Field with slightly modified simplification rules
paulson
parents: 14348
diff changeset
  1754
79f9fbef9106 Added lemmas to Ring_and_Field with slightly modified simplification rules
paulson
parents: 14348
diff changeset
  1755
text{*Simplify expressions equated with 1*}
24286
7619080e49f0 ATP blacklisting is now in theory data, attribute noatp
paulson
parents: 23879
diff changeset
  1756
lemma divide_eq_1_iff [simp,noatp]:
14353
79f9fbef9106 Added lemmas to Ring_and_Field with slightly modified simplification rules
paulson
parents: 14348
diff changeset
  1757
     "(a/b = 1) = (b \<noteq> 0 & a = (b::'a::{field,division_by_zero}))"
23482
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
  1758
apply (cases "b=0", simp)
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
  1759
apply (simp add: right_inverse_eq)
14353
79f9fbef9106 Added lemmas to Ring_and_Field with slightly modified simplification rules
paulson
parents: 14348
diff changeset
  1760
done
79f9fbef9106 Added lemmas to Ring_and_Field with slightly modified simplification rules
paulson
parents: 14348
diff changeset
  1761
24286
7619080e49f0 ATP blacklisting is now in theory data, attribute noatp
paulson
parents: 23879
diff changeset
  1762
lemma one_eq_divide_iff [simp,noatp]:
14353
79f9fbef9106 Added lemmas to Ring_and_Field with slightly modified simplification rules
paulson
parents: 14348
diff changeset
  1763
     "(1 = a/b) = (b \<noteq> 0 & a = (b::'a::{field,division_by_zero}))"
23482
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
  1764
by (simp add: eq_commute [of 1])
14353
79f9fbef9106 Added lemmas to Ring_and_Field with slightly modified simplification rules
paulson
parents: 14348
diff changeset
  1765
24286
7619080e49f0 ATP blacklisting is now in theory data, attribute noatp
paulson
parents: 23879
diff changeset
  1766
lemma zero_eq_1_divide_iff [simp,noatp]:
14353
79f9fbef9106 Added lemmas to Ring_and_Field with slightly modified simplification rules
paulson
parents: 14348
diff changeset
  1767
     "((0::'a::{ordered_field,division_by_zero}) = 1/a) = (a = 0)"
23482
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
  1768
apply (cases "a=0", simp)
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
  1769
apply (auto simp add: nonzero_eq_divide_eq)
14353
79f9fbef9106 Added lemmas to Ring_and_Field with slightly modified simplification rules
paulson
parents: 14348
diff changeset
  1770
done
79f9fbef9106 Added lemmas to Ring_and_Field with slightly modified simplification rules
paulson
parents: 14348
diff changeset
  1771
24286
7619080e49f0 ATP blacklisting is now in theory data, attribute noatp
paulson
parents: 23879
diff changeset
  1772
lemma one_divide_eq_0_iff [simp,noatp]:
14353
79f9fbef9106 Added lemmas to Ring_and_Field with slightly modified simplification rules
paulson
parents: 14348
diff changeset
  1773
     "(1/a = (0::'a::{ordered_field,division_by_zero})) = (a = 0)"
23482
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
  1774
apply (cases "a=0", simp)
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
  1775
apply (insert zero_neq_one [THEN not_sym])
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
  1776
apply (auto simp add: nonzero_divide_eq_eq)
14353
79f9fbef9106 Added lemmas to Ring_and_Field with slightly modified simplification rules
paulson
parents: 14348
diff changeset
  1777
done
79f9fbef9106 Added lemmas to Ring_and_Field with slightly modified simplification rules
paulson
parents: 14348
diff changeset
  1778
79f9fbef9106 Added lemmas to Ring_and_Field with slightly modified simplification rules
paulson
parents: 14348
diff changeset
  1779
text{*Simplify expressions such as @{text "0 < 1/x"} to @{text "0 < x"}*}
18623
9a5419d5ca01 simplified the special-case simprules
paulson
parents: 17085
diff changeset
  1780
lemmas zero_less_divide_1_iff = zero_less_divide_iff [of 1, simplified]
9a5419d5ca01 simplified the special-case simprules
paulson
parents: 17085
diff changeset
  1781
lemmas divide_less_0_1_iff = divide_less_0_iff [of 1, simplified]
9a5419d5ca01 simplified the special-case simprules
paulson
parents: 17085
diff changeset
  1782
lemmas zero_le_divide_1_iff = zero_le_divide_iff [of 1, simplified]
9a5419d5ca01 simplified the special-case simprules
paulson
parents: 17085
diff changeset
  1783
lemmas divide_le_0_1_iff = divide_le_0_iff [of 1, simplified]
17085
5b57f995a179 more simprules now have names
paulson
parents: 16775
diff changeset
  1784
5b57f995a179 more simprules now have names
paulson
parents: 16775
diff changeset
  1785
declare zero_less_divide_1_iff [simp]
24286
7619080e49f0 ATP blacklisting is now in theory data, attribute noatp
paulson
parents: 23879
diff changeset
  1786
declare divide_less_0_1_iff [simp,noatp]
17085
5b57f995a179 more simprules now have names
paulson
parents: 16775
diff changeset
  1787
declare zero_le_divide_1_iff [simp]
24286
7619080e49f0 ATP blacklisting is now in theory data, attribute noatp
paulson
parents: 23879
diff changeset
  1788
declare divide_le_0_1_iff [simp,noatp]
14353
79f9fbef9106 Added lemmas to Ring_and_Field with slightly modified simplification rules
paulson
parents: 14348
diff changeset
  1789
23389
aaca6a8e5414 tuned proofs: avoid implicit prems;
wenzelm
parents: 23326
diff changeset
  1790
14293
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1791
subsection {* Ordering Rules for Division *}
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1792
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1793
lemma divide_strict_right_mono:
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1794
     "[|a < b; 0 < c|] ==> a / c < b / (c::'a::ordered_field)"
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1795
by (simp add: order_less_imp_not_eq2 divide_inverse mult_strict_right_mono 
23482
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
  1796
              positive_imp_inverse_positive)
14293
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1797
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1798
lemma divide_right_mono:
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1799
     "[|a \<le> b; 0 \<le> c|] ==> a/c \<le> b/(c::'a::{ordered_field,division_by_zero})"
23482
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
  1800
by (force simp add: divide_strict_right_mono order_le_less)
14293
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1801
16775
c1b87ef4a1c3 added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
avigad
parents: 16568
diff changeset
  1802
lemma divide_right_mono_neg: "(a::'a::{division_by_zero,ordered_field}) <= b 
c1b87ef4a1c3 added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
avigad
parents: 16568
diff changeset
  1803
    ==> c <= 0 ==> b / c <= a / c"
23482
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
  1804
apply (drule divide_right_mono [of _ _ "- c"])
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
  1805
apply auto
16775
c1b87ef4a1c3 added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
avigad
parents: 16568
diff changeset
  1806
done
c1b87ef4a1c3 added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
avigad
parents: 16568
diff changeset
  1807
c1b87ef4a1c3 added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
avigad
parents: 16568
diff changeset
  1808
lemma divide_strict_right_mono_neg:
c1b87ef4a1c3 added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
avigad
parents: 16568
diff changeset
  1809
     "[|b < a; c < 0|] ==> a / c < b / (c::'a::ordered_field)"
23482
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
  1810
apply (drule divide_strict_right_mono [of _ _ "-c"], simp)
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
  1811
apply (simp add: order_less_imp_not_eq nonzero_minus_divide_right [symmetric])
16775
c1b87ef4a1c3 added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
avigad
parents: 16568
diff changeset
  1812
done
14293
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1813
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1814
text{*The last premise ensures that @{term a} and @{term b} 
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1815
      have the same sign*}
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1816
lemma divide_strict_left_mono:
23482
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
  1817
  "[|b < a; 0 < c; 0 < a*b|] ==> c / a < c / (b::'a::ordered_field)"
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
  1818
by(auto simp: field_simps times_divide_eq zero_less_mult_iff mult_strict_right_mono)
14293
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1819
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1820
lemma divide_left_mono:
23482
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
  1821
  "[|b \<le> a; 0 \<le> c; 0 < a*b|] ==> c / a \<le> c / (b::'a::ordered_field)"
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
  1822
by(auto simp: field_simps times_divide_eq zero_less_mult_iff mult_right_mono)
14293
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1823
16775
c1b87ef4a1c3 added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
avigad
parents: 16568
diff changeset
  1824
lemma divide_left_mono_neg: "(a::'a::{division_by_zero,ordered_field}) <= b 
c1b87ef4a1c3 added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
avigad
parents: 16568
diff changeset
  1825
    ==> c <= 0 ==> 0 < a * b ==> c / a <= c / b"
c1b87ef4a1c3 added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
avigad
parents: 16568
diff changeset
  1826
  apply (drule divide_left_mono [of _ _ "- c"])
c1b87ef4a1c3 added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
avigad
parents: 16568
diff changeset
  1827
  apply (auto simp add: mult_commute)
c1b87ef4a1c3 added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
avigad
parents: 16568
diff changeset
  1828
done
c1b87ef4a1c3 added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
avigad
parents: 16568
diff changeset
  1829
14293
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1830
lemma divide_strict_left_mono_neg:
23482
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
  1831
  "[|a < b; c < 0; 0 < a*b|] ==> c / a < c / (b::'a::ordered_field)"
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
  1832
by(auto simp: field_simps times_divide_eq zero_less_mult_iff mult_strict_right_mono_neg)
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
  1833
14293
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1834
16775
c1b87ef4a1c3 added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
avigad
parents: 16568
diff changeset
  1835
text{*Simplify quotients that are compared with the value 1.*}
c1b87ef4a1c3 added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
avigad
parents: 16568
diff changeset
  1836
24286
7619080e49f0 ATP blacklisting is now in theory data, attribute noatp
paulson
parents: 23879
diff changeset
  1837
lemma le_divide_eq_1 [noatp]:
16775
c1b87ef4a1c3 added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
avigad
parents: 16568
diff changeset
  1838
  fixes a :: "'a :: {ordered_field,division_by_zero}"
c1b87ef4a1c3 added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
avigad
parents: 16568
diff changeset
  1839
  shows "(1 \<le> b / a) = ((0 < a & a \<le> b) | (a < 0 & b \<le> a))"
c1b87ef4a1c3 added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
avigad
parents: 16568
diff changeset
  1840
by (auto simp add: le_divide_eq)
c1b87ef4a1c3 added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
avigad
parents: 16568
diff changeset
  1841
24286
7619080e49f0 ATP blacklisting is now in theory data, attribute noatp
paulson
parents: 23879
diff changeset
  1842
lemma divide_le_eq_1 [noatp]:
16775
c1b87ef4a1c3 added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
avigad
parents: 16568
diff changeset
  1843
  fixes a :: "'a :: {ordered_field,division_by_zero}"
c1b87ef4a1c3 added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
avigad
parents: 16568
diff changeset
  1844
  shows "(b / a \<le> 1) = ((0 < a & b \<le> a) | (a < 0 & a \<le> b) | a=0)"
c1b87ef4a1c3 added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
avigad
parents: 16568
diff changeset
  1845
by (auto simp add: divide_le_eq)
c1b87ef4a1c3 added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
avigad
parents: 16568
diff changeset
  1846
24286
7619080e49f0 ATP blacklisting is now in theory data, attribute noatp
paulson
parents: 23879
diff changeset
  1847
lemma less_divide_eq_1 [noatp]:
16775
c1b87ef4a1c3 added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
avigad
parents: 16568
diff changeset
  1848
  fixes a :: "'a :: {ordered_field,division_by_zero}"
c1b87ef4a1c3 added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
avigad
parents: 16568
diff changeset
  1849
  shows "(1 < b / a) = ((0 < a & a < b) | (a < 0 & b < a))"
c1b87ef4a1c3 added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
avigad
parents: 16568
diff changeset
  1850
by (auto simp add: less_divide_eq)
c1b87ef4a1c3 added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
avigad
parents: 16568
diff changeset
  1851
24286
7619080e49f0 ATP blacklisting is now in theory data, attribute noatp
paulson
parents: 23879
diff changeset
  1852
lemma divide_less_eq_1 [noatp]:
16775
c1b87ef4a1c3 added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
avigad
parents: 16568
diff changeset
  1853
  fixes a :: "'a :: {ordered_field,division_by_zero}"
c1b87ef4a1c3 added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
avigad
parents: 16568
diff changeset
  1854
  shows "(b / a < 1) = ((0 < a & b < a) | (a < 0 & a < b) | a=0)"
c1b87ef4a1c3 added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
avigad
parents: 16568
diff changeset
  1855
by (auto simp add: divide_less_eq)
c1b87ef4a1c3 added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
avigad
parents: 16568
diff changeset
  1856
23389
aaca6a8e5414 tuned proofs: avoid implicit prems;
wenzelm
parents: 23326
diff changeset
  1857
16775
c1b87ef4a1c3 added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
avigad
parents: 16568
diff changeset
  1858
subsection{*Conditional Simplification Rules: No Case Splits*}
c1b87ef4a1c3 added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
avigad
parents: 16568
diff changeset
  1859
24286
7619080e49f0 ATP blacklisting is now in theory data, attribute noatp
paulson
parents: 23879
diff changeset
  1860
lemma le_divide_eq_1_pos [simp,noatp]:
16775
c1b87ef4a1c3 added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
avigad
parents: 16568
diff changeset
  1861
  fixes a :: "'a :: {ordered_field,division_by_zero}"
18649
bb99c2e705ca tidied, and added missing thm divide_less_eq_1_neg
paulson
parents: 18623
diff changeset
  1862
  shows "0 < a \<Longrightarrow> (1 \<le> b/a) = (a \<le> b)"
16775
c1b87ef4a1c3 added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
avigad
parents: 16568
diff changeset
  1863
by (auto simp add: le_divide_eq)
c1b87ef4a1c3 added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
avigad
parents: 16568
diff changeset
  1864
24286
7619080e49f0 ATP blacklisting is now in theory data, attribute noatp
paulson
parents: 23879
diff changeset
  1865
lemma le_divide_eq_1_neg [simp,noatp]:
16775
c1b87ef4a1c3 added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
avigad
parents: 16568
diff changeset
  1866
  fixes a :: "'a :: {ordered_field,division_by_zero}"
18649
bb99c2e705ca tidied, and added missing thm divide_less_eq_1_neg
paulson
parents: 18623
diff changeset
  1867
  shows "a < 0 \<Longrightarrow> (1 \<le> b/a) = (b \<le> a)"
16775
c1b87ef4a1c3 added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
avigad
parents: 16568
diff changeset
  1868
by (auto simp add: le_divide_eq)
c1b87ef4a1c3 added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
avigad
parents: 16568
diff changeset
  1869
24286
7619080e49f0 ATP blacklisting is now in theory data, attribute noatp
paulson
parents: 23879
diff changeset
  1870
lemma divide_le_eq_1_pos [simp,noatp]:
16775
c1b87ef4a1c3 added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
avigad
parents: 16568
diff changeset
  1871
  fixes a :: "'a :: {ordered_field,division_by_zero}"
18649
bb99c2e705ca tidied, and added missing thm divide_less_eq_1_neg
paulson
parents: 18623
diff changeset
  1872
  shows "0 < a \<Longrightarrow> (b/a \<le> 1) = (b \<le> a)"
16775
c1b87ef4a1c3 added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
avigad
parents: 16568
diff changeset
  1873
by (auto simp add: divide_le_eq)
c1b87ef4a1c3 added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
avigad
parents: 16568
diff changeset
  1874
24286
7619080e49f0 ATP blacklisting is now in theory data, attribute noatp
paulson
parents: 23879
diff changeset
  1875
lemma divide_le_eq_1_neg [simp,noatp]:
16775
c1b87ef4a1c3 added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
avigad
parents: 16568
diff changeset
  1876
  fixes a :: "'a :: {ordered_field,division_by_zero}"
18649
bb99c2e705ca tidied, and added missing thm divide_less_eq_1_neg
paulson
parents: 18623
diff changeset
  1877
  shows "a < 0 \<Longrightarrow> (b/a \<le> 1) = (a \<le> b)"
16775
c1b87ef4a1c3 added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
avigad
parents: 16568
diff changeset
  1878
by (auto simp add: divide_le_eq)
c1b87ef4a1c3 added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
avigad
parents: 16568
diff changeset
  1879
24286
7619080e49f0 ATP blacklisting is now in theory data, attribute noatp
paulson
parents: 23879
diff changeset
  1880
lemma less_divide_eq_1_pos [simp,noatp]:
16775
c1b87ef4a1c3 added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
avigad
parents: 16568
diff changeset
  1881
  fixes a :: "'a :: {ordered_field,division_by_zero}"
18649
bb99c2e705ca tidied, and added missing thm divide_less_eq_1_neg
paulson
parents: 18623
diff changeset
  1882
  shows "0 < a \<Longrightarrow> (1 < b/a) = (a < b)"
16775
c1b87ef4a1c3 added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
avigad
parents: 16568
diff changeset
  1883
by (auto simp add: less_divide_eq)
c1b87ef4a1c3 added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
avigad
parents: 16568
diff changeset
  1884
24286
7619080e49f0 ATP blacklisting is now in theory data, attribute noatp
paulson
parents: 23879
diff changeset
  1885
lemma less_divide_eq_1_neg [simp,noatp]:
16775
c1b87ef4a1c3 added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
avigad
parents: 16568
diff changeset
  1886
  fixes a :: "'a :: {ordered_field,division_by_zero}"
18649
bb99c2e705ca tidied, and added missing thm divide_less_eq_1_neg
paulson
parents: 18623
diff changeset
  1887
  shows "a < 0 \<Longrightarrow> (1 < b/a) = (b < a)"
16775
c1b87ef4a1c3 added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
avigad
parents: 16568
diff changeset
  1888
by (auto simp add: less_divide_eq)
c1b87ef4a1c3 added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
avigad
parents: 16568
diff changeset
  1889
24286
7619080e49f0 ATP blacklisting is now in theory data, attribute noatp
paulson
parents: 23879
diff changeset
  1890
lemma divide_less_eq_1_pos [simp,noatp]:
16775
c1b87ef4a1c3 added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
avigad
parents: 16568
diff changeset
  1891
  fixes a :: "'a :: {ordered_field,division_by_zero}"
18649
bb99c2e705ca tidied, and added missing thm divide_less_eq_1_neg
paulson
parents: 18623
diff changeset
  1892
  shows "0 < a \<Longrightarrow> (b/a < 1) = (b < a)"
bb99c2e705ca tidied, and added missing thm divide_less_eq_1_neg
paulson
parents: 18623
diff changeset
  1893
by (auto simp add: divide_less_eq)
bb99c2e705ca tidied, and added missing thm divide_less_eq_1_neg
paulson
parents: 18623
diff changeset
  1894
24286
7619080e49f0 ATP blacklisting is now in theory data, attribute noatp
paulson
parents: 23879
diff changeset
  1895
lemma divide_less_eq_1_neg [simp,noatp]:
18649
bb99c2e705ca tidied, and added missing thm divide_less_eq_1_neg
paulson
parents: 18623
diff changeset
  1896
  fixes a :: "'a :: {ordered_field,division_by_zero}"
bb99c2e705ca tidied, and added missing thm divide_less_eq_1_neg
paulson
parents: 18623
diff changeset
  1897
  shows "a < 0 \<Longrightarrow> b/a < 1 <-> a < b"
16775
c1b87ef4a1c3 added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
avigad
parents: 16568
diff changeset
  1898
by (auto simp add: divide_less_eq)
c1b87ef4a1c3 added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
avigad
parents: 16568
diff changeset
  1899
24286
7619080e49f0 ATP blacklisting is now in theory data, attribute noatp
paulson
parents: 23879
diff changeset
  1900
lemma eq_divide_eq_1 [simp,noatp]:
16775
c1b87ef4a1c3 added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
avigad
parents: 16568
diff changeset
  1901
  fixes a :: "'a :: {ordered_field,division_by_zero}"
18649
bb99c2e705ca tidied, and added missing thm divide_less_eq_1_neg
paulson
parents: 18623
diff changeset
  1902
  shows "(1 = b/a) = ((a \<noteq> 0 & a = b))"
16775
c1b87ef4a1c3 added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
avigad
parents: 16568
diff changeset
  1903
by (auto simp add: eq_divide_eq)
c1b87ef4a1c3 added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
avigad
parents: 16568
diff changeset
  1904
24286
7619080e49f0 ATP blacklisting is now in theory data, attribute noatp
paulson
parents: 23879
diff changeset
  1905
lemma divide_eq_eq_1 [simp,noatp]:
16775
c1b87ef4a1c3 added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
avigad
parents: 16568
diff changeset
  1906
  fixes a :: "'a :: {ordered_field,division_by_zero}"
18649
bb99c2e705ca tidied, and added missing thm divide_less_eq_1_neg
paulson
parents: 18623
diff changeset
  1907
  shows "(b/a = 1) = ((a \<noteq> 0 & a = b))"
16775
c1b87ef4a1c3 added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
avigad
parents: 16568
diff changeset
  1908
by (auto simp add: divide_eq_eq)
c1b87ef4a1c3 added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
avigad
parents: 16568
diff changeset
  1909
23389
aaca6a8e5414 tuned proofs: avoid implicit prems;
wenzelm
parents: 23326
diff changeset
  1910
16775
c1b87ef4a1c3 added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
avigad
parents: 16568
diff changeset
  1911
subsection {* Reasoning about inequalities with division *}
c1b87ef4a1c3 added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
avigad
parents: 16568
diff changeset
  1912
c1b87ef4a1c3 added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
avigad
parents: 16568
diff changeset
  1913
lemma mult_right_le_one_le: "0 <= (x::'a::ordered_idom) ==> 0 <= y ==> y <= 1
c1b87ef4a1c3 added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
avigad
parents: 16568
diff changeset
  1914
    ==> x * y <= x"
c1b87ef4a1c3 added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
avigad
parents: 16568
diff changeset
  1915
  by (auto simp add: mult_compare_simps);
c1b87ef4a1c3 added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
avigad
parents: 16568
diff changeset
  1916
c1b87ef4a1c3 added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
avigad
parents: 16568
diff changeset
  1917
lemma mult_left_le_one_le: "0 <= (x::'a::ordered_idom) ==> 0 <= y ==> y <= 1
c1b87ef4a1c3 added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
avigad
parents: 16568
diff changeset
  1918
    ==> y * x <= x"
c1b87ef4a1c3 added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
avigad
parents: 16568
diff changeset
  1919
  by (auto simp add: mult_compare_simps);
c1b87ef4a1c3 added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
avigad
parents: 16568
diff changeset
  1920
c1b87ef4a1c3 added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
avigad
parents: 16568
diff changeset
  1921
lemma mult_imp_div_pos_le: "0 < (y::'a::ordered_field) ==> x <= z * y ==>
c1b87ef4a1c3 added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
avigad
parents: 16568
diff changeset
  1922
    x / y <= z";
c1b87ef4a1c3 added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
avigad
parents: 16568
diff changeset
  1923
  by (subst pos_divide_le_eq, assumption+);
c1b87ef4a1c3 added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
avigad
parents: 16568
diff changeset
  1924
c1b87ef4a1c3 added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
avigad
parents: 16568
diff changeset
  1925
lemma mult_imp_le_div_pos: "0 < (y::'a::ordered_field) ==> z * y <= x ==>
23482
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
  1926
    z <= x / y"
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
  1927
by(simp add:field_simps)
16775
c1b87ef4a1c3 added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
avigad
parents: 16568
diff changeset
  1928
c1b87ef4a1c3 added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
avigad
parents: 16568
diff changeset
  1929
lemma mult_imp_div_pos_less: "0 < (y::'a::ordered_field) ==> x < z * y ==>
c1b87ef4a1c3 added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
avigad
parents: 16568
diff changeset
  1930
    x / y < z"
23482
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
  1931
by(simp add:field_simps)
16775
c1b87ef4a1c3 added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
avigad
parents: 16568
diff changeset
  1932
c1b87ef4a1c3 added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
avigad
parents: 16568
diff changeset
  1933
lemma mult_imp_less_div_pos: "0 < (y::'a::ordered_field) ==> z * y < x ==>
c1b87ef4a1c3 added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
avigad
parents: 16568
diff changeset
  1934
    z < x / y"
23482
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
  1935
by(simp add:field_simps)
16775
c1b87ef4a1c3 added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
avigad
parents: 16568
diff changeset
  1936
c1b87ef4a1c3 added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
avigad
parents: 16568
diff changeset
  1937
lemma frac_le: "(0::'a::ordered_field) <= x ==> 
c1b87ef4a1c3 added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
avigad
parents: 16568
diff changeset
  1938
    x <= y ==> 0 < w ==> w <= z  ==> x / z <= y / w"
c1b87ef4a1c3 added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
avigad
parents: 16568
diff changeset
  1939
  apply (rule mult_imp_div_pos_le)
c1b87ef4a1c3 added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
avigad
parents: 16568
diff changeset
  1940
  apply simp;
c1b87ef4a1c3 added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
avigad
parents: 16568
diff changeset
  1941
  apply (subst times_divide_eq_left);
c1b87ef4a1c3 added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
avigad
parents: 16568
diff changeset
  1942
  apply (rule mult_imp_le_div_pos, assumption)
c1b87ef4a1c3 added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
avigad
parents: 16568
diff changeset
  1943
  apply (rule mult_mono)
c1b87ef4a1c3 added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
avigad
parents: 16568
diff changeset
  1944
  apply simp_all
14293
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1945
done
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1946
16775
c1b87ef4a1c3 added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
avigad
parents: 16568
diff changeset
  1947
lemma frac_less: "(0::'a::ordered_field) <= x ==> 
c1b87ef4a1c3 added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
avigad
parents: 16568
diff changeset
  1948
    x < y ==> 0 < w ==> w <= z  ==> x / z < y / w"
c1b87ef4a1c3 added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
avigad
parents: 16568
diff changeset
  1949
  apply (rule mult_imp_div_pos_less)
c1b87ef4a1c3 added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
avigad
parents: 16568
diff changeset
  1950
  apply simp;
c1b87ef4a1c3 added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
avigad
parents: 16568
diff changeset
  1951
  apply (subst times_divide_eq_left);
c1b87ef4a1c3 added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
avigad
parents: 16568
diff changeset
  1952
  apply (rule mult_imp_less_div_pos, assumption)
c1b87ef4a1c3 added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
avigad
parents: 16568
diff changeset
  1953
  apply (erule mult_less_le_imp_less)
c1b87ef4a1c3 added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
avigad
parents: 16568
diff changeset
  1954
  apply simp_all
c1b87ef4a1c3 added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
avigad
parents: 16568
diff changeset
  1955
done
c1b87ef4a1c3 added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
avigad
parents: 16568
diff changeset
  1956
c1b87ef4a1c3 added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
avigad
parents: 16568
diff changeset
  1957
lemma frac_less2: "(0::'a::ordered_field) < x ==> 
c1b87ef4a1c3 added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
avigad
parents: 16568
diff changeset
  1958
    x <= y ==> 0 < w ==> w < z  ==> x / z < y / w"
c1b87ef4a1c3 added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
avigad
parents: 16568
diff changeset
  1959
  apply (rule mult_imp_div_pos_less)
c1b87ef4a1c3 added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
avigad
parents: 16568
diff changeset
  1960
  apply simp_all
c1b87ef4a1c3 added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
avigad
parents: 16568
diff changeset
  1961
  apply (subst times_divide_eq_left);
c1b87ef4a1c3 added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
avigad
parents: 16568
diff changeset
  1962
  apply (rule mult_imp_less_div_pos, assumption)
c1b87ef4a1c3 added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
avigad
parents: 16568
diff changeset
  1963
  apply (erule mult_le_less_imp_less)
c1b87ef4a1c3 added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
avigad
parents: 16568
diff changeset
  1964
  apply simp_all
c1b87ef4a1c3 added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
avigad
parents: 16568
diff changeset
  1965
done
c1b87ef4a1c3 added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
avigad
parents: 16568
diff changeset
  1966
c1b87ef4a1c3 added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
avigad
parents: 16568
diff changeset
  1967
text{*It's not obvious whether these should be simprules or not. 
c1b87ef4a1c3 added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
avigad
parents: 16568
diff changeset
  1968
  Their effect is to gather terms into one big fraction, like
c1b87ef4a1c3 added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
avigad
parents: 16568
diff changeset
  1969
  a*b*c / x*y*z. The rationale for that is unclear, but many proofs 
c1b87ef4a1c3 added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
avigad
parents: 16568
diff changeset
  1970
  seem to need them.*}
c1b87ef4a1c3 added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
avigad
parents: 16568
diff changeset
  1971
c1b87ef4a1c3 added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
avigad
parents: 16568
diff changeset
  1972
declare times_divide_eq [simp]
14293
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1973
23389
aaca6a8e5414 tuned proofs: avoid implicit prems;
wenzelm
parents: 23326
diff changeset
  1974
14293
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1975
subsection {* Ordered Fields are Dense *}
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1976
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1977
lemma less_add_one: "a < (a+1::'a::ordered_semidom)"
14293
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1978
proof -
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1979
  have "a+0 < (a+1::'a::ordered_semidom)"
23482
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
  1980
    by (blast intro: zero_less_one add_strict_left_mono)
14293
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1981
  thus ?thesis by simp
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1982
qed
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1983
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1984
lemma zero_less_two: "0 < (1+1::'a::ordered_semidom)"
23482
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
  1985
by (blast intro: order_less_trans zero_less_one less_add_one)
14365
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14353
diff changeset
  1986
14293
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1987
lemma less_half_sum: "a < b ==> a < (a+b) / (1+1::'a::ordered_field)"
23482
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
  1988
by (simp add: field_simps zero_less_two)
14293
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1989
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1990
lemma gt_half_sum: "a < b ==> (a+b)/(1+1::'a::ordered_field) < b"
23482
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
  1991
by (simp add: field_simps zero_less_two)
14293
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1992
24422
c0b5ff9e9e4d moved class dense_linear_order to Orderings.thy
haftmann
parents: 24286
diff changeset
  1993
instance ordered_field < dense_linear_order
c0b5ff9e9e4d moved class dense_linear_order to Orderings.thy
haftmann
parents: 24286
diff changeset
  1994
proof
c0b5ff9e9e4d moved class dense_linear_order to Orderings.thy
haftmann
parents: 24286
diff changeset
  1995
  fix x y :: 'a
c0b5ff9e9e4d moved class dense_linear_order to Orderings.thy
haftmann
parents: 24286
diff changeset
  1996
  have "x < x + 1" by simp
c0b5ff9e9e4d moved class dense_linear_order to Orderings.thy
haftmann
parents: 24286
diff changeset
  1997
  then show "\<exists>y. x < y" .. 
c0b5ff9e9e4d moved class dense_linear_order to Orderings.thy
haftmann
parents: 24286
diff changeset
  1998
  have "x - 1 < x" by simp
c0b5ff9e9e4d moved class dense_linear_order to Orderings.thy
haftmann
parents: 24286
diff changeset
  1999
  then show "\<exists>y. y < x" ..
c0b5ff9e9e4d moved class dense_linear_order to Orderings.thy
haftmann
parents: 24286
diff changeset
  2000
  show "x < y \<Longrightarrow> \<exists>z>x. z < y" by (blast intro!: less_half_sum gt_half_sum)
c0b5ff9e9e4d moved class dense_linear_order to Orderings.thy
haftmann
parents: 24286
diff changeset
  2001
qed
14293
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  2002
15234
ec91a90c604e simplification tweaks for better arithmetic reasoning
paulson
parents: 15229
diff changeset
  2003
14293
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  2004
subsection {* Absolute Value *}
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  2005
24491
8d194c9198ae added constant sgn
nipkow
parents: 24427
diff changeset
  2006
lemma mult_sgn_abs: "sgn x * abs x = (x::'a::{ordered_idom,linorder})"
8d194c9198ae added constant sgn
nipkow
parents: 24427
diff changeset
  2007
using less_linear[of x 0]
24506
020db6ec334a final(?) iteration of sgn saga.
nipkow
parents: 24491
diff changeset
  2008
by(auto simp: sgn_if abs_if)
24491
8d194c9198ae added constant sgn
nipkow
parents: 24427
diff changeset
  2009
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  2010
lemma abs_one [simp]: "abs 1 = (1::'a::ordered_idom)"
23477
f4b83f03cac9 tuned and renamed group_eq_simps and ring_eq_simps
nipkow
parents: 23413
diff changeset
  2011
by (simp add: abs_if zero_less_one [THEN order_less_not_sym])
14294
f4d806fd72ce absolute value theorems moved to HOL/Ring_and_Field
paulson
parents: 14293
diff changeset
  2012
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  2013
lemma abs_le_mult: "abs (a * b) \<le> (abs a) * (abs (b::'a::lordered_ring))" 
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  2014
proof -
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  2015
  let ?x = "pprt a * pprt b - pprt a * nprt b - nprt a * pprt b + nprt a * nprt b"
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  2016
  let ?y = "pprt a * pprt b + pprt a * nprt b + nprt a * pprt b + nprt a * nprt b"
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  2017
  have a: "(abs a) * (abs b) = ?x"
23477
f4b83f03cac9 tuned and renamed group_eq_simps and ring_eq_simps
nipkow
parents: 23413
diff changeset
  2018
    by (simp only: abs_prts[of a] abs_prts[of b] ring_simps)
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  2019
  {
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  2020
    fix u v :: 'a
15481
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents: 15234
diff changeset
  2021
    have bh: "\<lbrakk>u = a; v = b\<rbrakk> \<Longrightarrow> 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents: 15234
diff changeset
  2022
              u * v = pprt a * pprt b + pprt a * nprt b + 
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents: 15234
diff changeset
  2023
                      nprt a * pprt b + nprt a * nprt b"
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  2024
      apply (subst prts[of u], subst prts[of v])
23477
f4b83f03cac9 tuned and renamed group_eq_simps and ring_eq_simps
nipkow
parents: 23413
diff changeset
  2025
      apply (simp add: ring_simps) 
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  2026
      done
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  2027
  }
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  2028
  note b = this[OF refl[of a] refl[of b]]
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  2029
  note addm = add_mono[of "0::'a" _ "0::'a", simplified]
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  2030
  note addm2 = add_mono[of _ "0::'a" _ "0::'a", simplified]
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  2031
  have xy: "- ?x <= ?y"
14754
a080eeeaec14 Modification / Installation of Provers/Arith/abel_cancel.ML for OrderedGroup.thy
obua
parents: 14738
diff changeset
  2032
    apply (simp)
a080eeeaec14 Modification / Installation of Provers/Arith/abel_cancel.ML for OrderedGroup.thy
obua
parents: 14738
diff changeset
  2033
    apply (rule_tac y="0::'a" in order_trans)
16568
e02fe7ae212b Changes due to new abel_cancel.ML
nipkow
parents: 15923
diff changeset
  2034
    apply (rule addm2)
16775
c1b87ef4a1c3 added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
avigad
parents: 16568
diff changeset
  2035
    apply (simp_all add: mult_nonneg_nonneg mult_nonpos_nonpos)
16568
e02fe7ae212b Changes due to new abel_cancel.ML
nipkow
parents: 15923
diff changeset
  2036
    apply (rule addm)
16775
c1b87ef4a1c3 added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
avigad
parents: 16568
diff changeset
  2037
    apply (simp_all add: mult_nonneg_nonneg mult_nonpos_nonpos)
14754
a080eeeaec14 Modification / Installation of Provers/Arith/abel_cancel.ML for OrderedGroup.thy
obua
parents: 14738
diff changeset
  2038
    done
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  2039
  have yx: "?y <= ?x"
16568
e02fe7ae212b Changes due to new abel_cancel.ML
nipkow
parents: 15923
diff changeset
  2040
    apply (simp add:diff_def)
14754
a080eeeaec14 Modification / Installation of Provers/Arith/abel_cancel.ML for OrderedGroup.thy
obua
parents: 14738
diff changeset
  2041
    apply (rule_tac y=0 in order_trans)
16775
c1b87ef4a1c3 added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
avigad
parents: 16568
diff changeset
  2042
    apply (rule addm2, (simp add: mult_nonneg_nonpos mult_nonneg_nonpos2)+)
c1b87ef4a1c3 added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
avigad
parents: 16568
diff changeset
  2043
    apply (rule addm, (simp add: mult_nonneg_nonpos mult_nonneg_nonpos2)+)
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  2044
    done
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  2045
  have i1: "a*b <= abs a * abs b" by (simp only: a b yx)
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  2046
  have i2: "- (abs a * abs b) <= a*b" by (simp only: a b xy)
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  2047
  show ?thesis
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  2048
    apply (rule abs_leI)
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  2049
    apply (simp add: i1)
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  2050
    apply (simp add: i2[simplified minus_le_iff])
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  2051
    done
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  2052
qed
14294
f4d806fd72ce absolute value theorems moved to HOL/Ring_and_Field
paulson
parents: 14293
diff changeset
  2053
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  2054
lemma abs_eq_mult: 
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  2055
  assumes "(0 \<le> a \<or> a \<le> 0) \<and> (0 \<le> b \<or> b \<le> 0)"
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  2056
  shows "abs (a*b) = abs a * abs (b::'a::lordered_ring)"
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  2057
proof -
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  2058
  have s: "(0 <= a*b) | (a*b <= 0)"
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  2059
    apply (auto)    
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  2060
    apply (rule_tac split_mult_pos_le)
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  2061
    apply (rule_tac contrapos_np[of "a*b <= 0"])
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  2062
    apply (simp)
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  2063
    apply (rule_tac split_mult_neg_le)
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  2064
    apply (insert prems)
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  2065
    apply (blast)
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  2066
    done
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  2067
  have mulprts: "a * b = (pprt a + nprt a) * (pprt b + nprt b)"
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  2068
    by (simp add: prts[symmetric])
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  2069
  show ?thesis
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  2070
  proof cases
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  2071
    assume "0 <= a * b"
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  2072
    then show ?thesis
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  2073
      apply (simp_all add: mulprts abs_prts)
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  2074
      apply (insert prems)
14754
a080eeeaec14 Modification / Installation of Provers/Arith/abel_cancel.ML for OrderedGroup.thy
obua
parents: 14738
diff changeset
  2075
      apply (auto simp add: 
23477
f4b83f03cac9 tuned and renamed group_eq_simps and ring_eq_simps
nipkow
parents: 23413
diff changeset
  2076
	ring_simps 
25078
a1ddc5206cb1 moved lemmas to OrderedGroup.thy
haftmann
parents: 25062
diff changeset
  2077
	iffD1[OF zero_le_iff_zero_nprt] iffD1[OF le_zero_iff_zero_pprt]
a1ddc5206cb1 moved lemmas to OrderedGroup.thy
haftmann
parents: 25062
diff changeset
  2078
	iffD1[OF le_zero_iff_pprt_id] iffD1[OF zero_le_iff_nprt_id])
16775
c1b87ef4a1c3 added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
avigad
parents: 16568
diff changeset
  2079
	apply(drule (1) mult_nonneg_nonpos[of a b], simp)
c1b87ef4a1c3 added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
avigad
parents: 16568
diff changeset
  2080
	apply(drule (1) mult_nonneg_nonpos2[of b a], simp)
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  2081
      done
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  2082
  next
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  2083
    assume "~(0 <= a*b)"
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  2084
    with s have "a*b <= 0" by simp
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  2085
    then show ?thesis
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  2086
      apply (simp_all add: mulprts abs_prts)
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  2087
      apply (insert prems)
23477
f4b83f03cac9 tuned and renamed group_eq_simps and ring_eq_simps
nipkow
parents: 23413
diff changeset
  2088
      apply (auto simp add: ring_simps)
16775
c1b87ef4a1c3 added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
avigad
parents: 16568
diff changeset
  2089
      apply(drule (1) mult_nonneg_nonneg[of a b],simp)
c1b87ef4a1c3 added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
avigad
parents: 16568
diff changeset
  2090
      apply(drule (1) mult_nonpos_nonpos[of a b],simp)
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  2091
      done
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  2092
  qed
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  2093
qed
14294
f4d806fd72ce absolute value theorems moved to HOL/Ring_and_Field
paulson
parents: 14293
diff changeset
  2094
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  2095
lemma abs_mult: "abs (a * b) = abs a * abs (b::'a::ordered_idom)" 
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  2096
by (simp add: abs_eq_mult linorder_linear)
14293
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  2097
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  2098
lemma abs_mult_self: "abs a * abs a = a * (a::'a::ordered_idom)"
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  2099
by (simp add: abs_if) 
14294
f4d806fd72ce absolute value theorems moved to HOL/Ring_and_Field
paulson
parents: 14293
diff changeset
  2100
f4d806fd72ce absolute value theorems moved to HOL/Ring_and_Field
paulson
parents: 14293
diff changeset
  2101
lemma nonzero_abs_inverse:
f4d806fd72ce absolute value theorems moved to HOL/Ring_and_Field
paulson
parents: 14293
diff changeset
  2102
     "a \<noteq> 0 ==> abs (inverse (a::'a::ordered_field)) = inverse (abs a)"
f4d806fd72ce absolute value theorems moved to HOL/Ring_and_Field
paulson
parents: 14293
diff changeset
  2103
apply (auto simp add: linorder_neq_iff abs_if nonzero_inverse_minus_eq 
f4d806fd72ce absolute value theorems moved to HOL/Ring_and_Field
paulson
parents: 14293
diff changeset
  2104
                      negative_imp_inverse_negative)
f4d806fd72ce absolute value theorems moved to HOL/Ring_and_Field
paulson
parents: 14293
diff changeset
  2105
apply (blast intro: positive_imp_inverse_positive elim: order_less_asym) 
f4d806fd72ce absolute value theorems moved to HOL/Ring_and_Field
paulson
parents: 14293
diff changeset
  2106
done
f4d806fd72ce absolute value theorems moved to HOL/Ring_and_Field
paulson
parents: 14293
diff changeset
  2107
f4d806fd72ce absolute value theorems moved to HOL/Ring_and_Field
paulson
parents: 14293
diff changeset
  2108
lemma abs_inverse [simp]:
f4d806fd72ce absolute value theorems moved to HOL/Ring_and_Field
paulson
parents: 14293
diff changeset
  2109
     "abs (inverse (a::'a::{ordered_field,division_by_zero})) = 
f4d806fd72ce absolute value theorems moved to HOL/Ring_and_Field
paulson
parents: 14293
diff changeset
  2110
      inverse (abs a)"
21328
73bb86d0f483 dropped Inductive dependency
haftmann
parents: 21258
diff changeset
  2111
apply (cases "a=0", simp) 
14294
f4d806fd72ce absolute value theorems moved to HOL/Ring_and_Field
paulson
parents: 14293
diff changeset
  2112
apply (simp add: nonzero_abs_inverse) 
f4d806fd72ce absolute value theorems moved to HOL/Ring_and_Field
paulson
parents: 14293
diff changeset
  2113
done
f4d806fd72ce absolute value theorems moved to HOL/Ring_and_Field
paulson
parents: 14293
diff changeset
  2114
f4d806fd72ce absolute value theorems moved to HOL/Ring_and_Field
paulson
parents: 14293
diff changeset
  2115
lemma nonzero_abs_divide:
f4d806fd72ce absolute value theorems moved to HOL/Ring_and_Field
paulson
parents: 14293
diff changeset
  2116
     "b \<noteq> 0 ==> abs (a / (b::'a::ordered_field)) = abs a / abs b"
f4d806fd72ce absolute value theorems moved to HOL/Ring_and_Field
paulson
parents: 14293
diff changeset
  2117
by (simp add: divide_inverse abs_mult nonzero_abs_inverse) 
f4d806fd72ce absolute value theorems moved to HOL/Ring_and_Field
paulson
parents: 14293
diff changeset
  2118
15234
ec91a90c604e simplification tweaks for better arithmetic reasoning
paulson
parents: 15229
diff changeset
  2119
lemma abs_divide [simp]:
14294
f4d806fd72ce absolute value theorems moved to HOL/Ring_and_Field
paulson
parents: 14293
diff changeset
  2120
     "abs (a / (b::'a::{ordered_field,division_by_zero})) = abs a / abs b"
21328
73bb86d0f483 dropped Inductive dependency
haftmann
parents: 21258
diff changeset
  2121
apply (cases "b=0", simp) 
14294
f4d806fd72ce absolute value theorems moved to HOL/Ring_and_Field
paulson
parents: 14293
diff changeset
  2122
apply (simp add: nonzero_abs_divide) 
f4d806fd72ce absolute value theorems moved to HOL/Ring_and_Field
paulson
parents: 14293
diff changeset
  2123
done
f4d806fd72ce absolute value theorems moved to HOL/Ring_and_Field
paulson
parents: 14293
diff changeset
  2124
f4d806fd72ce absolute value theorems moved to HOL/Ring_and_Field
paulson
parents: 14293
diff changeset
  2125
lemma abs_mult_less:
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  2126
     "[| abs a < c; abs b < d |] ==> abs a * abs b < c*(d::'a::ordered_idom)"
14294
f4d806fd72ce absolute value theorems moved to HOL/Ring_and_Field
paulson
parents: 14293
diff changeset
  2127
proof -
f4d806fd72ce absolute value theorems moved to HOL/Ring_and_Field
paulson
parents: 14293
diff changeset
  2128
  assume ac: "abs a < c"
f4d806fd72ce absolute value theorems moved to HOL/Ring_and_Field
paulson
parents: 14293
diff changeset
  2129
  hence cpos: "0<c" by (blast intro: order_le_less_trans abs_ge_zero)
f4d806fd72ce absolute value theorems moved to HOL/Ring_and_Field
paulson
parents: 14293
diff changeset
  2130
  assume "abs b < d"
f4d806fd72ce absolute value theorems moved to HOL/Ring_and_Field
paulson
parents: 14293
diff changeset
  2131
  thus ?thesis by (simp add: ac cpos mult_strict_mono) 
f4d806fd72ce absolute value theorems moved to HOL/Ring_and_Field
paulson
parents: 14293
diff changeset
  2132
qed
14293
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  2133
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  2134
lemma eq_minus_self_iff: "(a = -a) = (a = (0::'a::ordered_idom))"
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  2135
by (force simp add: order_eq_iff le_minus_self_iff minus_le_self_iff)
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  2136
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  2137
lemma less_minus_self_iff: "(a < -a) = (a < (0::'a::ordered_idom))"
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  2138
by (simp add: order_less_le le_minus_self_iff eq_minus_self_iff)
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  2139
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  2140
lemma abs_less_iff: "(abs a < b) = (a < b & -a < (b::'a::ordered_idom))" 
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  2141
apply (simp add: order_less_le abs_le_iff)  
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  2142
apply (auto simp add: abs_if minus_le_self_iff eq_minus_self_iff)
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  2143
apply (simp add: le_minus_self_iff linorder_neq_iff) 
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  2144
done
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  2145
16775
c1b87ef4a1c3 added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
avigad
parents: 16568
diff changeset
  2146
lemma abs_mult_pos: "(0::'a::ordered_idom) <= x ==> 
c1b87ef4a1c3 added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
avigad
parents: 16568
diff changeset
  2147
    (abs y) * x = abs (y * x)";
c1b87ef4a1c3 added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
avigad
parents: 16568
diff changeset
  2148
  apply (subst abs_mult);
c1b87ef4a1c3 added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
avigad
parents: 16568
diff changeset
  2149
  apply simp;
c1b87ef4a1c3 added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
avigad
parents: 16568
diff changeset
  2150
done;
c1b87ef4a1c3 added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
avigad
parents: 16568
diff changeset
  2151
c1b87ef4a1c3 added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
avigad
parents: 16568
diff changeset
  2152
lemma abs_div_pos: "(0::'a::{division_by_zero,ordered_field}) < y ==> 
c1b87ef4a1c3 added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
avigad
parents: 16568
diff changeset
  2153
    abs x / y = abs (x / y)";
c1b87ef4a1c3 added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
avigad
parents: 16568
diff changeset
  2154
  apply (subst abs_divide);
c1b87ef4a1c3 added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
avigad
parents: 16568
diff changeset
  2155
  apply (simp add: order_less_imp_le);
c1b87ef4a1c3 added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
avigad
parents: 16568
diff changeset
  2156
done;
c1b87ef4a1c3 added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
avigad
parents: 16568
diff changeset
  2157
23389
aaca6a8e5414 tuned proofs: avoid implicit prems;
wenzelm
parents: 23326
diff changeset
  2158
19404
9bf2cdc9e8e8 Moved stuff from Ring_and_Field to Matrix
obua
parents: 18649
diff changeset
  2159
subsection {* Bounds of products via negative and positive Part *}
15178
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15140
diff changeset
  2160
15580
900291ee0af8 Cleaning up HOL/Matrix
obua
parents: 15481
diff changeset
  2161
lemma mult_le_prts:
900291ee0af8 Cleaning up HOL/Matrix
obua
parents: 15481
diff changeset
  2162
  assumes
900291ee0af8 Cleaning up HOL/Matrix
obua
parents: 15481
diff changeset
  2163
  "a1 <= (a::'a::lordered_ring)"
900291ee0af8 Cleaning up HOL/Matrix
obua
parents: 15481
diff changeset
  2164
  "a <= a2"
900291ee0af8 Cleaning up HOL/Matrix
obua
parents: 15481
diff changeset
  2165
  "b1 <= b"
900291ee0af8 Cleaning up HOL/Matrix
obua
parents: 15481
diff changeset
  2166
  "b <= b2"
900291ee0af8 Cleaning up HOL/Matrix
obua
parents: 15481
diff changeset
  2167
  shows
900291ee0af8 Cleaning up HOL/Matrix
obua
parents: 15481
diff changeset
  2168
  "a * b <= pprt a2 * pprt b2 + pprt a1 * nprt b2 + nprt a2 * pprt b1 + nprt a1 * nprt b1"
900291ee0af8 Cleaning up HOL/Matrix
obua
parents: 15481
diff changeset
  2169
proof - 
900291ee0af8 Cleaning up HOL/Matrix
obua
parents: 15481
diff changeset
  2170
  have "a * b = (pprt a + nprt a) * (pprt b + nprt b)" 
900291ee0af8 Cleaning up HOL/Matrix
obua
parents: 15481
diff changeset
  2171
    apply (subst prts[symmetric])+
900291ee0af8 Cleaning up HOL/Matrix
obua
parents: 15481
diff changeset
  2172
    apply simp
900291ee0af8 Cleaning up HOL/Matrix
obua
parents: 15481
diff changeset
  2173
    done
900291ee0af8 Cleaning up HOL/Matrix
obua
parents: 15481
diff changeset
  2174
  then have "a * b = pprt a * pprt b + pprt a * nprt b + nprt a * pprt b + nprt a * nprt b"
23477
f4b83f03cac9 tuned and renamed group_eq_simps and ring_eq_simps
nipkow
parents: 23413
diff changeset
  2175
    by (simp add: ring_simps)
15580
900291ee0af8 Cleaning up HOL/Matrix
obua
parents: 15481
diff changeset
  2176
  moreover have "pprt a * pprt b <= pprt a2 * pprt b2"
900291ee0af8 Cleaning up HOL/Matrix
obua
parents: 15481
diff changeset
  2177
    by (simp_all add: prems mult_mono)
900291ee0af8 Cleaning up HOL/Matrix
obua
parents: 15481
diff changeset
  2178
  moreover have "pprt a * nprt b <= pprt a1 * nprt b2"
900291ee0af8 Cleaning up HOL/Matrix
obua
parents: 15481
diff changeset
  2179
  proof -
900291ee0af8 Cleaning up HOL/Matrix
obua
parents: 15481
diff changeset
  2180
    have "pprt a * nprt b <= pprt a * nprt b2"
900291ee0af8 Cleaning up HOL/Matrix
obua
parents: 15481
diff changeset
  2181
      by (simp add: mult_left_mono prems)
900291ee0af8 Cleaning up HOL/Matrix
obua
parents: 15481
diff changeset
  2182
    moreover have "pprt a * nprt b2 <= pprt a1 * nprt b2"
900291ee0af8 Cleaning up HOL/Matrix
obua
parents: 15481
diff changeset
  2183
      by (simp add: mult_right_mono_neg prems)
900291ee0af8 Cleaning up HOL/Matrix
obua
parents: 15481
diff changeset
  2184
    ultimately show ?thesis
900291ee0af8 Cleaning up HOL/Matrix
obua
parents: 15481
diff changeset
  2185
      by simp
900291ee0af8 Cleaning up HOL/Matrix
obua
parents: 15481
diff changeset
  2186
  qed
900291ee0af8 Cleaning up HOL/Matrix
obua
parents: 15481
diff changeset
  2187
  moreover have "nprt a * pprt b <= nprt a2 * pprt b1"
900291ee0af8 Cleaning up HOL/Matrix
obua
parents: 15481
diff changeset
  2188
  proof - 
900291ee0af8 Cleaning up HOL/Matrix
obua
parents: 15481
diff changeset
  2189
    have "nprt a * pprt b <= nprt a2 * pprt b"
900291ee0af8 Cleaning up HOL/Matrix
obua
parents: 15481
diff changeset
  2190
      by (simp add: mult_right_mono prems)
900291ee0af8 Cleaning up HOL/Matrix
obua
parents: 15481
diff changeset
  2191
    moreover have "nprt a2 * pprt b <= nprt a2 * pprt b1"
900291ee0af8 Cleaning up HOL/Matrix
obua
parents: 15481
diff changeset
  2192
      by (simp add: mult_left_mono_neg prems)
900291ee0af8 Cleaning up HOL/Matrix
obua
parents: 15481
diff changeset
  2193
    ultimately show ?thesis
900291ee0af8 Cleaning up HOL/Matrix
obua
parents: 15481
diff changeset
  2194
      by simp
900291ee0af8 Cleaning up HOL/Matrix
obua
parents: 15481
diff changeset
  2195
  qed
900291ee0af8 Cleaning up HOL/Matrix
obua
parents: 15481
diff changeset
  2196
  moreover have "nprt a * nprt b <= nprt a1 * nprt b1"
900291ee0af8 Cleaning up HOL/Matrix
obua
parents: 15481
diff changeset
  2197
  proof -
900291ee0af8 Cleaning up HOL/Matrix
obua
parents: 15481
diff changeset
  2198
    have "nprt a * nprt b <= nprt a * nprt b1"
900291ee0af8 Cleaning up HOL/Matrix
obua
parents: 15481
diff changeset
  2199
      by (simp add: mult_left_mono_neg prems)
900291ee0af8 Cleaning up HOL/Matrix
obua
parents: 15481
diff changeset
  2200
    moreover have "nprt a * nprt b1 <= nprt a1 * nprt b1"
900291ee0af8 Cleaning up HOL/Matrix
obua
parents: 15481
diff changeset
  2201
      by (simp add: mult_right_mono_neg prems)
900291ee0af8 Cleaning up HOL/Matrix
obua
parents: 15481
diff changeset
  2202
    ultimately show ?thesis
900291ee0af8 Cleaning up HOL/Matrix
obua
parents: 15481
diff changeset
  2203
      by simp
900291ee0af8 Cleaning up HOL/Matrix
obua
parents: 15481
diff changeset
  2204
  qed
900291ee0af8 Cleaning up HOL/Matrix
obua
parents: 15481
diff changeset
  2205
  ultimately show ?thesis
900291ee0af8 Cleaning up HOL/Matrix
obua
parents: 15481
diff changeset
  2206
    by - (rule add_mono | simp)+
900291ee0af8 Cleaning up HOL/Matrix
obua
parents: 15481
diff changeset
  2207
qed
19404
9bf2cdc9e8e8 Moved stuff from Ring_and_Field to Matrix
obua
parents: 18649
diff changeset
  2208
9bf2cdc9e8e8 Moved stuff from Ring_and_Field to Matrix
obua
parents: 18649
diff changeset
  2209
lemma mult_ge_prts:
15178
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15140
diff changeset
  2210
  assumes
19404
9bf2cdc9e8e8 Moved stuff from Ring_and_Field to Matrix
obua
parents: 18649
diff changeset
  2211
  "a1 <= (a::'a::lordered_ring)"
9bf2cdc9e8e8 Moved stuff from Ring_and_Field to Matrix
obua
parents: 18649
diff changeset
  2212
  "a <= a2"
9bf2cdc9e8e8 Moved stuff from Ring_and_Field to Matrix
obua
parents: 18649
diff changeset
  2213
  "b1 <= b"
9bf2cdc9e8e8 Moved stuff from Ring_and_Field to Matrix
obua
parents: 18649
diff changeset
  2214
  "b <= b2"
15178
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15140
diff changeset
  2215
  shows
19404
9bf2cdc9e8e8 Moved stuff from Ring_and_Field to Matrix
obua
parents: 18649
diff changeset
  2216
  "a * b >= nprt a1 * pprt b2 + nprt a2 * nprt b2 + pprt a1 * pprt b1 + pprt a2 * nprt b1"
9bf2cdc9e8e8 Moved stuff from Ring_and_Field to Matrix
obua
parents: 18649
diff changeset
  2217
proof - 
9bf2cdc9e8e8 Moved stuff from Ring_and_Field to Matrix
obua
parents: 18649
diff changeset
  2218
  from prems have a1:"- a2 <= -a" by auto
9bf2cdc9e8e8 Moved stuff from Ring_and_Field to Matrix
obua
parents: 18649
diff changeset
  2219
  from prems have a2: "-a <= -a1" by auto
9bf2cdc9e8e8 Moved stuff from Ring_and_Field to Matrix
obua
parents: 18649
diff changeset
  2220
  from mult_le_prts[of "-a2" "-a" "-a1" "b1" b "b2", OF a1 a2 prems(3) prems(4), simplified nprt_neg pprt_neg] 
9bf2cdc9e8e8 Moved stuff from Ring_and_Field to Matrix
obua
parents: 18649
diff changeset
  2221
  have le: "- (a * b) <= - nprt a1 * pprt b2 + - nprt a2 * nprt b2 + - pprt a1 * pprt b1 + - pprt a2 * nprt b1" by simp  
9bf2cdc9e8e8 Moved stuff from Ring_and_Field to Matrix
obua
parents: 18649
diff changeset
  2222
  then have "-(- nprt a1 * pprt b2 + - nprt a2 * nprt b2 + - pprt a1 * pprt b1 + - pprt a2 * nprt b1) <= a * b"
9bf2cdc9e8e8 Moved stuff from Ring_and_Field to Matrix
obua
parents: 18649
diff changeset
  2223
    by (simp only: minus_le_iff)
9bf2cdc9e8e8 Moved stuff from Ring_and_Field to Matrix
obua
parents: 18649
diff changeset
  2224
  then show ?thesis by simp
15178
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15140
diff changeset
  2225
qed
5f621aa35c25 Matrix theory, linear programming
obua
parents: 15140
diff changeset
  2226
14265
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
  2227
end