src/HOL/Fields.thy
author haftmann
Mon, 08 Feb 2010 17:12:38 +0100
changeset 35050 9f841f20dca6
parent 35043 src/HOL/Ring_and_Field.thy@07dbdf60d5ad
child 35084 e25eedfc15ce
permissions -rw-r--r--
renamed OrderedGroup to Groups; split theory Ring_and_Field into Rings Fields
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
35050
9f841f20dca6 renamed OrderedGroup to Groups; split theory Ring_and_Field into Rings Fields
haftmann
parents: 35043
diff changeset
     1
(*  Title:      HOL/Fields.thy
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 30961
diff changeset
     2
    Author:     Gertrud Bauer
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 30961
diff changeset
     3
    Author:     Steven Obua
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 30961
diff changeset
     4
    Author:     Tobias Nipkow
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 30961
diff changeset
     5
    Author:     Lawrence C Paulson
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 30961
diff changeset
     6
    Author:     Markus Wenzel
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 30961
diff changeset
     7
    Author:     Jeremy Avigad
14265
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
     8
*)
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
     9
35050
9f841f20dca6 renamed OrderedGroup to Groups; split theory Ring_and_Field into Rings Fields
haftmann
parents: 35043
diff changeset
    10
header {* Fields *}
25152
bfde2f8c0f63 partially localized
haftmann
parents: 25078
diff changeset
    11
35050
9f841f20dca6 renamed OrderedGroup to Groups; split theory Ring_and_Field into Rings Fields
haftmann
parents: 35043
diff changeset
    12
theory Fields
9f841f20dca6 renamed OrderedGroup to Groups; split theory Ring_and_Field into Rings Fields
haftmann
parents: 35043
diff changeset
    13
imports Rings
25186
f4d1ebffd025 localized further
haftmann
parents: 25152
diff changeset
    14
begin
14421
ee97b6463cb4 new Ring_and_Field hierarchy, eliminating redundant axioms
paulson
parents: 14398
diff changeset
    15
22987
550709aa8e66 instance division_ring < no_zero_divisors; clean up field instance proofs
huffman
parents: 22842
diff changeset
    16
class field = comm_ring_1 + inverse +
25062
af5ef0d4d655 global class syntax
haftmann
parents: 24748
diff changeset
    17
  assumes field_inverse:  "a \<noteq> 0 \<Longrightarrow> inverse a * a = 1"
af5ef0d4d655 global class syntax
haftmann
parents: 24748
diff changeset
    18
  assumes divide_inverse: "a / b = a * inverse b"
25267
1f745c599b5c proper reinitialisation after subclass
haftmann
parents: 25238
diff changeset
    19
begin
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
    20
25267
1f745c599b5c proper reinitialisation after subclass
haftmann
parents: 25238
diff changeset
    21
subclass division_ring
28823
dcbef866c9e2 tuned unfold_locales invocation
haftmann
parents: 28559
diff changeset
    22
proof
22987
550709aa8e66 instance division_ring < no_zero_divisors; clean up field instance proofs
huffman
parents: 22842
diff changeset
    23
  fix a :: 'a
550709aa8e66 instance division_ring < no_zero_divisors; clean up field instance proofs
huffman
parents: 22842
diff changeset
    24
  assume "a \<noteq> 0"
550709aa8e66 instance division_ring < no_zero_divisors; clean up field instance proofs
huffman
parents: 22842
diff changeset
    25
  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
    26
  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
    27
qed
25230
022029099a83 continued localization
haftmann
parents: 25193
diff changeset
    28
27516
9a5d4a8d4aac by intro_locales -> ..
huffman
parents: 26274
diff changeset
    29
subclass idom ..
25230
022029099a83 continued localization
haftmann
parents: 25193
diff changeset
    30
022029099a83 continued localization
haftmann
parents: 25193
diff changeset
    31
lemma right_inverse_eq: "b \<noteq> 0 \<Longrightarrow> a / b = 1 \<longleftrightarrow> a = b"
022029099a83 continued localization
haftmann
parents: 25193
diff changeset
    32
proof
022029099a83 continued localization
haftmann
parents: 25193
diff changeset
    33
  assume neq: "b \<noteq> 0"
022029099a83 continued localization
haftmann
parents: 25193
diff changeset
    34
  {
022029099a83 continued localization
haftmann
parents: 25193
diff changeset
    35
    hence "a = (a / b) * b" by (simp add: divide_inverse mult_ac)
022029099a83 continued localization
haftmann
parents: 25193
diff changeset
    36
    also assume "a / b = 1"
022029099a83 continued localization
haftmann
parents: 25193
diff changeset
    37
    finally show "a = b" by simp
022029099a83 continued localization
haftmann
parents: 25193
diff changeset
    38
  next
022029099a83 continued localization
haftmann
parents: 25193
diff changeset
    39
    assume "a = b"
022029099a83 continued localization
haftmann
parents: 25193
diff changeset
    40
    with neq show "a / b = 1" by (simp add: divide_inverse)
022029099a83 continued localization
haftmann
parents: 25193
diff changeset
    41
  }
022029099a83 continued localization
haftmann
parents: 25193
diff changeset
    42
qed
022029099a83 continued localization
haftmann
parents: 25193
diff changeset
    43
022029099a83 continued localization
haftmann
parents: 25193
diff changeset
    44
lemma nonzero_inverse_eq_divide: "a \<noteq> 0 \<Longrightarrow> inverse a = 1 / a"
29667
53103fc8ffa3 Replaced group_ and ring_simps by algebra_simps;
nipkow
parents: 29465
diff changeset
    45
by (simp add: divide_inverse)
25230
022029099a83 continued localization
haftmann
parents: 25193
diff changeset
    46
022029099a83 continued localization
haftmann
parents: 25193
diff changeset
    47
lemma divide_self [simp]: "a \<noteq> 0 \<Longrightarrow> a / a = 1"
29667
53103fc8ffa3 Replaced group_ and ring_simps by algebra_simps;
nipkow
parents: 29465
diff changeset
    48
by (simp add: divide_inverse)
25230
022029099a83 continued localization
haftmann
parents: 25193
diff changeset
    49
022029099a83 continued localization
haftmann
parents: 25193
diff changeset
    50
lemma divide_zero_left [simp]: "0 / a = 0"
29667
53103fc8ffa3 Replaced group_ and ring_simps by algebra_simps;
nipkow
parents: 29465
diff changeset
    51
by (simp add: divide_inverse)
25230
022029099a83 continued localization
haftmann
parents: 25193
diff changeset
    52
022029099a83 continued localization
haftmann
parents: 25193
diff changeset
    53
lemma inverse_eq_divide: "inverse a = 1 / a"
29667
53103fc8ffa3 Replaced group_ and ring_simps by algebra_simps;
nipkow
parents: 29465
diff changeset
    54
by (simp add: divide_inverse)
25230
022029099a83 continued localization
haftmann
parents: 25193
diff changeset
    55
022029099a83 continued localization
haftmann
parents: 25193
diff changeset
    56
lemma add_divide_distrib: "(a+b) / c = a/c + b/c"
30630
4fbe1401bac2 move field lemmas into class locale context
huffman
parents: 30242
diff changeset
    57
by (simp add: divide_inverse algebra_simps)
4fbe1401bac2 move field lemmas into class locale context
huffman
parents: 30242
diff changeset
    58
4fbe1401bac2 move field lemmas into class locale context
huffman
parents: 30242
diff changeset
    59
text{*There is no slick version using division by zero.*}
4fbe1401bac2 move field lemmas into class locale context
huffman
parents: 30242
diff changeset
    60
lemma inverse_add:
4fbe1401bac2 move field lemmas into class locale context
huffman
parents: 30242
diff changeset
    61
  "[| a \<noteq> 0;  b \<noteq> 0 |]
4fbe1401bac2 move field lemmas into class locale context
huffman
parents: 30242
diff changeset
    62
   ==> inverse a + inverse b = (a + b) * inverse a * inverse b"
4fbe1401bac2 move field lemmas into class locale context
huffman
parents: 30242
diff changeset
    63
by (simp add: division_ring_inverse_add mult_ac)
4fbe1401bac2 move field lemmas into class locale context
huffman
parents: 30242
diff changeset
    64
4fbe1401bac2 move field lemmas into class locale context
huffman
parents: 30242
diff changeset
    65
lemma nonzero_mult_divide_mult_cancel_left [simp, noatp]:
4fbe1401bac2 move field lemmas into class locale context
huffman
parents: 30242
diff changeset
    66
assumes [simp]: "b\<noteq>0" and [simp]: "c\<noteq>0" shows "(c*a)/(c*b) = a/b"
4fbe1401bac2 move field lemmas into class locale context
huffman
parents: 30242
diff changeset
    67
proof -
4fbe1401bac2 move field lemmas into class locale context
huffman
parents: 30242
diff changeset
    68
  have "(c*a)/(c*b) = c * a * (inverse b * inverse c)"
4fbe1401bac2 move field lemmas into class locale context
huffman
parents: 30242
diff changeset
    69
    by (simp add: divide_inverse nonzero_inverse_mult_distrib)
4fbe1401bac2 move field lemmas into class locale context
huffman
parents: 30242
diff changeset
    70
  also have "... =  a * inverse b * (inverse c * c)"
4fbe1401bac2 move field lemmas into class locale context
huffman
parents: 30242
diff changeset
    71
    by (simp only: mult_ac)
4fbe1401bac2 move field lemmas into class locale context
huffman
parents: 30242
diff changeset
    72
  also have "... =  a * inverse b" by simp
4fbe1401bac2 move field lemmas into class locale context
huffman
parents: 30242
diff changeset
    73
    finally show ?thesis by (simp add: divide_inverse)
4fbe1401bac2 move field lemmas into class locale context
huffman
parents: 30242
diff changeset
    74
qed
4fbe1401bac2 move field lemmas into class locale context
huffman
parents: 30242
diff changeset
    75
4fbe1401bac2 move field lemmas into class locale context
huffman
parents: 30242
diff changeset
    76
lemma nonzero_mult_divide_mult_cancel_right [simp, noatp]:
4fbe1401bac2 move field lemmas into class locale context
huffman
parents: 30242
diff changeset
    77
  "\<lbrakk>b \<noteq> 0; c \<noteq> 0\<rbrakk> \<Longrightarrow> (a * c) / (b * c) = a / b"
4fbe1401bac2 move field lemmas into class locale context
huffman
parents: 30242
diff changeset
    78
by (simp add: mult_commute [of _ c])
4fbe1401bac2 move field lemmas into class locale context
huffman
parents: 30242
diff changeset
    79
4fbe1401bac2 move field lemmas into class locale context
huffman
parents: 30242
diff changeset
    80
lemma divide_1 [simp]: "a / 1 = a"
4fbe1401bac2 move field lemmas into class locale context
huffman
parents: 30242
diff changeset
    81
by (simp add: divide_inverse)
4fbe1401bac2 move field lemmas into class locale context
huffman
parents: 30242
diff changeset
    82
4fbe1401bac2 move field lemmas into class locale context
huffman
parents: 30242
diff changeset
    83
lemma times_divide_eq_right: "a * (b / c) = (a * b) / c"
4fbe1401bac2 move field lemmas into class locale context
huffman
parents: 30242
diff changeset
    84
by (simp add: divide_inverse mult_assoc)
4fbe1401bac2 move field lemmas into class locale context
huffman
parents: 30242
diff changeset
    85
4fbe1401bac2 move field lemmas into class locale context
huffman
parents: 30242
diff changeset
    86
lemma times_divide_eq_left: "(b / c) * a = (b * a) / c"
4fbe1401bac2 move field lemmas into class locale context
huffman
parents: 30242
diff changeset
    87
by (simp add: divide_inverse mult_ac)
4fbe1401bac2 move field lemmas into class locale context
huffman
parents: 30242
diff changeset
    88
4fbe1401bac2 move field lemmas into class locale context
huffman
parents: 30242
diff changeset
    89
text {* These are later declared as simp rules. *}
4fbe1401bac2 move field lemmas into class locale context
huffman
parents: 30242
diff changeset
    90
lemmas times_divide_eq [noatp] = times_divide_eq_right times_divide_eq_left
4fbe1401bac2 move field lemmas into class locale context
huffman
parents: 30242
diff changeset
    91
4fbe1401bac2 move field lemmas into class locale context
huffman
parents: 30242
diff changeset
    92
lemma add_frac_eq:
4fbe1401bac2 move field lemmas into class locale context
huffman
parents: 30242
diff changeset
    93
  assumes "y \<noteq> 0" and "z \<noteq> 0"
4fbe1401bac2 move field lemmas into class locale context
huffman
parents: 30242
diff changeset
    94
  shows "x / y + w / z = (x * z + w * y) / (y * z)"
4fbe1401bac2 move field lemmas into class locale context
huffman
parents: 30242
diff changeset
    95
proof -
4fbe1401bac2 move field lemmas into class locale context
huffman
parents: 30242
diff changeset
    96
  have "x / y + w / z = (x * z) / (y * z) + (y * w) / (y * z)"
4fbe1401bac2 move field lemmas into class locale context
huffman
parents: 30242
diff changeset
    97
    using assms by simp
4fbe1401bac2 move field lemmas into class locale context
huffman
parents: 30242
diff changeset
    98
  also have "\<dots> = (x * z + y * w) / (y * z)"
4fbe1401bac2 move field lemmas into class locale context
huffman
parents: 30242
diff changeset
    99
    by (simp only: add_divide_distrib)
4fbe1401bac2 move field lemmas into class locale context
huffman
parents: 30242
diff changeset
   100
  finally show ?thesis
4fbe1401bac2 move field lemmas into class locale context
huffman
parents: 30242
diff changeset
   101
    by (simp only: mult_commute)
4fbe1401bac2 move field lemmas into class locale context
huffman
parents: 30242
diff changeset
   102
qed
4fbe1401bac2 move field lemmas into class locale context
huffman
parents: 30242
diff changeset
   103
4fbe1401bac2 move field lemmas into class locale context
huffman
parents: 30242
diff changeset
   104
text{*Special Cancellation Simprules for Division*}
4fbe1401bac2 move field lemmas into class locale context
huffman
parents: 30242
diff changeset
   105
4fbe1401bac2 move field lemmas into class locale context
huffman
parents: 30242
diff changeset
   106
lemma nonzero_mult_divide_cancel_right [simp, noatp]:
4fbe1401bac2 move field lemmas into class locale context
huffman
parents: 30242
diff changeset
   107
  "b \<noteq> 0 \<Longrightarrow> a * b / b = a"
4fbe1401bac2 move field lemmas into class locale context
huffman
parents: 30242
diff changeset
   108
using nonzero_mult_divide_mult_cancel_right [of 1 b a] by simp
4fbe1401bac2 move field lemmas into class locale context
huffman
parents: 30242
diff changeset
   109
4fbe1401bac2 move field lemmas into class locale context
huffman
parents: 30242
diff changeset
   110
lemma nonzero_mult_divide_cancel_left [simp, noatp]:
4fbe1401bac2 move field lemmas into class locale context
huffman
parents: 30242
diff changeset
   111
  "a \<noteq> 0 \<Longrightarrow> a * b / a = b"
4fbe1401bac2 move field lemmas into class locale context
huffman
parents: 30242
diff changeset
   112
using nonzero_mult_divide_mult_cancel_left [of 1 a b] by simp
4fbe1401bac2 move field lemmas into class locale context
huffman
parents: 30242
diff changeset
   113
4fbe1401bac2 move field lemmas into class locale context
huffman
parents: 30242
diff changeset
   114
lemma nonzero_divide_mult_cancel_right [simp, noatp]:
4fbe1401bac2 move field lemmas into class locale context
huffman
parents: 30242
diff changeset
   115
  "\<lbrakk>a \<noteq> 0; b \<noteq> 0\<rbrakk> \<Longrightarrow> b / (a * b) = 1 / a"
4fbe1401bac2 move field lemmas into class locale context
huffman
parents: 30242
diff changeset
   116
using nonzero_mult_divide_mult_cancel_right [of a b 1] by simp
4fbe1401bac2 move field lemmas into class locale context
huffman
parents: 30242
diff changeset
   117
4fbe1401bac2 move field lemmas into class locale context
huffman
parents: 30242
diff changeset
   118
lemma nonzero_divide_mult_cancel_left [simp, noatp]:
4fbe1401bac2 move field lemmas into class locale context
huffman
parents: 30242
diff changeset
   119
  "\<lbrakk>a \<noteq> 0; b \<noteq> 0\<rbrakk> \<Longrightarrow> a / (a * b) = 1 / b"
4fbe1401bac2 move field lemmas into class locale context
huffman
parents: 30242
diff changeset
   120
using nonzero_mult_divide_mult_cancel_left [of b a 1] by simp
4fbe1401bac2 move field lemmas into class locale context
huffman
parents: 30242
diff changeset
   121
4fbe1401bac2 move field lemmas into class locale context
huffman
parents: 30242
diff changeset
   122
lemma nonzero_mult_divide_mult_cancel_left2 [simp, noatp]:
4fbe1401bac2 move field lemmas into class locale context
huffman
parents: 30242
diff changeset
   123
  "\<lbrakk>b \<noteq> 0; c \<noteq> 0\<rbrakk> \<Longrightarrow> (c * a) / (b * c) = a / b"
4fbe1401bac2 move field lemmas into class locale context
huffman
parents: 30242
diff changeset
   124
using nonzero_mult_divide_mult_cancel_left [of b c a] by (simp add: mult_ac)
4fbe1401bac2 move field lemmas into class locale context
huffman
parents: 30242
diff changeset
   125
4fbe1401bac2 move field lemmas into class locale context
huffman
parents: 30242
diff changeset
   126
lemma nonzero_mult_divide_mult_cancel_right2 [simp, noatp]:
4fbe1401bac2 move field lemmas into class locale context
huffman
parents: 30242
diff changeset
   127
  "\<lbrakk>b \<noteq> 0; c \<noteq> 0\<rbrakk> \<Longrightarrow> (a * c) / (c * b) = a / b"
4fbe1401bac2 move field lemmas into class locale context
huffman
parents: 30242
diff changeset
   128
using nonzero_mult_divide_mult_cancel_right [of b c a] by (simp add: mult_ac)
4fbe1401bac2 move field lemmas into class locale context
huffman
parents: 30242
diff changeset
   129
4fbe1401bac2 move field lemmas into class locale context
huffman
parents: 30242
diff changeset
   130
lemma minus_divide_left: "- (a / b) = (-a) / b"
4fbe1401bac2 move field lemmas into class locale context
huffman
parents: 30242
diff changeset
   131
by (simp add: divide_inverse)
4fbe1401bac2 move field lemmas into class locale context
huffman
parents: 30242
diff changeset
   132
4fbe1401bac2 move field lemmas into class locale context
huffman
parents: 30242
diff changeset
   133
lemma nonzero_minus_divide_right: "b \<noteq> 0 ==> - (a / b) = a / (- b)"
4fbe1401bac2 move field lemmas into class locale context
huffman
parents: 30242
diff changeset
   134
by (simp add: divide_inverse nonzero_inverse_minus_eq)
4fbe1401bac2 move field lemmas into class locale context
huffman
parents: 30242
diff changeset
   135
4fbe1401bac2 move field lemmas into class locale context
huffman
parents: 30242
diff changeset
   136
lemma nonzero_minus_divide_divide: "b \<noteq> 0 ==> (-a) / (-b) = a / b"
4fbe1401bac2 move field lemmas into class locale context
huffman
parents: 30242
diff changeset
   137
by (simp add: divide_inverse nonzero_inverse_minus_eq)
4fbe1401bac2 move field lemmas into class locale context
huffman
parents: 30242
diff changeset
   138
4fbe1401bac2 move field lemmas into class locale context
huffman
parents: 30242
diff changeset
   139
lemma divide_minus_left [simp, noatp]: "(-a) / b = - (a / b)"
4fbe1401bac2 move field lemmas into class locale context
huffman
parents: 30242
diff changeset
   140
by (simp add: divide_inverse)
4fbe1401bac2 move field lemmas into class locale context
huffman
parents: 30242
diff changeset
   141
4fbe1401bac2 move field lemmas into class locale context
huffman
parents: 30242
diff changeset
   142
lemma diff_divide_distrib: "(a - b) / c = a / c - b / c"
4fbe1401bac2 move field lemmas into class locale context
huffman
parents: 30242
diff changeset
   143
by (simp add: diff_minus add_divide_distrib)
4fbe1401bac2 move field lemmas into class locale context
huffman
parents: 30242
diff changeset
   144
4fbe1401bac2 move field lemmas into class locale context
huffman
parents: 30242
diff changeset
   145
lemma add_divide_eq_iff:
4fbe1401bac2 move field lemmas into class locale context
huffman
parents: 30242
diff changeset
   146
  "z \<noteq> 0 \<Longrightarrow> x + y / z = (z * x + y) / z"
4fbe1401bac2 move field lemmas into class locale context
huffman
parents: 30242
diff changeset
   147
by (simp add: add_divide_distrib)
4fbe1401bac2 move field lemmas into class locale context
huffman
parents: 30242
diff changeset
   148
4fbe1401bac2 move field lemmas into class locale context
huffman
parents: 30242
diff changeset
   149
lemma divide_add_eq_iff:
4fbe1401bac2 move field lemmas into class locale context
huffman
parents: 30242
diff changeset
   150
  "z \<noteq> 0 \<Longrightarrow> x / z + y = (x + z * y) / z"
4fbe1401bac2 move field lemmas into class locale context
huffman
parents: 30242
diff changeset
   151
by (simp add: add_divide_distrib)
4fbe1401bac2 move field lemmas into class locale context
huffman
parents: 30242
diff changeset
   152
4fbe1401bac2 move field lemmas into class locale context
huffman
parents: 30242
diff changeset
   153
lemma diff_divide_eq_iff:
4fbe1401bac2 move field lemmas into class locale context
huffman
parents: 30242
diff changeset
   154
  "z \<noteq> 0 \<Longrightarrow> x - y / z = (z * x - y) / z"
4fbe1401bac2 move field lemmas into class locale context
huffman
parents: 30242
diff changeset
   155
by (simp add: diff_divide_distrib)
4fbe1401bac2 move field lemmas into class locale context
huffman
parents: 30242
diff changeset
   156
4fbe1401bac2 move field lemmas into class locale context
huffman
parents: 30242
diff changeset
   157
lemma divide_diff_eq_iff:
4fbe1401bac2 move field lemmas into class locale context
huffman
parents: 30242
diff changeset
   158
  "z \<noteq> 0 \<Longrightarrow> x / z - y = (x - z * y) / z"
4fbe1401bac2 move field lemmas into class locale context
huffman
parents: 30242
diff changeset
   159
by (simp add: diff_divide_distrib)
4fbe1401bac2 move field lemmas into class locale context
huffman
parents: 30242
diff changeset
   160
4fbe1401bac2 move field lemmas into class locale context
huffman
parents: 30242
diff changeset
   161
lemma nonzero_eq_divide_eq: "c \<noteq> 0 \<Longrightarrow> a = b / c \<longleftrightarrow> a * c = b"
4fbe1401bac2 move field lemmas into class locale context
huffman
parents: 30242
diff changeset
   162
proof -
4fbe1401bac2 move field lemmas into class locale context
huffman
parents: 30242
diff changeset
   163
  assume [simp]: "c \<noteq> 0"
4fbe1401bac2 move field lemmas into class locale context
huffman
parents: 30242
diff changeset
   164
  have "a = b / c \<longleftrightarrow> a * c = (b / c) * c" by simp
4fbe1401bac2 move field lemmas into class locale context
huffman
parents: 30242
diff changeset
   165
  also have "... \<longleftrightarrow> a * c = b" by (simp add: divide_inverse mult_assoc)
4fbe1401bac2 move field lemmas into class locale context
huffman
parents: 30242
diff changeset
   166
  finally show ?thesis .
4fbe1401bac2 move field lemmas into class locale context
huffman
parents: 30242
diff changeset
   167
qed
4fbe1401bac2 move field lemmas into class locale context
huffman
parents: 30242
diff changeset
   168
4fbe1401bac2 move field lemmas into class locale context
huffman
parents: 30242
diff changeset
   169
lemma nonzero_divide_eq_eq: "c \<noteq> 0 \<Longrightarrow> b / c = a \<longleftrightarrow> b = a * c"
4fbe1401bac2 move field lemmas into class locale context
huffman
parents: 30242
diff changeset
   170
proof -
4fbe1401bac2 move field lemmas into class locale context
huffman
parents: 30242
diff changeset
   171
  assume [simp]: "c \<noteq> 0"
4fbe1401bac2 move field lemmas into class locale context
huffman
parents: 30242
diff changeset
   172
  have "b / c = a \<longleftrightarrow> (b / c) * c = a * c" by simp
4fbe1401bac2 move field lemmas into class locale context
huffman
parents: 30242
diff changeset
   173
  also have "... \<longleftrightarrow> b = a * c" by (simp add: divide_inverse mult_assoc) 
4fbe1401bac2 move field lemmas into class locale context
huffman
parents: 30242
diff changeset
   174
  finally show ?thesis .
4fbe1401bac2 move field lemmas into class locale context
huffman
parents: 30242
diff changeset
   175
qed
4fbe1401bac2 move field lemmas into class locale context
huffman
parents: 30242
diff changeset
   176
4fbe1401bac2 move field lemmas into class locale context
huffman
parents: 30242
diff changeset
   177
lemma divide_eq_imp: "c \<noteq> 0 \<Longrightarrow> b = a * c \<Longrightarrow> b / c = a"
4fbe1401bac2 move field lemmas into class locale context
huffman
parents: 30242
diff changeset
   178
by simp
4fbe1401bac2 move field lemmas into class locale context
huffman
parents: 30242
diff changeset
   179
4fbe1401bac2 move field lemmas into class locale context
huffman
parents: 30242
diff changeset
   180
lemma eq_divide_imp: "c \<noteq> 0 \<Longrightarrow> a * c = b \<Longrightarrow> a = b / c"
4fbe1401bac2 move field lemmas into class locale context
huffman
parents: 30242
diff changeset
   181
by (erule subst, simp)
4fbe1401bac2 move field lemmas into class locale context
huffman
parents: 30242
diff changeset
   182
4fbe1401bac2 move field lemmas into class locale context
huffman
parents: 30242
diff changeset
   183
lemmas field_eq_simps[noatp] = algebra_simps
4fbe1401bac2 move field lemmas into class locale context
huffman
parents: 30242
diff changeset
   184
  (* pull / out*)
4fbe1401bac2 move field lemmas into class locale context
huffman
parents: 30242
diff changeset
   185
  add_divide_eq_iff divide_add_eq_iff
4fbe1401bac2 move field lemmas into class locale context
huffman
parents: 30242
diff changeset
   186
  diff_divide_eq_iff divide_diff_eq_iff
4fbe1401bac2 move field lemmas into class locale context
huffman
parents: 30242
diff changeset
   187
  (* multiply eqn *)
4fbe1401bac2 move field lemmas into class locale context
huffman
parents: 30242
diff changeset
   188
  nonzero_eq_divide_eq nonzero_divide_eq_eq
4fbe1401bac2 move field lemmas into class locale context
huffman
parents: 30242
diff changeset
   189
(* is added later:
4fbe1401bac2 move field lemmas into class locale context
huffman
parents: 30242
diff changeset
   190
  times_divide_eq_left times_divide_eq_right
4fbe1401bac2 move field lemmas into class locale context
huffman
parents: 30242
diff changeset
   191
*)
4fbe1401bac2 move field lemmas into class locale context
huffman
parents: 30242
diff changeset
   192
4fbe1401bac2 move field lemmas into class locale context
huffman
parents: 30242
diff changeset
   193
text{*An example:*}
4fbe1401bac2 move field lemmas into class locale context
huffman
parents: 30242
diff changeset
   194
lemma "\<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"
4fbe1401bac2 move field lemmas into class locale context
huffman
parents: 30242
diff changeset
   195
apply(subgoal_tac "(c-d)*(e-f)*(a-b) \<noteq> 0")
4fbe1401bac2 move field lemmas into class locale context
huffman
parents: 30242
diff changeset
   196
 apply(simp add:field_eq_simps)
4fbe1401bac2 move field lemmas into class locale context
huffman
parents: 30242
diff changeset
   197
apply(simp)
4fbe1401bac2 move field lemmas into class locale context
huffman
parents: 30242
diff changeset
   198
done
4fbe1401bac2 move field lemmas into class locale context
huffman
parents: 30242
diff changeset
   199
4fbe1401bac2 move field lemmas into class locale context
huffman
parents: 30242
diff changeset
   200
lemma diff_frac_eq:
4fbe1401bac2 move field lemmas into class locale context
huffman
parents: 30242
diff changeset
   201
  "y \<noteq> 0 \<Longrightarrow> z \<noteq> 0 \<Longrightarrow> x / y - w / z = (x * z - w * y) / (y * z)"
4fbe1401bac2 move field lemmas into class locale context
huffman
parents: 30242
diff changeset
   202
by (simp add: field_eq_simps times_divide_eq)
4fbe1401bac2 move field lemmas into class locale context
huffman
parents: 30242
diff changeset
   203
4fbe1401bac2 move field lemmas into class locale context
huffman
parents: 30242
diff changeset
   204
lemma frac_eq_eq:
4fbe1401bac2 move field lemmas into class locale context
huffman
parents: 30242
diff changeset
   205
  "y \<noteq> 0 \<Longrightarrow> z \<noteq> 0 \<Longrightarrow> (x / y = w / z) = (x * z = w * y)"
4fbe1401bac2 move field lemmas into class locale context
huffman
parents: 30242
diff changeset
   206
by (simp add: field_eq_simps times_divide_eq)
25230
022029099a83 continued localization
haftmann
parents: 25193
diff changeset
   207
022029099a83 continued localization
haftmann
parents: 25193
diff changeset
   208
end
022029099a83 continued localization
haftmann
parents: 25193
diff changeset
   209
22390
378f34b1e380 now using "class"
haftmann
parents: 21328
diff changeset
   210
class division_by_zero = zero + inverse +
25062
af5ef0d4d655 global class syntax
haftmann
parents: 24748
diff changeset
   211
  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
   212
25230
022029099a83 continued localization
haftmann
parents: 25193
diff changeset
   213
lemma divide_zero [simp]:
022029099a83 continued localization
haftmann
parents: 25193
diff changeset
   214
  "a / 0 = (0::'a::{field,division_by_zero})"
29667
53103fc8ffa3 Replaced group_ and ring_simps by algebra_simps;
nipkow
parents: 29465
diff changeset
   215
by (simp add: divide_inverse)
25230
022029099a83 continued localization
haftmann
parents: 25193
diff changeset
   216
022029099a83 continued localization
haftmann
parents: 25193
diff changeset
   217
lemma divide_self_if [simp]:
022029099a83 continued localization
haftmann
parents: 25193
diff changeset
   218
  "a / (a::'a::{field,division_by_zero}) = (if a=0 then 0 else 1)"
29667
53103fc8ffa3 Replaced group_ and ring_simps by algebra_simps;
nipkow
parents: 29465
diff changeset
   219
by simp
25230
022029099a83 continued localization
haftmann
parents: 25193
diff changeset
   220
35028
108662d50512 more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents: 34146
diff changeset
   221
class linordered_field = field + linordered_idom
25230
022029099a83 continued localization
haftmann
parents: 25193
diff changeset
   222
14268
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   223
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
   224
   "(inverse a = 0) = (a = (0::'a::{division_ring,division_by_zero}))"
26274
2bdb61a28971 continued localization
haftmann
parents: 26234
diff changeset
   225
by (force dest: inverse_zero_imp_zero) 
14268
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   226
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   227
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
   228
   "inverse(-a) = -inverse(a::'a::{division_ring,division_by_zero})"
14377
f454b3004f8f tidying up, especially the Complex numbers
paulson
parents: 14370
diff changeset
   229
proof cases
f454b3004f8f tidying up, especially the Complex numbers
paulson
parents: 14370
diff changeset
   230
  assume "a=0" thus ?thesis by (simp add: inverse_zero)
f454b3004f8f tidying up, especially the Complex numbers
paulson
parents: 14370
diff changeset
   231
next
f454b3004f8f tidying up, especially the Complex numbers
paulson
parents: 14370
diff changeset
   232
  assume "a\<noteq>0" 
f454b3004f8f tidying up, especially the Complex numbers
paulson
parents: 14370
diff changeset
   233
  thus ?thesis by (simp add: nonzero_inverse_minus_eq)
f454b3004f8f tidying up, especially the Complex numbers
paulson
parents: 14370
diff changeset
   234
qed
14268
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   235
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   236
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
   237
  "inverse a = inverse b ==> a = (b::'a::{division_ring,division_by_zero})"
21328
73bb86d0f483 dropped Inductive dependency
haftmann
parents: 21258
diff changeset
   238
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
   239
 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
   240
              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
   241
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
   242
done
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   243
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   244
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
   245
  "(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
   246
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
   247
14270
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   248
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
   249
     "inverse(inverse (a::'a::{division_ring,division_by_zero})) = a"
14270
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   250
  proof cases
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   251
    assume "a=0" thus ?thesis by simp
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   252
  next
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   253
    assume "a\<noteq>0" 
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   254
    thus ?thesis by (simp add: nonzero_inverse_inverse_eq)
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   255
  qed
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   256
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   257
text{*This version builds in division by zero while also re-orienting
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   258
      the right-hand side.*}
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   259
lemma inverse_mult_distrib [simp]:
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   260
     "inverse(a*b) = inverse(a) * inverse(b::'a::{field,division_by_zero})"
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   261
  proof cases
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   262
    assume "a \<noteq> 0 & b \<noteq> 0" 
29667
53103fc8ffa3 Replaced group_ and ring_simps by algebra_simps;
nipkow
parents: 29465
diff changeset
   263
    thus ?thesis by (simp add: nonzero_inverse_mult_distrib mult_commute)
14270
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   264
  next
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   265
    assume "~ (a \<noteq> 0 & b \<noteq> 0)" 
29667
53103fc8ffa3 Replaced group_ and ring_simps by algebra_simps;
nipkow
parents: 29465
diff changeset
   266
    thus ?thesis by force
14270
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   267
  qed
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
   268
14365
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14353
diff changeset
   269
lemma inverse_divide [simp]:
23477
f4b83f03cac9 tuned and renamed group_eq_simps and ring_eq_simps
nipkow
parents: 23413
diff changeset
   270
  "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
   271
by (simp add: divide_inverse mult_commute)
14365
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14353
diff changeset
   272
23389
aaca6a8e5414 tuned proofs: avoid implicit prems;
wenzelm
parents: 23326
diff changeset
   273
16775
c1b87ef4a1c3 added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
avigad
parents: 16568
diff changeset
   274
subsection {* Calculations with fractions *}
c1b87ef4a1c3 added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
avigad
parents: 16568
diff changeset
   275
23413
5caa2710dd5b tuned laws for cancellation in divisions for fields.
nipkow
parents: 23406
diff changeset
   276
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
   277
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
   278
because the latter are covered by a simproc. *}
5caa2710dd5b tuned laws for cancellation in divisions for fields.
nipkow
parents: 23406
diff changeset
   279
5caa2710dd5b tuned laws for cancellation in divisions for fields.
nipkow
parents: 23406
diff changeset
   280
lemma mult_divide_mult_cancel_left:
23477
f4b83f03cac9 tuned and renamed group_eq_simps and ring_eq_simps
nipkow
parents: 23413
diff changeset
   281
  "c\<noteq>0 ==> (c*a) / (c*b) = a / (b::'a::{field,division_by_zero})"
21328
73bb86d0f483 dropped Inductive dependency
haftmann
parents: 21258
diff changeset
   282
apply (cases "b = 0")
23413
5caa2710dd5b tuned laws for cancellation in divisions for fields.
nipkow
parents: 23406
diff changeset
   283
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
   284
done
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
   285
23413
5caa2710dd5b tuned laws for cancellation in divisions for fields.
nipkow
parents: 23406
diff changeset
   286
lemma mult_divide_mult_cancel_right:
23477
f4b83f03cac9 tuned and renamed group_eq_simps and ring_eq_simps
nipkow
parents: 23413
diff changeset
   287
  "c\<noteq>0 ==> (a*c) / (b*c) = a / (b::'a::{field,division_by_zero})"
21328
73bb86d0f483 dropped Inductive dependency
haftmann
parents: 21258
diff changeset
   288
apply (cases "b = 0")
23413
5caa2710dd5b tuned laws for cancellation in divisions for fields.
nipkow
parents: 23406
diff changeset
   289
apply (simp_all add: nonzero_mult_divide_mult_cancel_right)
14321
55c688d2eefa new theorems
paulson
parents: 14305
diff changeset
   290
done
23413
5caa2710dd5b tuned laws for cancellation in divisions for fields.
nipkow
parents: 23406
diff changeset
   291
24286
7619080e49f0 ATP blacklisting is now in theory data, attribute noatp
paulson
parents: 23879
diff changeset
   292
lemma divide_divide_eq_right [simp,noatp]:
23477
f4b83f03cac9 tuned and renamed group_eq_simps and ring_eq_simps
nipkow
parents: 23413
diff changeset
   293
  "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
   294
by (simp add: divide_inverse mult_ac)
14288
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
   295
24286
7619080e49f0 ATP blacklisting is now in theory data, attribute noatp
paulson
parents: 23879
diff changeset
   296
lemma divide_divide_eq_left [simp,noatp]:
23477
f4b83f03cac9 tuned and renamed group_eq_simps and ring_eq_simps
nipkow
parents: 23413
diff changeset
   297
  "(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
   298
by (simp add: divide_inverse mult_assoc)
14288
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
   299
23389
aaca6a8e5414 tuned proofs: avoid implicit prems;
wenzelm
parents: 23326
diff changeset
   300
15234
ec91a90c604e simplification tweaks for better arithmetic reasoning
paulson
parents: 15229
diff changeset
   301
subsubsection{*Special Cancellation Simprules for Division*}
ec91a90c604e simplification tweaks for better arithmetic reasoning
paulson
parents: 15229
diff changeset
   302
24427
bc5cf3b09ff3 revised blacklisting for ATP linkup
paulson
parents: 24422
diff changeset
   303
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
   304
fixes c :: "'a :: {field,division_by_zero}"
f4b83f03cac9 tuned and renamed group_eq_simps and ring_eq_simps
nipkow
parents: 23413
diff changeset
   305
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
   306
by (simp add: mult_divide_mult_cancel_left)
5caa2710dd5b tuned laws for cancellation in divisions for fields.
nipkow
parents: 23406
diff changeset
   307
15234
ec91a90c604e simplification tweaks for better arithmetic reasoning
paulson
parents: 15229
diff changeset
   308
14293
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
   309
subsection {* Division and Unary Minus *}
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
   310
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
   311
lemma minus_divide_right: "- (a/b) = a / -(b::'a::{field,division_by_zero})"
29407
5ef7e97fd9e4 move lemmas mult_minus{left,right} inside class ring
huffman
parents: 29406
diff changeset
   312
by (simp add: divide_inverse)
14430
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14421
diff changeset
   313
30630
4fbe1401bac2 move field lemmas into class locale context
huffman
parents: 30242
diff changeset
   314
lemma divide_minus_right [simp, noatp]:
4fbe1401bac2 move field lemmas into class locale context
huffman
parents: 30242
diff changeset
   315
  "a / -(b::'a::{field,division_by_zero}) = -(a / b)"
4fbe1401bac2 move field lemmas into class locale context
huffman
parents: 30242
diff changeset
   316
by (simp add: divide_inverse)
4fbe1401bac2 move field lemmas into class locale context
huffman
parents: 30242
diff changeset
   317
4fbe1401bac2 move field lemmas into class locale context
huffman
parents: 30242
diff changeset
   318
lemma minus_divide_divide:
23477
f4b83f03cac9 tuned and renamed group_eq_simps and ring_eq_simps
nipkow
parents: 23413
diff changeset
   319
  "(-a)/(-b) = a / (b::'a::{field,division_by_zero})"
21328
73bb86d0f483 dropped Inductive dependency
haftmann
parents: 21258
diff changeset
   320
apply (cases "b=0", simp) 
14293
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
   321
apply (simp add: nonzero_minus_divide_divide) 
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
   322
done
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
   323
23482
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
   324
lemma eq_divide_eq:
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
   325
  "((a::'a::{field,division_by_zero}) = b/c) = (if c\<noteq>0 then a*c = b else a=0)"
30630
4fbe1401bac2 move field lemmas into class locale context
huffman
parents: 30242
diff changeset
   326
by (simp add: nonzero_eq_divide_eq)
23482
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
   327
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
   328
lemma divide_eq_eq:
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
   329
  "(b/c = (a::'a::{field,division_by_zero})) = (if c\<noteq>0 then b = a*c else a=0)"
30630
4fbe1401bac2 move field lemmas into class locale context
huffman
parents: 30242
diff changeset
   330
by (force simp add: nonzero_divide_eq_eq)
14293
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
   331
23389
aaca6a8e5414 tuned proofs: avoid implicit prems;
wenzelm
parents: 23326
diff changeset
   332
14268
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   333
subsection {* Ordered Fields *}
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   334
14277
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
   335
lemma positive_imp_inverse_positive: 
35028
108662d50512 more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents: 34146
diff changeset
   336
assumes a_gt_0: "0 < a"  shows "0 < inverse (a::'a::linordered_field)"
23482
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
   337
proof -
14268
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   338
  have "0 < a * inverse a" 
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   339
    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
   340
  thus "0 < inverse a" 
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   341
    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
   342
qed
14268
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   343
14277
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
   344
lemma negative_imp_inverse_negative:
35028
108662d50512 more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents: 34146
diff changeset
   345
  "a < 0 ==> inverse a < (0::'a::linordered_field)"
23482
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
   346
by (insert positive_imp_inverse_positive [of "-a"], 
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
   347
    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
   348
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   349
lemma inverse_le_imp_le:
23482
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
   350
assumes invle: "inverse a \<le> inverse b" and apos:  "0 < a"
35028
108662d50512 more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents: 34146
diff changeset
   351
shows "b \<le> (a::'a::linordered_field)"
23482
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
   352
proof (rule classical)
14268
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   353
  assume "~ b \<le> a"
23482
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
   354
  hence "a < b"  by (simp add: linorder_not_le)
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
   355
  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
   356
  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
   357
    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
   358
  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
   359
    by (simp add: bpos order_less_imp_le mult_right_mono)
23482
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
   360
  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
   361
qed
14268
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   362
14277
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
   363
lemma inverse_positive_imp_positive:
23482
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
   364
assumes inv_gt_0: "0 < inverse a" and nz: "a \<noteq> 0"
35028
108662d50512 more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents: 34146
diff changeset
   365
shows "0 < (a::'a::linordered_field)"
23389
aaca6a8e5414 tuned proofs: avoid implicit prems;
wenzelm
parents: 23326
diff changeset
   366
proof -
14277
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
   367
  have "0 < inverse (inverse a)"
23389
aaca6a8e5414 tuned proofs: avoid implicit prems;
wenzelm
parents: 23326
diff changeset
   368
    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
   369
  thus "0 < a"
23389
aaca6a8e5414 tuned proofs: avoid implicit prems;
wenzelm
parents: 23326
diff changeset
   370
    using nz by (simp add: nonzero_inverse_inverse_eq)
aaca6a8e5414 tuned proofs: avoid implicit prems;
wenzelm
parents: 23326
diff changeset
   371
qed
14277
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
   372
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
   373
lemma inverse_positive_iff_positive [simp]:
35028
108662d50512 more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents: 34146
diff changeset
   374
  "(0 < inverse a) = (0 < (a::'a::{linordered_field,division_by_zero}))"
21328
73bb86d0f483 dropped Inductive dependency
haftmann
parents: 21258
diff changeset
   375
apply (cases "a = 0", simp)
14277
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
   376
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
   377
done
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
   378
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
   379
lemma inverse_negative_imp_negative:
23482
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
   380
assumes inv_less_0: "inverse a < 0" and nz:  "a \<noteq> 0"
35028
108662d50512 more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents: 34146
diff changeset
   381
shows "a < (0::'a::linordered_field)"
23389
aaca6a8e5414 tuned proofs: avoid implicit prems;
wenzelm
parents: 23326
diff changeset
   382
proof -
14277
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
   383
  have "inverse (inverse a) < 0"
23389
aaca6a8e5414 tuned proofs: avoid implicit prems;
wenzelm
parents: 23326
diff changeset
   384
    using inv_less_0 by (rule negative_imp_inverse_negative)
23482
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
   385
  thus "a < 0" using nz by (simp add: nonzero_inverse_inverse_eq)
23389
aaca6a8e5414 tuned proofs: avoid implicit prems;
wenzelm
parents: 23326
diff changeset
   386
qed
14277
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
   387
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
   388
lemma inverse_negative_iff_negative [simp]:
35028
108662d50512 more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents: 34146
diff changeset
   389
  "(inverse a < 0) = (a < (0::'a::{linordered_field,division_by_zero}))"
21328
73bb86d0f483 dropped Inductive dependency
haftmann
parents: 21258
diff changeset
   390
apply (cases "a = 0", simp)
14277
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
   391
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
   392
done
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
   393
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
   394
lemma inverse_nonnegative_iff_nonnegative [simp]:
35028
108662d50512 more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents: 34146
diff changeset
   395
  "(0 \<le> inverse a) = (0 \<le> (a::'a::{linordered_field,division_by_zero}))"
14277
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
   396
by (simp add: linorder_not_less [symmetric])
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
   397
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
   398
lemma inverse_nonpositive_iff_nonpositive [simp]:
35028
108662d50512 more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents: 34146
diff changeset
   399
  "(inverse a \<le> 0) = (a \<le> (0::'a::{linordered_field,division_by_zero}))"
14277
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
   400
by (simp add: linorder_not_less [symmetric])
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
   401
35043
07dbdf60d5ad dropped accidental duplication of "lin" prefix from cs. 108662d50512
haftmann
parents: 35032
diff changeset
   402
lemma linordered_field_no_lb: "\<forall> x. \<exists>y. y < (x::'a::linordered_field)"
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
   403
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
   404
  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
   405
  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
   406
  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
   407
  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
   408
  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
   409
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
   410
35043
07dbdf60d5ad dropped accidental duplication of "lin" prefix from cs. 108662d50512
haftmann
parents: 35032
diff changeset
   411
lemma linordered_field_no_ub: "\<forall> x. \<exists>y. y > (x::'a::linordered_field)"
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
   412
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
   413
  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
   414
  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
   415
  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
   416
  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
   417
  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
   418
qed
14277
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
   419
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
   420
subsection{*Anti-Monotonicity of @{term inverse}*}
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
   421
14268
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   422
lemma less_imp_inverse_less:
23482
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
   423
assumes less: "a < b" and apos:  "0 < a"
35028
108662d50512 more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents: 34146
diff changeset
   424
shows "inverse b < inverse (a::'a::linordered_field)"
23482
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
   425
proof (rule ccontr)
14268
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   426
  assume "~ inverse b < inverse a"
29667
53103fc8ffa3 Replaced group_ and ring_simps by algebra_simps;
nipkow
parents: 29465
diff changeset
   427
  hence "inverse a \<le> inverse b" 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
   428
  hence "~ (a < b)"
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   429
    by (simp add: linorder_not_less inverse_le_imp_le [OF _ apos])
29667
53103fc8ffa3 Replaced group_ and ring_simps by algebra_simps;
nipkow
parents: 29465
diff changeset
   430
  thus False by (rule notE [OF _ less])
23482
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
   431
qed
14268
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   432
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   433
lemma inverse_less_imp_less:
35028
108662d50512 more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents: 34146
diff changeset
   434
  "[|inverse a < inverse b; 0 < a|] ==> b < (a::'a::linordered_field)"
14268
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   435
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
   436
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
   437
done
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   438
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   439
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
   440
lemma inverse_less_iff_less [simp,noatp]:
35028
108662d50512 more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents: 34146
diff changeset
   441
  "[|0 < a; 0 < b|] ==> (inverse a < inverse b) = (b < (a::'a::linordered_field))"
14268
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   442
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
   443
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   444
lemma le_imp_inverse_le:
35028
108662d50512 more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents: 34146
diff changeset
   445
  "[|a \<le> b; 0 < a|] ==> inverse b \<le> inverse (a::'a::linordered_field)"
23482
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
   446
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
   447
24286
7619080e49f0 ATP blacklisting is now in theory data, attribute noatp
paulson
parents: 23879
diff changeset
   448
lemma inverse_le_iff_le [simp,noatp]:
35028
108662d50512 more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents: 34146
diff changeset
   449
 "[|0 < a; 0 < b|] ==> (inverse a \<le> inverse b) = (b \<le> (a::'a::linordered_field))"
14268
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   450
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
   451
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   452
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   453
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
   454
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
   455
lemma inverse_le_imp_le_neg:
35028
108662d50512 more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents: 34146
diff changeset
   456
  "[|inverse a \<le> inverse b; b < 0|] ==> b \<le> (a::'a::linordered_field)"
23482
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
   457
apply (rule classical) 
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
   458
apply (subgoal_tac "a < 0") 
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
   459
 prefer 2 apply (force simp add: linorder_not_le intro: order_less_trans) 
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
   460
apply (insert inverse_le_imp_le [of "-b" "-a"])
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
   461
apply (simp add: order_less_imp_not_eq nonzero_inverse_minus_eq) 
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
   462
done
14268
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   463
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   464
lemma less_imp_inverse_less_neg:
35028
108662d50512 more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents: 34146
diff changeset
   465
   "[|a < b; b < 0|] ==> inverse b < inverse (a::'a::linordered_field)"
23482
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
   466
apply (subgoal_tac "a < 0") 
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
   467
 prefer 2 apply (blast intro: order_less_trans) 
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
   468
apply (insert less_imp_inverse_less [of "-b" "-a"])
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
   469
apply (simp add: order_less_imp_not_eq nonzero_inverse_minus_eq) 
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
   470
done
14268
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   471
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   472
lemma inverse_less_imp_less_neg:
35028
108662d50512 more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents: 34146
diff changeset
   473
   "[|inverse a < inverse b; b < 0|] ==> b < (a::'a::linordered_field)"
23482
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
   474
apply (rule classical) 
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
   475
apply (subgoal_tac "a < 0") 
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
   476
 prefer 2
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
   477
 apply (force simp add: linorder_not_less intro: order_le_less_trans) 
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
   478
apply (insert inverse_less_imp_less [of "-b" "-a"])
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
   479
apply (simp add: order_less_imp_not_eq nonzero_inverse_minus_eq) 
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
   480
done
14268
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   481
24286
7619080e49f0 ATP blacklisting is now in theory data, attribute noatp
paulson
parents: 23879
diff changeset
   482
lemma inverse_less_iff_less_neg [simp,noatp]:
35028
108662d50512 more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents: 34146
diff changeset
   483
  "[|a < 0; b < 0|] ==> (inverse a < inverse b) = (b < (a::'a::linordered_field))"
23482
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
   484
apply (insert inverse_less_iff_less [of "-b" "-a"])
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
   485
apply (simp del: inverse_less_iff_less 
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
   486
            add: order_less_imp_not_eq nonzero_inverse_minus_eq)
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
   487
done
14268
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   488
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   489
lemma le_imp_inverse_le_neg:
35028
108662d50512 more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents: 34146
diff changeset
   490
  "[|a \<le> b; b < 0|] ==> inverse b \<le> inverse (a::'a::linordered_field)"
23482
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
   491
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
   492
24286
7619080e49f0 ATP blacklisting is now in theory data, attribute noatp
paulson
parents: 23879
diff changeset
   493
lemma inverse_le_iff_le_neg [simp,noatp]:
35028
108662d50512 more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents: 34146
diff changeset
   494
 "[|a < 0; b < 0|] ==> (inverse a \<le> inverse b) = (b \<le> (a::'a::linordered_field))"
14268
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   495
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
   496
14277
ad66687ece6e more field division lemmas transferred from Real to Ring_and_Field
paulson
parents: 14272
diff changeset
   497
14365
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14353
diff changeset
   498
subsection{*Inverses and the Number One*}
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14353
diff changeset
   499
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14353
diff changeset
   500
lemma one_less_inverse_iff:
35028
108662d50512 more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents: 34146
diff changeset
   501
  "(1 < inverse x) = (0 < x & x < (1::'a::{linordered_field,division_by_zero}))"
23482
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
   502
proof cases
14365
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14353
diff changeset
   503
  assume "0 < x"
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14353
diff changeset
   504
    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
   505
    show ?thesis by simp
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14353
diff changeset
   506
next
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14353
diff changeset
   507
  assume notless: "~ (0 < x)"
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14353
diff changeset
   508
  have "~ (1 < inverse x)"
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14353
diff changeset
   509
  proof
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14353
diff changeset
   510
    assume "1 < inverse x"
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14353
diff changeset
   511
    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
   512
    also have "... < 1" by (rule zero_less_one) 
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14353
diff changeset
   513
    finally show False by auto
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14353
diff changeset
   514
  qed
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14353
diff changeset
   515
  with notless show ?thesis by simp
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14353
diff changeset
   516
qed
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14353
diff changeset
   517
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14353
diff changeset
   518
lemma inverse_eq_1_iff [simp]:
23482
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
   519
  "(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
   520
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
   521
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14353
diff changeset
   522
lemma one_le_inverse_iff:
35028
108662d50512 more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents: 34146
diff changeset
   523
  "(1 \<le> inverse x) = (0 < x & x \<le> (1::'a::{linordered_field,division_by_zero}))"
14365
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14353
diff changeset
   524
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
   525
                    eq_commute [of 1]) 
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14353
diff changeset
   526
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14353
diff changeset
   527
lemma inverse_less_1_iff:
35028
108662d50512 more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents: 34146
diff changeset
   528
  "(inverse x < 1) = (x \<le> 0 | 1 < (x::'a::{linordered_field,division_by_zero}))"
14365
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14353
diff changeset
   529
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
   530
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14353
diff changeset
   531
lemma inverse_le_1_iff:
35028
108662d50512 more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents: 34146
diff changeset
   532
  "(inverse x \<le> 1) = (x \<le> 0 | 1 \<le> (x::'a::{linordered_field,division_by_zero}))"
14365
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14353
diff changeset
   533
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
   534
23389
aaca6a8e5414 tuned proofs: avoid implicit prems;
wenzelm
parents: 23326
diff changeset
   535
14288
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
   536
subsection{*Simplification of Inequalities Involving Literal Divisors*}
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
   537
35028
108662d50512 more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents: 34146
diff changeset
   538
lemma pos_le_divide_eq: "0 < (c::'a::linordered_field) ==> (a \<le> b/c) = (a*c \<le> b)"
14288
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
   539
proof -
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
   540
  assume less: "0<c"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
   541
  hence "(a \<le> b/c) = (a*c \<le> (b/c)*c)"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
   542
    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
   543
  also have "... = (a*c \<le> b)"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
   544
    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
   545
  finally show ?thesis .
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
   546
qed
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
   547
35028
108662d50512 more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents: 34146
diff changeset
   548
lemma neg_le_divide_eq: "c < (0::'a::linordered_field) ==> (a \<le> b/c) = (b \<le> a*c)"
14288
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
   549
proof -
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
   550
  assume less: "c<0"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
   551
  hence "(a \<le> b/c) = ((b/c)*c \<le> a*c)"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
   552
    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
   553
  also have "... = (b \<le> a*c)"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
   554
    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
   555
  finally show ?thesis .
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
   556
qed
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
   557
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
   558
lemma le_divide_eq:
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
   559
  "(a \<le> b/c) = 
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
   560
   (if 0 < c then a*c \<le> b
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
   561
             else if c < 0 then b \<le> a*c
35028
108662d50512 more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents: 34146
diff changeset
   562
             else  a \<le> (0::'a::{linordered_field,division_by_zero}))"
21328
73bb86d0f483 dropped Inductive dependency
haftmann
parents: 21258
diff changeset
   563
apply (cases "c=0", simp) 
14288
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
   564
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
   565
done
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
   566
35028
108662d50512 more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents: 34146
diff changeset
   567
lemma pos_divide_le_eq: "0 < (c::'a::linordered_field) ==> (b/c \<le> a) = (b \<le> a*c)"
14288
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
   568
proof -
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
   569
  assume less: "0<c"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
   570
  hence "(b/c \<le> a) = ((b/c)*c \<le> a*c)"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
   571
    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
   572
  also have "... = (b \<le> a*c)"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
   573
    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
   574
  finally show ?thesis .
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
   575
qed
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
   576
35028
108662d50512 more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents: 34146
diff changeset
   577
lemma neg_divide_le_eq: "c < (0::'a::linordered_field) ==> (b/c \<le> a) = (a*c \<le> b)"
14288
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
   578
proof -
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
   579
  assume less: "c<0"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
   580
  hence "(b/c \<le> a) = (a*c \<le> (b/c)*c)"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
   581
    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
   582
  also have "... = (a*c \<le> b)"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
   583
    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
   584
  finally show ?thesis .
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
   585
qed
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
   586
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
   587
lemma divide_le_eq:
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
   588
  "(b/c \<le> a) = 
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
   589
   (if 0 < c then b \<le> a*c
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
   590
             else if c < 0 then a*c \<le> b
35028
108662d50512 more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents: 34146
diff changeset
   591
             else 0 \<le> (a::'a::{linordered_field,division_by_zero}))"
21328
73bb86d0f483 dropped Inductive dependency
haftmann
parents: 21258
diff changeset
   592
apply (cases "c=0", simp) 
14288
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
   593
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
   594
done
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
   595
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
   596
lemma pos_less_divide_eq:
35028
108662d50512 more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents: 34146
diff changeset
   597
     "0 < (c::'a::linordered_field) ==> (a < b/c) = (a*c < b)"
14288
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
   598
proof -
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
   599
  assume less: "0<c"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
   600
  hence "(a < b/c) = (a*c < (b/c)*c)"
15234
ec91a90c604e simplification tweaks for better arithmetic reasoning
paulson
parents: 15229
diff changeset
   601
    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
   602
  also have "... = (a*c < b)"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
   603
    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
   604
  finally show ?thesis .
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
   605
qed
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
   606
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
   607
lemma neg_less_divide_eq:
35028
108662d50512 more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents: 34146
diff changeset
   608
 "c < (0::'a::linordered_field) ==> (a < b/c) = (b < a*c)"
14288
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
   609
proof -
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
   610
  assume less: "c<0"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
   611
  hence "(a < b/c) = ((b/c)*c < a*c)"
15234
ec91a90c604e simplification tweaks for better arithmetic reasoning
paulson
parents: 15229
diff changeset
   612
    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
   613
  also have "... = (b < a*c)"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
   614
    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
   615
  finally show ?thesis .
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
   616
qed
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
   617
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
   618
lemma less_divide_eq:
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
   619
  "(a < b/c) = 
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
   620
   (if 0 < c then a*c < b
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
   621
             else if c < 0 then b < a*c
35028
108662d50512 more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents: 34146
diff changeset
   622
             else  a < (0::'a::{linordered_field,division_by_zero}))"
21328
73bb86d0f483 dropped Inductive dependency
haftmann
parents: 21258
diff changeset
   623
apply (cases "c=0", simp) 
14288
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
   624
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
   625
done
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
   626
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
   627
lemma pos_divide_less_eq:
35028
108662d50512 more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents: 34146
diff changeset
   628
     "0 < (c::'a::linordered_field) ==> (b/c < a) = (b < a*c)"
14288
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
   629
proof -
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
   630
  assume less: "0<c"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
   631
  hence "(b/c < a) = ((b/c)*c < a*c)"
15234
ec91a90c604e simplification tweaks for better arithmetic reasoning
paulson
parents: 15229
diff changeset
   632
    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
   633
  also have "... = (b < a*c)"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
   634
    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
   635
  finally show ?thesis .
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
   636
qed
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
   637
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
   638
lemma neg_divide_less_eq:
35028
108662d50512 more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents: 34146
diff changeset
   639
 "c < (0::'a::linordered_field) ==> (b/c < a) = (a*c < b)"
14288
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
   640
proof -
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
   641
  assume less: "c<0"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
   642
  hence "(b/c < a) = (a*c < (b/c)*c)"
15234
ec91a90c604e simplification tweaks for better arithmetic reasoning
paulson
parents: 15229
diff changeset
   643
    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
   644
  also have "... = (a*c < b)"
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
   645
    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
   646
  finally show ?thesis .
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
   647
qed
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
   648
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
   649
lemma divide_less_eq:
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
   650
  "(b/c < a) = 
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
   651
   (if 0 < c then b < a*c
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
   652
             else if c < 0 then a*c < b
35028
108662d50512 more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents: 34146
diff changeset
   653
             else 0 < (a::'a::{linordered_field,division_by_zero}))"
21328
73bb86d0f483 dropped Inductive dependency
haftmann
parents: 21258
diff changeset
   654
apply (cases "c=0", simp) 
14288
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
   655
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
   656
done
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
   657
23482
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
   658
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
   659
subsection{*Field simplification*}
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
   660
29667
53103fc8ffa3 Replaced group_ and ring_simps by algebra_simps;
nipkow
parents: 29465
diff changeset
   661
text{* Lemmas @{text field_simps} multiply with denominators in in(equations)
53103fc8ffa3 Replaced group_ and ring_simps by algebra_simps;
nipkow
parents: 29465
diff changeset
   662
if they can be proved to be non-zero (for equations) or positive/negative
53103fc8ffa3 Replaced group_ and ring_simps by algebra_simps;
nipkow
parents: 29465
diff changeset
   663
(for inequations). Can be too aggressive and is therefore separate from the
53103fc8ffa3 Replaced group_ and ring_simps by algebra_simps;
nipkow
parents: 29465
diff changeset
   664
more benign @{text algebra_simps}. *}
14288
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
   665
29833
409138c4de12 added noatps
nipkow
parents: 29700
diff changeset
   666
lemmas field_simps[noatp] = field_eq_simps
23482
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
   667
  (* multiply ineqn *)
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
   668
  pos_divide_less_eq neg_divide_less_eq
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
   669
  pos_less_divide_eq neg_less_divide_eq
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
   670
  pos_divide_le_eq neg_divide_le_eq
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
   671
  pos_le_divide_eq neg_le_divide_eq
14288
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
   672
23482
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
   673
text{* Lemmas @{text sign_simps} is a first attempt to automate proofs
23483
a9356b40fbd3 tex problem fixed
nipkow
parents: 23482
diff changeset
   674
of positivity/negativity needed for @{text field_simps}. Have not added @{text
23482
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
   675
sign_simps} to @{text field_simps} because the former can lead to case
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
   676
explosions. *}
14288
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
   677
29833
409138c4de12 added noatps
nipkow
parents: 29700
diff changeset
   678
lemmas sign_simps[noatp] = group_simps
23482
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
   679
  zero_less_mult_iff  mult_less_0_iff
14288
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
   680
23482
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
   681
(* Only works once linear arithmetic is installed:
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
   682
text{*An example:*}
35028
108662d50512 more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents: 34146
diff changeset
   683
lemma fixes a b c d e f :: "'a::linordered_field"
23482
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
   684
shows "\<lbrakk>a>b; c<d; e<f; 0 < u \<rbrakk> \<Longrightarrow>
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
   685
 ((a-b)*(c-d)*(e-f))/((c-d)*(e-f)*(a-b)) <
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
   686
 ((e-f)*(a-b)*(c-d))/((e-f)*(a-b)*(c-d)) + u"
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
   687
apply(subgoal_tac "(c-d)*(e-f)*(a-b) > 0")
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
   688
 prefer 2 apply(simp add:sign_simps)
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
   689
apply(subgoal_tac "(c-d)*(e-f)*(a-b)*u > 0")
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
   690
 prefer 2 apply(simp add:sign_simps)
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
   691
apply(simp add:field_simps)
16775
c1b87ef4a1c3 added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
avigad
parents: 16568
diff changeset
   692
done
23482
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
   693
*)
16775
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
16775
c1b87ef4a1c3 added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
avigad
parents: 16568
diff changeset
   696
subsection{*Division and Signs*}
c1b87ef4a1c3 added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
avigad
parents: 16568
diff changeset
   697
c1b87ef4a1c3 added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
avigad
parents: 16568
diff changeset
   698
lemma zero_less_divide_iff:
35028
108662d50512 more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents: 34146
diff changeset
   699
     "((0::'a::{linordered_field,division_by_zero}) < 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
   700
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
   701
c1b87ef4a1c3 added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
avigad
parents: 16568
diff changeset
   702
lemma divide_less_0_iff:
35028
108662d50512 more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents: 34146
diff changeset
   703
     "(a/b < (0::'a::{linordered_field,division_by_zero})) = 
16775
c1b87ef4a1c3 added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
avigad
parents: 16568
diff changeset
   704
      (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
   705
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
   706
c1b87ef4a1c3 added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
avigad
parents: 16568
diff changeset
   707
lemma zero_le_divide_iff:
35028
108662d50512 more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents: 34146
diff changeset
   708
     "((0::'a::{linordered_field,division_by_zero}) \<le> a/b) =
16775
c1b87ef4a1c3 added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
avigad
parents: 16568
diff changeset
   709
      (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
   710
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
   711
c1b87ef4a1c3 added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
avigad
parents: 16568
diff changeset
   712
lemma divide_le_0_iff:
35028
108662d50512 more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents: 34146
diff changeset
   713
     "(a/b \<le> (0::'a::{linordered_field,division_by_zero})) =
16775
c1b87ef4a1c3 added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
avigad
parents: 16568
diff changeset
   714
      (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
   715
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
   716
24286
7619080e49f0 ATP blacklisting is now in theory data, attribute noatp
paulson
parents: 23879
diff changeset
   717
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
   718
     "(a/b = 0) = (a=0 | b=(0::'a::{field,division_by_zero}))"
23482
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
   719
by (simp add: divide_inverse)
16775
c1b87ef4a1c3 added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
avigad
parents: 16568
diff changeset
   720
23482
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
   721
lemma divide_pos_pos:
35028
108662d50512 more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents: 34146
diff changeset
   722
  "0 < (x::'a::linordered_field) ==> 0 < y ==> 0 < x / y"
23482
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
   723
by(simp add:field_simps)
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
   724
16775
c1b87ef4a1c3 added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
avigad
parents: 16568
diff changeset
   725
23482
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
   726
lemma divide_nonneg_pos:
35028
108662d50512 more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents: 34146
diff changeset
   727
  "0 <= (x::'a::linordered_field) ==> 0 < y ==> 0 <= x / y"
23482
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
   728
by(simp add:field_simps)
16775
c1b87ef4a1c3 added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
avigad
parents: 16568
diff changeset
   729
23482
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
   730
lemma divide_neg_pos:
35028
108662d50512 more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents: 34146
diff changeset
   731
  "(x::'a::linordered_field) < 0 ==> 0 < y ==> x / y < 0"
23482
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
   732
by(simp add:field_simps)
16775
c1b87ef4a1c3 added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
avigad
parents: 16568
diff changeset
   733
23482
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
   734
lemma divide_nonpos_pos:
35028
108662d50512 more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents: 34146
diff changeset
   735
  "(x::'a::linordered_field) <= 0 ==> 0 < y ==> x / y <= 0"
23482
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
   736
by(simp add:field_simps)
16775
c1b87ef4a1c3 added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
avigad
parents: 16568
diff changeset
   737
23482
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
   738
lemma divide_pos_neg:
35028
108662d50512 more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents: 34146
diff changeset
   739
  "0 < (x::'a::linordered_field) ==> y < 0 ==> x / y < 0"
23482
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
   740
by(simp add:field_simps)
16775
c1b87ef4a1c3 added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
avigad
parents: 16568
diff changeset
   741
23482
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
   742
lemma divide_nonneg_neg:
35028
108662d50512 more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents: 34146
diff changeset
   743
  "0 <= (x::'a::linordered_field) ==> y < 0 ==> x / y <= 0" 
23482
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
   744
by(simp add:field_simps)
16775
c1b87ef4a1c3 added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
avigad
parents: 16568
diff changeset
   745
23482
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
   746
lemma divide_neg_neg:
35028
108662d50512 more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents: 34146
diff changeset
   747
  "(x::'a::linordered_field) < 0 ==> y < 0 ==> 0 < x / y"
23482
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
   748
by(simp add:field_simps)
16775
c1b87ef4a1c3 added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
avigad
parents: 16568
diff changeset
   749
23482
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
   750
lemma divide_nonpos_neg:
35028
108662d50512 more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents: 34146
diff changeset
   751
  "(x::'a::linordered_field) <= 0 ==> y < 0 ==> 0 <= x / y"
23482
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
   752
by(simp add:field_simps)
15234
ec91a90c604e simplification tweaks for better arithmetic reasoning
paulson
parents: 15229
diff changeset
   753
23389
aaca6a8e5414 tuned proofs: avoid implicit prems;
wenzelm
parents: 23326
diff changeset
   754
14288
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
   755
subsection{*Cancellation Laws for Division*}
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
   756
24286
7619080e49f0 ATP blacklisting is now in theory data, attribute noatp
paulson
parents: 23879
diff changeset
   757
lemma divide_cancel_right [simp,noatp]:
14288
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
   758
     "(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
   759
apply (cases "c=0", simp)
23496
84e9216a6d0e removed redundant lemmas
nipkow
parents: 23483
diff changeset
   760
apply (simp add: divide_inverse)
14288
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
   761
done
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
   762
24286
7619080e49f0 ATP blacklisting is now in theory data, attribute noatp
paulson
parents: 23879
diff changeset
   763
lemma divide_cancel_left [simp,noatp]:
14288
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
   764
     "(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
   765
apply (cases "c=0", simp)
23496
84e9216a6d0e removed redundant lemmas
nipkow
parents: 23483
diff changeset
   766
apply (simp add: divide_inverse)
14288
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
   767
done
d149e3cbdb39 Moving some theorems from Real/RealArith0.ML
paulson
parents: 14284
diff changeset
   768
23389
aaca6a8e5414 tuned proofs: avoid implicit prems;
wenzelm
parents: 23326
diff changeset
   769
14353
79f9fbef9106 Added lemmas to Ring_and_Field with slightly modified simplification rules
paulson
parents: 14348
diff changeset
   770
subsection {* Division and the Number One *}
79f9fbef9106 Added lemmas to Ring_and_Field with slightly modified simplification rules
paulson
parents: 14348
diff changeset
   771
79f9fbef9106 Added lemmas to Ring_and_Field with slightly modified simplification rules
paulson
parents: 14348
diff changeset
   772
text{*Simplify expressions equated with 1*}
24286
7619080e49f0 ATP blacklisting is now in theory data, attribute noatp
paulson
parents: 23879
diff changeset
   773
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
   774
     "(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
   775
apply (cases "b=0", simp)
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
   776
apply (simp add: right_inverse_eq)
14353
79f9fbef9106 Added lemmas to Ring_and_Field with slightly modified simplification rules
paulson
parents: 14348
diff changeset
   777
done
79f9fbef9106 Added lemmas to Ring_and_Field with slightly modified simplification rules
paulson
parents: 14348
diff changeset
   778
24286
7619080e49f0 ATP blacklisting is now in theory data, attribute noatp
paulson
parents: 23879
diff changeset
   779
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
   780
     "(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
   781
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
   782
24286
7619080e49f0 ATP blacklisting is now in theory data, attribute noatp
paulson
parents: 23879
diff changeset
   783
lemma zero_eq_1_divide_iff [simp,noatp]:
35028
108662d50512 more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents: 34146
diff changeset
   784
     "((0::'a::{linordered_field,division_by_zero}) = 1/a) = (a = 0)"
23482
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
   785
apply (cases "a=0", simp)
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
   786
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
   787
done
79f9fbef9106 Added lemmas to Ring_and_Field with slightly modified simplification rules
paulson
parents: 14348
diff changeset
   788
24286
7619080e49f0 ATP blacklisting is now in theory data, attribute noatp
paulson
parents: 23879
diff changeset
   789
lemma one_divide_eq_0_iff [simp,noatp]:
35028
108662d50512 more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents: 34146
diff changeset
   790
     "(1/a = (0::'a::{linordered_field,division_by_zero})) = (a = 0)"
23482
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
   791
apply (cases "a=0", simp)
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
   792
apply (insert zero_neq_one [THEN not_sym])
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
   793
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
   794
done
79f9fbef9106 Added lemmas to Ring_and_Field with slightly modified simplification rules
paulson
parents: 14348
diff changeset
   795
79f9fbef9106 Added lemmas to Ring_and_Field with slightly modified simplification rules
paulson
parents: 14348
diff changeset
   796
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
   797
lemmas zero_less_divide_1_iff = zero_less_divide_iff [of 1, simplified]
9a5419d5ca01 simplified the special-case simprules
paulson
parents: 17085
diff changeset
   798
lemmas divide_less_0_1_iff = divide_less_0_iff [of 1, simplified]
9a5419d5ca01 simplified the special-case simprules
paulson
parents: 17085
diff changeset
   799
lemmas zero_le_divide_1_iff = zero_le_divide_iff [of 1, simplified]
9a5419d5ca01 simplified the special-case simprules
paulson
parents: 17085
diff changeset
   800
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
   801
29833
409138c4de12 added noatps
nipkow
parents: 29700
diff changeset
   802
declare zero_less_divide_1_iff [simp,noatp]
24286
7619080e49f0 ATP blacklisting is now in theory data, attribute noatp
paulson
parents: 23879
diff changeset
   803
declare divide_less_0_1_iff [simp,noatp]
29833
409138c4de12 added noatps
nipkow
parents: 29700
diff changeset
   804
declare zero_le_divide_1_iff [simp,noatp]
24286
7619080e49f0 ATP blacklisting is now in theory data, attribute noatp
paulson
parents: 23879
diff changeset
   805
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
   806
23389
aaca6a8e5414 tuned proofs: avoid implicit prems;
wenzelm
parents: 23326
diff changeset
   807
14293
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
   808
subsection {* Ordering Rules for Division *}
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
   809
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
   810
lemma divide_strict_right_mono:
35028
108662d50512 more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents: 34146
diff changeset
   811
     "[|a < b; 0 < c|] ==> a / c < b / (c::'a::linordered_field)"
14293
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
   812
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
   813
              positive_imp_inverse_positive)
14293
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
   814
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
   815
lemma divide_right_mono:
35028
108662d50512 more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents: 34146
diff changeset
   816
     "[|a \<le> b; 0 \<le> c|] ==> a/c \<le> b/(c::'a::{linordered_field,division_by_zero})"
23482
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
   817
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
   818
35028
108662d50512 more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents: 34146
diff changeset
   819
lemma divide_right_mono_neg: "(a::'a::{division_by_zero,linordered_field}) <= b 
16775
c1b87ef4a1c3 added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
avigad
parents: 16568
diff changeset
   820
    ==> c <= 0 ==> b / c <= a / c"
23482
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
   821
apply (drule divide_right_mono [of _ _ "- c"])
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
   822
apply auto
16775
c1b87ef4a1c3 added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
avigad
parents: 16568
diff changeset
   823
done
c1b87ef4a1c3 added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
avigad
parents: 16568
diff changeset
   824
c1b87ef4a1c3 added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
avigad
parents: 16568
diff changeset
   825
lemma divide_strict_right_mono_neg:
35028
108662d50512 more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents: 34146
diff changeset
   826
     "[|b < a; c < 0|] ==> a / c < b / (c::'a::linordered_field)"
23482
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
   827
apply (drule divide_strict_right_mono [of _ _ "-c"], simp)
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
   828
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
   829
done
14293
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
   830
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
   831
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
   832
      have the same sign*}
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
   833
lemma divide_strict_left_mono:
35028
108662d50512 more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents: 34146
diff changeset
   834
  "[|b < a; 0 < c; 0 < a*b|] ==> c / a < c / (b::'a::linordered_field)"
23482
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
   835
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
   836
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
   837
lemma divide_left_mono:
35028
108662d50512 more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents: 34146
diff changeset
   838
  "[|b \<le> a; 0 \<le> c; 0 < a*b|] ==> c / a \<le> c / (b::'a::linordered_field)"
23482
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
   839
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
   840
35028
108662d50512 more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents: 34146
diff changeset
   841
lemma divide_left_mono_neg: "(a::'a::{division_by_zero,linordered_field}) <= b 
16775
c1b87ef4a1c3 added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
avigad
parents: 16568
diff changeset
   842
    ==> 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
   843
  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
   844
  apply (auto simp add: mult_commute)
c1b87ef4a1c3 added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
avigad
parents: 16568
diff changeset
   845
done
c1b87ef4a1c3 added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
avigad
parents: 16568
diff changeset
   846
14293
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
   847
lemma divide_strict_left_mono_neg:
35028
108662d50512 more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents: 34146
diff changeset
   848
  "[|a < b; c < 0; 0 < a*b|] ==> c / a < c / (b::'a::linordered_field)"
23482
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
   849
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
   850
14293
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
   851
16775
c1b87ef4a1c3 added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
avigad
parents: 16568
diff changeset
   852
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
   853
24286
7619080e49f0 ATP blacklisting is now in theory data, attribute noatp
paulson
parents: 23879
diff changeset
   854
lemma le_divide_eq_1 [noatp]:
35028
108662d50512 more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents: 34146
diff changeset
   855
  fixes a :: "'a :: {linordered_field,division_by_zero}"
16775
c1b87ef4a1c3 added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
avigad
parents: 16568
diff changeset
   856
  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
   857
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
   858
24286
7619080e49f0 ATP blacklisting is now in theory data, attribute noatp
paulson
parents: 23879
diff changeset
   859
lemma divide_le_eq_1 [noatp]:
35028
108662d50512 more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents: 34146
diff changeset
   860
  fixes a :: "'a :: {linordered_field,division_by_zero}"
16775
c1b87ef4a1c3 added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
avigad
parents: 16568
diff changeset
   861
  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
   862
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
   863
24286
7619080e49f0 ATP blacklisting is now in theory data, attribute noatp
paulson
parents: 23879
diff changeset
   864
lemma less_divide_eq_1 [noatp]:
35028
108662d50512 more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents: 34146
diff changeset
   865
  fixes a :: "'a :: {linordered_field,division_by_zero}"
16775
c1b87ef4a1c3 added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
avigad
parents: 16568
diff changeset
   866
  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
   867
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
   868
24286
7619080e49f0 ATP blacklisting is now in theory data, attribute noatp
paulson
parents: 23879
diff changeset
   869
lemma divide_less_eq_1 [noatp]:
35028
108662d50512 more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents: 34146
diff changeset
   870
  fixes a :: "'a :: {linordered_field,division_by_zero}"
16775
c1b87ef4a1c3 added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
avigad
parents: 16568
diff changeset
   871
  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
   872
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
   873
23389
aaca6a8e5414 tuned proofs: avoid implicit prems;
wenzelm
parents: 23326
diff changeset
   874
16775
c1b87ef4a1c3 added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
avigad
parents: 16568
diff changeset
   875
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
   876
24286
7619080e49f0 ATP blacklisting is now in theory data, attribute noatp
paulson
parents: 23879
diff changeset
   877
lemma le_divide_eq_1_pos [simp,noatp]:
35028
108662d50512 more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents: 34146
diff changeset
   878
  fixes a :: "'a :: {linordered_field,division_by_zero}"
18649
bb99c2e705ca tidied, and added missing thm divide_less_eq_1_neg
paulson
parents: 18623
diff changeset
   879
  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
   880
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
   881
24286
7619080e49f0 ATP blacklisting is now in theory data, attribute noatp
paulson
parents: 23879
diff changeset
   882
lemma le_divide_eq_1_neg [simp,noatp]:
35028
108662d50512 more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents: 34146
diff changeset
   883
  fixes a :: "'a :: {linordered_field,division_by_zero}"
18649
bb99c2e705ca tidied, and added missing thm divide_less_eq_1_neg
paulson
parents: 18623
diff changeset
   884
  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
   885
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
   886
24286
7619080e49f0 ATP blacklisting is now in theory data, attribute noatp
paulson
parents: 23879
diff changeset
   887
lemma divide_le_eq_1_pos [simp,noatp]:
35028
108662d50512 more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents: 34146
diff changeset
   888
  fixes a :: "'a :: {linordered_field,division_by_zero}"
18649
bb99c2e705ca tidied, and added missing thm divide_less_eq_1_neg
paulson
parents: 18623
diff changeset
   889
  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
   890
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
   891
24286
7619080e49f0 ATP blacklisting is now in theory data, attribute noatp
paulson
parents: 23879
diff changeset
   892
lemma divide_le_eq_1_neg [simp,noatp]:
35028
108662d50512 more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents: 34146
diff changeset
   893
  fixes a :: "'a :: {linordered_field,division_by_zero}"
18649
bb99c2e705ca tidied, and added missing thm divide_less_eq_1_neg
paulson
parents: 18623
diff changeset
   894
  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
   895
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
   896
24286
7619080e49f0 ATP blacklisting is now in theory data, attribute noatp
paulson
parents: 23879
diff changeset
   897
lemma less_divide_eq_1_pos [simp,noatp]:
35028
108662d50512 more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents: 34146
diff changeset
   898
  fixes a :: "'a :: {linordered_field,division_by_zero}"
18649
bb99c2e705ca tidied, and added missing thm divide_less_eq_1_neg
paulson
parents: 18623
diff changeset
   899
  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
   900
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
   901
24286
7619080e49f0 ATP blacklisting is now in theory data, attribute noatp
paulson
parents: 23879
diff changeset
   902
lemma less_divide_eq_1_neg [simp,noatp]:
35028
108662d50512 more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents: 34146
diff changeset
   903
  fixes a :: "'a :: {linordered_field,division_by_zero}"
18649
bb99c2e705ca tidied, and added missing thm divide_less_eq_1_neg
paulson
parents: 18623
diff changeset
   904
  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
   905
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
   906
24286
7619080e49f0 ATP blacklisting is now in theory data, attribute noatp
paulson
parents: 23879
diff changeset
   907
lemma divide_less_eq_1_pos [simp,noatp]:
35028
108662d50512 more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents: 34146
diff changeset
   908
  fixes a :: "'a :: {linordered_field,division_by_zero}"
18649
bb99c2e705ca tidied, and added missing thm divide_less_eq_1_neg
paulson
parents: 18623
diff changeset
   909
  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
   910
by (auto simp add: divide_less_eq)
bb99c2e705ca tidied, and added missing thm divide_less_eq_1_neg
paulson
parents: 18623
diff changeset
   911
24286
7619080e49f0 ATP blacklisting is now in theory data, attribute noatp
paulson
parents: 23879
diff changeset
   912
lemma divide_less_eq_1_neg [simp,noatp]:
35028
108662d50512 more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents: 34146
diff changeset
   913
  fixes a :: "'a :: {linordered_field,division_by_zero}"
18649
bb99c2e705ca tidied, and added missing thm divide_less_eq_1_neg
paulson
parents: 18623
diff changeset
   914
  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
   915
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
   916
24286
7619080e49f0 ATP blacklisting is now in theory data, attribute noatp
paulson
parents: 23879
diff changeset
   917
lemma eq_divide_eq_1 [simp,noatp]:
35028
108662d50512 more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents: 34146
diff changeset
   918
  fixes a :: "'a :: {linordered_field,division_by_zero}"
18649
bb99c2e705ca tidied, and added missing thm divide_less_eq_1_neg
paulson
parents: 18623
diff changeset
   919
  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
   920
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
   921
24286
7619080e49f0 ATP blacklisting is now in theory data, attribute noatp
paulson
parents: 23879
diff changeset
   922
lemma divide_eq_eq_1 [simp,noatp]:
35028
108662d50512 more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents: 34146
diff changeset
   923
  fixes a :: "'a :: {linordered_field,division_by_zero}"
18649
bb99c2e705ca tidied, and added missing thm divide_less_eq_1_neg
paulson
parents: 18623
diff changeset
   924
  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
   925
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
   926
23389
aaca6a8e5414 tuned proofs: avoid implicit prems;
wenzelm
parents: 23326
diff changeset
   927
16775
c1b87ef4a1c3 added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
avigad
parents: 16568
diff changeset
   928
subsection {* Reasoning about inequalities with division *}
c1b87ef4a1c3 added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
avigad
parents: 16568
diff changeset
   929
35028
108662d50512 more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents: 34146
diff changeset
   930
lemma mult_imp_div_pos_le: "0 < (y::'a::linordered_field) ==> x <= z * y ==>
33319
74f0dcc0b5fb moved algebraic classes to Ring_and_Field
haftmann
parents: 32960
diff changeset
   931
    x / y <= z"
74f0dcc0b5fb moved algebraic classes to Ring_and_Field
haftmann
parents: 32960
diff changeset
   932
by (subst pos_divide_le_eq, assumption+)
16775
c1b87ef4a1c3 added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
avigad
parents: 16568
diff changeset
   933
35028
108662d50512 more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents: 34146
diff changeset
   934
lemma mult_imp_le_div_pos: "0 < (y::'a::linordered_field) ==> z * y <= x ==>
23482
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
   935
    z <= x / y"
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
   936
by(simp add:field_simps)
16775
c1b87ef4a1c3 added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
avigad
parents: 16568
diff changeset
   937
35028
108662d50512 more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents: 34146
diff changeset
   938
lemma mult_imp_div_pos_less: "0 < (y::'a::linordered_field) ==> x < z * y ==>
16775
c1b87ef4a1c3 added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
avigad
parents: 16568
diff changeset
   939
    x / y < z"
23482
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
   940
by(simp add:field_simps)
16775
c1b87ef4a1c3 added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
avigad
parents: 16568
diff changeset
   941
35028
108662d50512 more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents: 34146
diff changeset
   942
lemma mult_imp_less_div_pos: "0 < (y::'a::linordered_field) ==> z * y < x ==>
16775
c1b87ef4a1c3 added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
avigad
parents: 16568
diff changeset
   943
    z < x / y"
23482
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
   944
by(simp add:field_simps)
16775
c1b87ef4a1c3 added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
avigad
parents: 16568
diff changeset
   945
35028
108662d50512 more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents: 34146
diff changeset
   946
lemma frac_le: "(0::'a::linordered_field) <= x ==> 
16775
c1b87ef4a1c3 added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
avigad
parents: 16568
diff changeset
   947
    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
   948
  apply (rule mult_imp_div_pos_le)
25230
022029099a83 continued localization
haftmann
parents: 25193
diff changeset
   949
  apply simp
022029099a83 continued localization
haftmann
parents: 25193
diff changeset
   950
  apply (subst times_divide_eq_left)
16775
c1b87ef4a1c3 added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
avigad
parents: 16568
diff changeset
   951
  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
   952
  apply (rule mult_mono)
c1b87ef4a1c3 added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
avigad
parents: 16568
diff changeset
   953
  apply simp_all
14293
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
   954
done
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
   955
35028
108662d50512 more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents: 34146
diff changeset
   956
lemma frac_less: "(0::'a::linordered_field) <= x ==> 
16775
c1b87ef4a1c3 added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
avigad
parents: 16568
diff changeset
   957
    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
   958
  apply (rule mult_imp_div_pos_less)
33319
74f0dcc0b5fb moved algebraic classes to Ring_and_Field
haftmann
parents: 32960
diff changeset
   959
  apply simp
74f0dcc0b5fb moved algebraic classes to Ring_and_Field
haftmann
parents: 32960
diff changeset
   960
  apply (subst times_divide_eq_left)
16775
c1b87ef4a1c3 added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
avigad
parents: 16568
diff changeset
   961
  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
   962
  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
   963
  apply simp_all
c1b87ef4a1c3 added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
avigad
parents: 16568
diff changeset
   964
done
c1b87ef4a1c3 added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
avigad
parents: 16568
diff changeset
   965
35028
108662d50512 more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents: 34146
diff changeset
   966
lemma frac_less2: "(0::'a::linordered_field) < x ==> 
16775
c1b87ef4a1c3 added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
avigad
parents: 16568
diff changeset
   967
    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
   968
  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
   969
  apply simp_all
33319
74f0dcc0b5fb moved algebraic classes to Ring_and_Field
haftmann
parents: 32960
diff changeset
   970
  apply (subst times_divide_eq_left)
16775
c1b87ef4a1c3 added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
avigad
parents: 16568
diff changeset
   971
  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
   972
  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
   973
  apply simp_all
c1b87ef4a1c3 added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
avigad
parents: 16568
diff changeset
   974
done
c1b87ef4a1c3 added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
avigad
parents: 16568
diff changeset
   975
c1b87ef4a1c3 added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
avigad
parents: 16568
diff changeset
   976
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
   977
  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
   978
  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
   979
  seem to need them.*}
c1b87ef4a1c3 added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
avigad
parents: 16568
diff changeset
   980
c1b87ef4a1c3 added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
avigad
parents: 16568
diff changeset
   981
declare times_divide_eq [simp]
14293
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
   982
23389
aaca6a8e5414 tuned proofs: avoid implicit prems;
wenzelm
parents: 23326
diff changeset
   983
14293
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
   984
subsection {* Ordered Fields are Dense *}
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
   985
35028
108662d50512 more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents: 34146
diff changeset
   986
lemma less_half_sum: "a < b ==> a < (a+b) / (1+1::'a::linordered_field)"
23482
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
   987
by (simp add: field_simps zero_less_two)
14293
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
   988
35028
108662d50512 more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents: 34146
diff changeset
   989
lemma gt_half_sum: "a < b ==> (a+b)/(1+1::'a::linordered_field) < b"
23482
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
   990
by (simp add: field_simps zero_less_two)
14293
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
   991
35028
108662d50512 more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents: 34146
diff changeset
   992
instance linordered_field < dense_linorder
24422
c0b5ff9e9e4d moved class dense_linear_order to Orderings.thy
haftmann
parents: 24286
diff changeset
   993
proof
c0b5ff9e9e4d moved class dense_linear_order to Orderings.thy
haftmann
parents: 24286
diff changeset
   994
  fix x y :: 'a
c0b5ff9e9e4d moved class dense_linear_order to Orderings.thy
haftmann
parents: 24286
diff changeset
   995
  have "x < x + 1" by simp
c0b5ff9e9e4d moved class dense_linear_order to Orderings.thy
haftmann
parents: 24286
diff changeset
   996
  then show "\<exists>y. x < y" .. 
c0b5ff9e9e4d moved class dense_linear_order to Orderings.thy
haftmann
parents: 24286
diff changeset
   997
  have "x - 1 < x" by simp
c0b5ff9e9e4d moved class dense_linear_order to Orderings.thy
haftmann
parents: 24286
diff changeset
   998
  then show "\<exists>y. y < x" ..
c0b5ff9e9e4d moved class dense_linear_order to Orderings.thy
haftmann
parents: 24286
diff changeset
   999
  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
  1000
qed
14293
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1001
15234
ec91a90c604e simplification tweaks for better arithmetic reasoning
paulson
parents: 15229
diff changeset
  1002
14293
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1003
subsection {* Absolute Value *}
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1004
14294
f4d806fd72ce absolute value theorems moved to HOL/Ring_and_Field
paulson
parents: 14293
diff changeset
  1005
lemma nonzero_abs_inverse:
35028
108662d50512 more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents: 34146
diff changeset
  1006
     "a \<noteq> 0 ==> abs (inverse (a::'a::linordered_field)) = inverse (abs a)"
14294
f4d806fd72ce absolute value theorems moved to HOL/Ring_and_Field
paulson
parents: 14293
diff changeset
  1007
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
  1008
                      negative_imp_inverse_negative)
f4d806fd72ce absolute value theorems moved to HOL/Ring_and_Field
paulson
parents: 14293
diff changeset
  1009
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
  1010
done
f4d806fd72ce absolute value theorems moved to HOL/Ring_and_Field
paulson
parents: 14293
diff changeset
  1011
f4d806fd72ce absolute value theorems moved to HOL/Ring_and_Field
paulson
parents: 14293
diff changeset
  1012
lemma abs_inverse [simp]:
35028
108662d50512 more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents: 34146
diff changeset
  1013
     "abs (inverse (a::'a::{linordered_field,division_by_zero})) = 
14294
f4d806fd72ce absolute value theorems moved to HOL/Ring_and_Field
paulson
parents: 14293
diff changeset
  1014
      inverse (abs a)"
21328
73bb86d0f483 dropped Inductive dependency
haftmann
parents: 21258
diff changeset
  1015
apply (cases "a=0", simp) 
14294
f4d806fd72ce absolute value theorems moved to HOL/Ring_and_Field
paulson
parents: 14293
diff changeset
  1016
apply (simp add: nonzero_abs_inverse) 
f4d806fd72ce absolute value theorems moved to HOL/Ring_and_Field
paulson
parents: 14293
diff changeset
  1017
done
f4d806fd72ce absolute value theorems moved to HOL/Ring_and_Field
paulson
parents: 14293
diff changeset
  1018
f4d806fd72ce absolute value theorems moved to HOL/Ring_and_Field
paulson
parents: 14293
diff changeset
  1019
lemma nonzero_abs_divide:
35028
108662d50512 more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents: 34146
diff changeset
  1020
     "b \<noteq> 0 ==> abs (a / (b::'a::linordered_field)) = abs a / abs b"
14294
f4d806fd72ce absolute value theorems moved to HOL/Ring_and_Field
paulson
parents: 14293
diff changeset
  1021
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
  1022
15234
ec91a90c604e simplification tweaks for better arithmetic reasoning
paulson
parents: 15229
diff changeset
  1023
lemma abs_divide [simp]:
35028
108662d50512 more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents: 34146
diff changeset
  1024
     "abs (a / (b::'a::{linordered_field,division_by_zero})) = abs a / abs b"
21328
73bb86d0f483 dropped Inductive dependency
haftmann
parents: 21258
diff changeset
  1025
apply (cases "b=0", simp) 
14294
f4d806fd72ce absolute value theorems moved to HOL/Ring_and_Field
paulson
parents: 14293
diff changeset
  1026
apply (simp add: nonzero_abs_divide) 
f4d806fd72ce absolute value theorems moved to HOL/Ring_and_Field
paulson
parents: 14293
diff changeset
  1027
done
f4d806fd72ce absolute value theorems moved to HOL/Ring_and_Field
paulson
parents: 14293
diff changeset
  1028
35028
108662d50512 more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents: 34146
diff changeset
  1029
lemma abs_div_pos: "(0::'a::{division_by_zero,linordered_field}) < y ==> 
25304
7491c00f0915 removed subclass edge ordered_ring < lordered_ring
haftmann
parents: 25267
diff changeset
  1030
    abs x / y = abs (x / y)"
7491c00f0915 removed subclass edge ordered_ring < lordered_ring
haftmann
parents: 25267
diff changeset
  1031
  apply (subst abs_divide)
7491c00f0915 removed subclass edge ordered_ring < lordered_ring
haftmann
parents: 25267
diff changeset
  1032
  apply (simp add: order_less_imp_le)
7491c00f0915 removed subclass edge ordered_ring < lordered_ring
haftmann
parents: 25267
diff changeset
  1033
done
16775
c1b87ef4a1c3 added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
avigad
parents: 16568
diff changeset
  1034
33364
2bd12592c5e8 tuned code setup
haftmann
parents: 33319
diff changeset
  1035
code_modulename SML
35050
9f841f20dca6 renamed OrderedGroup to Groups; split theory Ring_and_Field into Rings Fields
haftmann
parents: 35043
diff changeset
  1036
  Fields Arith
33364
2bd12592c5e8 tuned code setup
haftmann
parents: 33319
diff changeset
  1037
2bd12592c5e8 tuned code setup
haftmann
parents: 33319
diff changeset
  1038
code_modulename OCaml
35050
9f841f20dca6 renamed OrderedGroup to Groups; split theory Ring_and_Field into Rings Fields
haftmann
parents: 35043
diff changeset
  1039
  Fields Arith
33364
2bd12592c5e8 tuned code setup
haftmann
parents: 33319
diff changeset
  1040
2bd12592c5e8 tuned code setup
haftmann
parents: 33319
diff changeset
  1041
code_modulename Haskell
35050
9f841f20dca6 renamed OrderedGroup to Groups; split theory Ring_and_Field into Rings Fields
haftmann
parents: 35043
diff changeset
  1042
  Fields Arith
33364
2bd12592c5e8 tuned code setup
haftmann
parents: 33319
diff changeset
  1043
14265
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
  1044
end