src/HOL/Rings.thy
author paulson <lp15@cam.ac.uk>
Tue, 30 Jun 2015 13:56:16 +0100
changeset 60615 e5fa1d5d3952
parent 60570 7ed2cde6806d
child 60685 cb21b7022b00
permissions -rw-r--r--
Useful lemmas. The theorem concerning swapping the variables in a double integral.
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/Rings.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
58889
5b7a9633cfa8 modernized header uniformly as section;
wenzelm
parents: 58776
diff changeset
    10
section {* Rings *}
14265
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
    11
35050
9f841f20dca6 renamed OrderedGroup to Groups; split theory Ring_and_Field into Rings Fields
haftmann
parents: 35043
diff changeset
    12
theory Rings
9f841f20dca6 renamed OrderedGroup to Groups; split theory Ring_and_Field into Rings Fields
haftmann
parents: 35043
diff changeset
    13
imports Groups
15131
c69542757a4d New theory header syntax.
nipkow
parents: 15077
diff changeset
    14
begin
14504
7a3d80e276d4 new type class abelian_group
paulson
parents: 14475
diff changeset
    15
22390
378f34b1e380 now using "class"
haftmann
parents: 21328
diff changeset
    16
class semiring = ab_semigroup_add + semigroup_mult +
58776
95e58e04e534 use NO_MATCH-simproc for distribution rules in field_simps, otherwise field_simps on '(a / (c + d)) * (e + f)' can be non-terminating
hoelzl
parents: 58649
diff changeset
    17
  assumes distrib_right[algebra_simps]: "(a + b) * c = a * c + b * c"
95e58e04e534 use NO_MATCH-simproc for distribution rules in field_simps, otherwise field_simps on '(a / (c + d)) * (e + f)' can be non-terminating
hoelzl
parents: 58649
diff changeset
    18
  assumes distrib_left[algebra_simps]: "a * (b + c) = a * b + a * c"
25152
bfde2f8c0f63 partially localized
haftmann
parents: 25078
diff changeset
    19
begin
bfde2f8c0f63 partially localized
haftmann
parents: 25078
diff changeset
    20
bfde2f8c0f63 partially localized
haftmann
parents: 25078
diff changeset
    21
text{*For the @{text combine_numerals} simproc*}
bfde2f8c0f63 partially localized
haftmann
parents: 25078
diff changeset
    22
lemma combine_common_factor:
bfde2f8c0f63 partially localized
haftmann
parents: 25078
diff changeset
    23
  "a * e + (b * e + c) = (a + b) * e + c"
57514
bdc2c6b40bf2 prefer ac_simps collections over separate name bindings for add and mult
haftmann
parents: 57512
diff changeset
    24
by (simp add: distrib_right ac_simps)
25152
bfde2f8c0f63 partially localized
haftmann
parents: 25078
diff changeset
    25
bfde2f8c0f63 partially localized
haftmann
parents: 25078
diff changeset
    26
end
14504
7a3d80e276d4 new type class abelian_group
paulson
parents: 14475
diff changeset
    27
22390
378f34b1e380 now using "class"
haftmann
parents: 21328
diff changeset
    28
class mult_zero = times + zero +
25062
af5ef0d4d655 global class syntax
haftmann
parents: 24748
diff changeset
    29
  assumes mult_zero_left [simp]: "0 * a = 0"
af5ef0d4d655 global class syntax
haftmann
parents: 24748
diff changeset
    30
  assumes mult_zero_right [simp]: "a * 0 = 0"
58195
1fee63e0377d added various facts
haftmann
parents: 57514
diff changeset
    31
begin
1fee63e0377d added various facts
haftmann
parents: 57514
diff changeset
    32
1fee63e0377d added various facts
haftmann
parents: 57514
diff changeset
    33
lemma mult_not_zero:
1fee63e0377d added various facts
haftmann
parents: 57514
diff changeset
    34
  "a * b \<noteq> 0 \<Longrightarrow> a \<noteq> 0 \<and> b \<noteq> 0"
1fee63e0377d added various facts
haftmann
parents: 57514
diff changeset
    35
  by auto
1fee63e0377d added various facts
haftmann
parents: 57514
diff changeset
    36
1fee63e0377d added various facts
haftmann
parents: 57514
diff changeset
    37
end
21199
2d83f93c3580 * Added annihilation axioms ("x * 0 = 0") to axclass semiring_0.
krauss
parents: 20633
diff changeset
    38
58198
099ca49b5231 generalized
haftmann
parents: 58195
diff changeset
    39
class semiring_0 = semiring + comm_monoid_add + mult_zero
099ca49b5231 generalized
haftmann
parents: 58195
diff changeset
    40
29904
856f16a3b436 add class cancel_comm_monoid_add
huffman
parents: 29833
diff changeset
    41
class semiring_0_cancel = semiring + cancel_comm_monoid_add
25186
f4d1ebffd025 localized further
haftmann
parents: 25152
diff changeset
    42
begin
14504
7a3d80e276d4 new type class abelian_group
paulson
parents: 14475
diff changeset
    43
25186
f4d1ebffd025 localized further
haftmann
parents: 25152
diff changeset
    44
subclass semiring_0
28823
dcbef866c9e2 tuned unfold_locales invocation
haftmann
parents: 28559
diff changeset
    45
proof
21199
2d83f93c3580 * Added annihilation axioms ("x * 0 = 0") to axclass semiring_0.
krauss
parents: 20633
diff changeset
    46
  fix a :: 'a
49962
a8cc904a6820 Renamed {left,right}_distrib to distrib_{right,left}.
webertj
parents: 44921
diff changeset
    47
  have "0 * a + 0 * a = 0 * a + 0" by (simp add: distrib_right [symmetric])
29667
53103fc8ffa3 Replaced group_ and ring_simps by algebra_simps;
nipkow
parents: 29465
diff changeset
    48
  thus "0 * a = 0" by (simp only: add_left_cancel)
25152
bfde2f8c0f63 partially localized
haftmann
parents: 25078
diff changeset
    49
next
bfde2f8c0f63 partially localized
haftmann
parents: 25078
diff changeset
    50
  fix a :: 'a
49962
a8cc904a6820 Renamed {left,right}_distrib to distrib_{right,left}.
webertj
parents: 44921
diff changeset
    51
  have "a * 0 + a * 0 = a * 0 + 0" by (simp add: distrib_left [symmetric])
29667
53103fc8ffa3 Replaced group_ and ring_simps by algebra_simps;
nipkow
parents: 29465
diff changeset
    52
  thus "a * 0 = 0" by (simp only: add_left_cancel)
21199
2d83f93c3580 * Added annihilation axioms ("x * 0 = 0") to axclass semiring_0.
krauss
parents: 20633
diff changeset
    53
qed
14940
b9ab8babd8b3 Further development of matrix theory
obua
parents: 14770
diff changeset
    54
25186
f4d1ebffd025 localized further
haftmann
parents: 25152
diff changeset
    55
end
25152
bfde2f8c0f63 partially localized
haftmann
parents: 25078
diff changeset
    56
22390
378f34b1e380 now using "class"
haftmann
parents: 21328
diff changeset
    57
class comm_semiring = ab_semigroup_add + ab_semigroup_mult +
25062
af5ef0d4d655 global class syntax
haftmann
parents: 24748
diff changeset
    58
  assumes distrib: "(a + b) * c = a * c + b * c"
25152
bfde2f8c0f63 partially localized
haftmann
parents: 25078
diff changeset
    59
begin
14504
7a3d80e276d4 new type class abelian_group
paulson
parents: 14475
diff changeset
    60
25152
bfde2f8c0f63 partially localized
haftmann
parents: 25078
diff changeset
    61
subclass semiring
28823
dcbef866c9e2 tuned unfold_locales invocation
haftmann
parents: 28559
diff changeset
    62
proof
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
    63
  fix a b c :: 'a
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
    64
  show "(a + b) * c = a * c + b * c" by (simp add: distrib)
57514
bdc2c6b40bf2 prefer ac_simps collections over separate name bindings for add and mult
haftmann
parents: 57512
diff changeset
    65
  have "a * (b + c) = (b + c) * a" by (simp add: ac_simps)
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
    66
  also have "... = b * a + c * a" by (simp only: distrib)
57514
bdc2c6b40bf2 prefer ac_simps collections over separate name bindings for add and mult
haftmann
parents: 57512
diff changeset
    67
  also have "... = a * b + a * c" by (simp add: ac_simps)
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
    68
  finally show "a * (b + c) = a * b + a * c" by blast
14504
7a3d80e276d4 new type class abelian_group
paulson
parents: 14475
diff changeset
    69
qed
7a3d80e276d4 new type class abelian_group
paulson
parents: 14475
diff changeset
    70
25152
bfde2f8c0f63 partially localized
haftmann
parents: 25078
diff changeset
    71
end
14504
7a3d80e276d4 new type class abelian_group
paulson
parents: 14475
diff changeset
    72
25152
bfde2f8c0f63 partially localized
haftmann
parents: 25078
diff changeset
    73
class comm_semiring_0 = comm_semiring + comm_monoid_add + mult_zero
bfde2f8c0f63 partially localized
haftmann
parents: 25078
diff changeset
    74
begin
bfde2f8c0f63 partially localized
haftmann
parents: 25078
diff changeset
    75
27516
9a5d4a8d4aac by intro_locales -> ..
huffman
parents: 26274
diff changeset
    76
subclass semiring_0 ..
25152
bfde2f8c0f63 partially localized
haftmann
parents: 25078
diff changeset
    77
bfde2f8c0f63 partially localized
haftmann
parents: 25078
diff changeset
    78
end
14504
7a3d80e276d4 new type class abelian_group
paulson
parents: 14475
diff changeset
    79
29904
856f16a3b436 add class cancel_comm_monoid_add
huffman
parents: 29833
diff changeset
    80
class comm_semiring_0_cancel = comm_semiring + cancel_comm_monoid_add
25186
f4d1ebffd025 localized further
haftmann
parents: 25152
diff changeset
    81
begin
14940
b9ab8babd8b3 Further development of matrix theory
obua
parents: 14770
diff changeset
    82
27516
9a5d4a8d4aac by intro_locales -> ..
huffman
parents: 26274
diff changeset
    83
subclass semiring_0_cancel ..
14940
b9ab8babd8b3 Further development of matrix theory
obua
parents: 14770
diff changeset
    84
28141
193c3ea0f63b instances comm_semiring_0_cancel < comm_semiring_0, comm_ring < comm_semiring_0_cancel
huffman
parents: 27651
diff changeset
    85
subclass comm_semiring_0 ..
193c3ea0f63b instances comm_semiring_0_cancel < comm_semiring_0, comm_ring < comm_semiring_0_cancel
huffman
parents: 27651
diff changeset
    86
25186
f4d1ebffd025 localized further
haftmann
parents: 25152
diff changeset
    87
end
21199
2d83f93c3580 * Added annihilation axioms ("x * 0 = 0") to axclass semiring_0.
krauss
parents: 20633
diff changeset
    88
22390
378f34b1e380 now using "class"
haftmann
parents: 21328
diff changeset
    89
class zero_neq_one = zero + one +
25062
af5ef0d4d655 global class syntax
haftmann
parents: 24748
diff changeset
    90
  assumes zero_neq_one [simp]: "0 \<noteq> 1"
26193
37a7eb7fd5f7 continued localization
haftmann
parents: 25917
diff changeset
    91
begin
37a7eb7fd5f7 continued localization
haftmann
parents: 25917
diff changeset
    92
37a7eb7fd5f7 continued localization
haftmann
parents: 25917
diff changeset
    93
lemma one_neq_zero [simp]: "1 \<noteq> 0"
29667
53103fc8ffa3 Replaced group_ and ring_simps by algebra_simps;
nipkow
parents: 29465
diff changeset
    94
by (rule not_sym) (rule zero_neq_one)
26193
37a7eb7fd5f7 continued localization
haftmann
parents: 25917
diff changeset
    95
54225
8a49a5a30284 generalized of_bool conversion
haftmann
parents: 54147
diff changeset
    96
definition of_bool :: "bool \<Rightarrow> 'a"
8a49a5a30284 generalized of_bool conversion
haftmann
parents: 54147
diff changeset
    97
where
60562
24af00b010cf Amalgamation of the class comm_semiring_1_diff_distrib into comm_semiring_1_cancel. Moving axiom le_add_diff_inverse2 from semiring_numeral_div to linordered_semidom.
paulson <lp15@cam.ac.uk>
parents: 60529
diff changeset
    98
  "of_bool p = (if p then 1 else 0)"
54225
8a49a5a30284 generalized of_bool conversion
haftmann
parents: 54147
diff changeset
    99
8a49a5a30284 generalized of_bool conversion
haftmann
parents: 54147
diff changeset
   100
lemma of_bool_eq [simp, code]:
8a49a5a30284 generalized of_bool conversion
haftmann
parents: 54147
diff changeset
   101
  "of_bool False = 0"
8a49a5a30284 generalized of_bool conversion
haftmann
parents: 54147
diff changeset
   102
  "of_bool True = 1"
8a49a5a30284 generalized of_bool conversion
haftmann
parents: 54147
diff changeset
   103
  by (simp_all add: of_bool_def)
8a49a5a30284 generalized of_bool conversion
haftmann
parents: 54147
diff changeset
   104
8a49a5a30284 generalized of_bool conversion
haftmann
parents: 54147
diff changeset
   105
lemma of_bool_eq_iff:
8a49a5a30284 generalized of_bool conversion
haftmann
parents: 54147
diff changeset
   106
  "of_bool p = of_bool q \<longleftrightarrow> p = q"
8a49a5a30284 generalized of_bool conversion
haftmann
parents: 54147
diff changeset
   107
  by (simp add: of_bool_def)
8a49a5a30284 generalized of_bool conversion
haftmann
parents: 54147
diff changeset
   108
55187
6d0d93316daf split rules for of_bool, similar to if
haftmann
parents: 54703
diff changeset
   109
lemma split_of_bool [split]:
6d0d93316daf split rules for of_bool, similar to if
haftmann
parents: 54703
diff changeset
   110
  "P (of_bool p) \<longleftrightarrow> (p \<longrightarrow> P 1) \<and> (\<not> p \<longrightarrow> P 0)"
6d0d93316daf split rules for of_bool, similar to if
haftmann
parents: 54703
diff changeset
   111
  by (cases p) simp_all
6d0d93316daf split rules for of_bool, similar to if
haftmann
parents: 54703
diff changeset
   112
6d0d93316daf split rules for of_bool, similar to if
haftmann
parents: 54703
diff changeset
   113
lemma split_of_bool_asm:
6d0d93316daf split rules for of_bool, similar to if
haftmann
parents: 54703
diff changeset
   114
  "P (of_bool p) \<longleftrightarrow> \<not> (p \<and> \<not> P 1 \<or> \<not> p \<and> \<not> P 0)"
6d0d93316daf split rules for of_bool, similar to if
haftmann
parents: 54703
diff changeset
   115
  by (cases p) simp_all
60562
24af00b010cf Amalgamation of the class comm_semiring_1_diff_distrib into comm_semiring_1_cancel. Moving axiom le_add_diff_inverse2 from semiring_numeral_div to linordered_semidom.
paulson <lp15@cam.ac.uk>
parents: 60529
diff changeset
   116
24af00b010cf Amalgamation of the class comm_semiring_1_diff_distrib into comm_semiring_1_cancel. Moving axiom le_add_diff_inverse2 from semiring_numeral_div to linordered_semidom.
paulson <lp15@cam.ac.uk>
parents: 60529
diff changeset
   117
end
14265
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   118
22390
378f34b1e380 now using "class"
haftmann
parents: 21328
diff changeset
   119
class semiring_1 = zero_neq_one + semiring_0 + monoid_mult
14504
7a3d80e276d4 new type class abelian_group
paulson
parents: 14475
diff changeset
   120
27651
16a26996c30e moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
haftmann
parents: 27516
diff changeset
   121
text {* Abstract divisibility *}
16a26996c30e moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
haftmann
parents: 27516
diff changeset
   122
16a26996c30e moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
haftmann
parents: 27516
diff changeset
   123
class dvd = times
16a26996c30e moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
haftmann
parents: 27516
diff changeset
   124
begin
16a26996c30e moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
haftmann
parents: 27516
diff changeset
   125
50420
f1a27e82af16 corrected nonsensical associativity of `` and dvd
nipkow
parents: 49962
diff changeset
   126
definition dvd :: "'a \<Rightarrow> 'a \<Rightarrow> bool" (infix "dvd" 50) where
37767
a2b7a20d6ea3 dropped superfluous [code del]s
haftmann
parents: 36977
diff changeset
   127
  "b dvd a \<longleftrightarrow> (\<exists>k. a = b * k)"
27651
16a26996c30e moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
haftmann
parents: 27516
diff changeset
   128
16a26996c30e moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
haftmann
parents: 27516
diff changeset
   129
lemma dvdI [intro?]: "a = b * k \<Longrightarrow> b dvd a"
16a26996c30e moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
haftmann
parents: 27516
diff changeset
   130
  unfolding dvd_def ..
16a26996c30e moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
haftmann
parents: 27516
diff changeset
   131
16a26996c30e moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
haftmann
parents: 27516
diff changeset
   132
lemma dvdE [elim?]: "b dvd a \<Longrightarrow> (\<And>k. a = b * k \<Longrightarrow> P) \<Longrightarrow> P"
60562
24af00b010cf Amalgamation of the class comm_semiring_1_diff_distrib into comm_semiring_1_cancel. Moving axiom le_add_diff_inverse2 from semiring_numeral_div to linordered_semidom.
paulson <lp15@cam.ac.uk>
parents: 60529
diff changeset
   133
  unfolding dvd_def by blast
27651
16a26996c30e moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
haftmann
parents: 27516
diff changeset
   134
16a26996c30e moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
haftmann
parents: 27516
diff changeset
   135
end
16a26996c30e moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
haftmann
parents: 27516
diff changeset
   136
59009
348561aa3869 generalized lemmas (particularly concerning dvd) as far as appropriate
haftmann
parents: 59000
diff changeset
   137
context comm_monoid_mult
25152
bfde2f8c0f63 partially localized
haftmann
parents: 25078
diff changeset
   138
begin
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   139
59009
348561aa3869 generalized lemmas (particularly concerning dvd) as far as appropriate
haftmann
parents: 59000
diff changeset
   140
subclass dvd .
25152
bfde2f8c0f63 partially localized
haftmann
parents: 25078
diff changeset
   141
59009
348561aa3869 generalized lemmas (particularly concerning dvd) as far as appropriate
haftmann
parents: 59000
diff changeset
   142
lemma dvd_refl [simp]:
348561aa3869 generalized lemmas (particularly concerning dvd) as far as appropriate
haftmann
parents: 59000
diff changeset
   143
  "a dvd a"
28559
55c003a5600a tuned default rules of (dvd)
haftmann
parents: 28141
diff changeset
   144
proof
55c003a5600a tuned default rules of (dvd)
haftmann
parents: 28141
diff changeset
   145
  show "a = a * 1" by simp
27651
16a26996c30e moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
haftmann
parents: 27516
diff changeset
   146
qed
16a26996c30e moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
haftmann
parents: 27516
diff changeset
   147
16a26996c30e moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
haftmann
parents: 27516
diff changeset
   148
lemma dvd_trans:
16a26996c30e moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
haftmann
parents: 27516
diff changeset
   149
  assumes "a dvd b" and "b dvd c"
16a26996c30e moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
haftmann
parents: 27516
diff changeset
   150
  shows "a dvd c"
16a26996c30e moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
haftmann
parents: 27516
diff changeset
   151
proof -
28559
55c003a5600a tuned default rules of (dvd)
haftmann
parents: 28141
diff changeset
   152
  from assms obtain v where "b = a * v" by (auto elim!: dvdE)
55c003a5600a tuned default rules of (dvd)
haftmann
parents: 28141
diff changeset
   153
  moreover from assms obtain w where "c = b * w" by (auto elim!: dvdE)
57512
cc97b347b301 reduced name variants for assoc and commute on plus and mult
haftmann
parents: 56544
diff changeset
   154
  ultimately have "c = a * (v * w)" by (simp add: mult.assoc)
28559
55c003a5600a tuned default rules of (dvd)
haftmann
parents: 28141
diff changeset
   155
  then show ?thesis ..
27651
16a26996c30e moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
haftmann
parents: 27516
diff changeset
   156
qed
16a26996c30e moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
haftmann
parents: 27516
diff changeset
   157
59009
348561aa3869 generalized lemmas (particularly concerning dvd) as far as appropriate
haftmann
parents: 59000
diff changeset
   158
lemma one_dvd [simp]:
348561aa3869 generalized lemmas (particularly concerning dvd) as far as appropriate
haftmann
parents: 59000
diff changeset
   159
  "1 dvd a"
348561aa3869 generalized lemmas (particularly concerning dvd) as far as appropriate
haftmann
parents: 59000
diff changeset
   160
  by (auto intro!: dvdI)
28559
55c003a5600a tuned default rules of (dvd)
haftmann
parents: 28141
diff changeset
   161
59009
348561aa3869 generalized lemmas (particularly concerning dvd) as far as appropriate
haftmann
parents: 59000
diff changeset
   162
lemma dvd_mult [simp]:
348561aa3869 generalized lemmas (particularly concerning dvd) as far as appropriate
haftmann
parents: 59000
diff changeset
   163
  "a dvd c \<Longrightarrow> a dvd (b * c)"
348561aa3869 generalized lemmas (particularly concerning dvd) as far as appropriate
haftmann
parents: 59000
diff changeset
   164
  by (auto intro!: mult.left_commute dvdI elim!: dvdE)
27651
16a26996c30e moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
haftmann
parents: 27516
diff changeset
   165
59009
348561aa3869 generalized lemmas (particularly concerning dvd) as far as appropriate
haftmann
parents: 59000
diff changeset
   166
lemma dvd_mult2 [simp]:
348561aa3869 generalized lemmas (particularly concerning dvd) as far as appropriate
haftmann
parents: 59000
diff changeset
   167
  "a dvd b \<Longrightarrow> a dvd (b * c)"
60562
24af00b010cf Amalgamation of the class comm_semiring_1_diff_distrib into comm_semiring_1_cancel. Moving axiom le_add_diff_inverse2 from semiring_numeral_div to linordered_semidom.
paulson <lp15@cam.ac.uk>
parents: 60529
diff changeset
   168
  using dvd_mult [of a b c] by (simp add: ac_simps)
27651
16a26996c30e moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
haftmann
parents: 27516
diff changeset
   169
59009
348561aa3869 generalized lemmas (particularly concerning dvd) as far as appropriate
haftmann
parents: 59000
diff changeset
   170
lemma dvd_triv_right [simp]:
348561aa3869 generalized lemmas (particularly concerning dvd) as far as appropriate
haftmann
parents: 59000
diff changeset
   171
  "a dvd b * a"
348561aa3869 generalized lemmas (particularly concerning dvd) as far as appropriate
haftmann
parents: 59000
diff changeset
   172
  by (rule dvd_mult) (rule dvd_refl)
27651
16a26996c30e moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
haftmann
parents: 27516
diff changeset
   173
59009
348561aa3869 generalized lemmas (particularly concerning dvd) as far as appropriate
haftmann
parents: 59000
diff changeset
   174
lemma dvd_triv_left [simp]:
348561aa3869 generalized lemmas (particularly concerning dvd) as far as appropriate
haftmann
parents: 59000
diff changeset
   175
  "a dvd a * b"
348561aa3869 generalized lemmas (particularly concerning dvd) as far as appropriate
haftmann
parents: 59000
diff changeset
   176
  by (rule dvd_mult2) (rule dvd_refl)
27651
16a26996c30e moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
haftmann
parents: 27516
diff changeset
   177
16a26996c30e moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
haftmann
parents: 27516
diff changeset
   178
lemma mult_dvd_mono:
30042
31039ee583fa Removed subsumed lemmas
nipkow
parents: 29981
diff changeset
   179
  assumes "a dvd b"
31039ee583fa Removed subsumed lemmas
nipkow
parents: 29981
diff changeset
   180
    and "c dvd d"
27651
16a26996c30e moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
haftmann
parents: 27516
diff changeset
   181
  shows "a * c dvd b * d"
16a26996c30e moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
haftmann
parents: 27516
diff changeset
   182
proof -
30042
31039ee583fa Removed subsumed lemmas
nipkow
parents: 29981
diff changeset
   183
  from `a dvd b` obtain b' where "b = a * b'" ..
31039ee583fa Removed subsumed lemmas
nipkow
parents: 29981
diff changeset
   184
  moreover from `c dvd d` obtain d' where "d = c * d'" ..
57514
bdc2c6b40bf2 prefer ac_simps collections over separate name bindings for add and mult
haftmann
parents: 57512
diff changeset
   185
  ultimately have "b * d = (a * c) * (b' * d')" by (simp add: ac_simps)
27651
16a26996c30e moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
haftmann
parents: 27516
diff changeset
   186
  then show ?thesis ..
16a26996c30e moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
haftmann
parents: 27516
diff changeset
   187
qed
16a26996c30e moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
haftmann
parents: 27516
diff changeset
   188
59009
348561aa3869 generalized lemmas (particularly concerning dvd) as far as appropriate
haftmann
parents: 59000
diff changeset
   189
lemma dvd_mult_left:
348561aa3869 generalized lemmas (particularly concerning dvd) as far as appropriate
haftmann
parents: 59000
diff changeset
   190
  "a * b dvd c \<Longrightarrow> a dvd c"
348561aa3869 generalized lemmas (particularly concerning dvd) as far as appropriate
haftmann
parents: 59000
diff changeset
   191
  by (simp add: dvd_def mult.assoc) blast
27651
16a26996c30e moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
haftmann
parents: 27516
diff changeset
   192
59009
348561aa3869 generalized lemmas (particularly concerning dvd) as far as appropriate
haftmann
parents: 59000
diff changeset
   193
lemma dvd_mult_right:
348561aa3869 generalized lemmas (particularly concerning dvd) as far as appropriate
haftmann
parents: 59000
diff changeset
   194
  "a * b dvd c \<Longrightarrow> b dvd c"
348561aa3869 generalized lemmas (particularly concerning dvd) as far as appropriate
haftmann
parents: 59000
diff changeset
   195
  using dvd_mult_left [of b a c] by (simp add: ac_simps)
60562
24af00b010cf Amalgamation of the class comm_semiring_1_diff_distrib into comm_semiring_1_cancel. Moving axiom le_add_diff_inverse2 from semiring_numeral_div to linordered_semidom.
paulson <lp15@cam.ac.uk>
parents: 60529
diff changeset
   196
59009
348561aa3869 generalized lemmas (particularly concerning dvd) as far as appropriate
haftmann
parents: 59000
diff changeset
   197
end
348561aa3869 generalized lemmas (particularly concerning dvd) as far as appropriate
haftmann
parents: 59000
diff changeset
   198
348561aa3869 generalized lemmas (particularly concerning dvd) as far as appropriate
haftmann
parents: 59000
diff changeset
   199
class comm_semiring_1 = zero_neq_one + comm_semiring_0 + comm_monoid_mult
348561aa3869 generalized lemmas (particularly concerning dvd) as far as appropriate
haftmann
parents: 59000
diff changeset
   200
begin
348561aa3869 generalized lemmas (particularly concerning dvd) as far as appropriate
haftmann
parents: 59000
diff changeset
   201
348561aa3869 generalized lemmas (particularly concerning dvd) as far as appropriate
haftmann
parents: 59000
diff changeset
   202
subclass semiring_1 ..
27651
16a26996c30e moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
haftmann
parents: 27516
diff changeset
   203
59009
348561aa3869 generalized lemmas (particularly concerning dvd) as far as appropriate
haftmann
parents: 59000
diff changeset
   204
lemma dvd_0_left_iff [simp]:
348561aa3869 generalized lemmas (particularly concerning dvd) as far as appropriate
haftmann
parents: 59000
diff changeset
   205
  "0 dvd a \<longleftrightarrow> a = 0"
348561aa3869 generalized lemmas (particularly concerning dvd) as far as appropriate
haftmann
parents: 59000
diff changeset
   206
  by (auto intro: dvd_refl elim!: dvdE)
27651
16a26996c30e moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
haftmann
parents: 27516
diff changeset
   207
59009
348561aa3869 generalized lemmas (particularly concerning dvd) as far as appropriate
haftmann
parents: 59000
diff changeset
   208
lemma dvd_0_right [iff]:
348561aa3869 generalized lemmas (particularly concerning dvd) as far as appropriate
haftmann
parents: 59000
diff changeset
   209
  "a dvd 0"
348561aa3869 generalized lemmas (particularly concerning dvd) as far as appropriate
haftmann
parents: 59000
diff changeset
   210
proof
348561aa3869 generalized lemmas (particularly concerning dvd) as far as appropriate
haftmann
parents: 59000
diff changeset
   211
  show "0 = a * 0" by simp
348561aa3869 generalized lemmas (particularly concerning dvd) as far as appropriate
haftmann
parents: 59000
diff changeset
   212
qed
348561aa3869 generalized lemmas (particularly concerning dvd) as far as appropriate
haftmann
parents: 59000
diff changeset
   213
348561aa3869 generalized lemmas (particularly concerning dvd) as far as appropriate
haftmann
parents: 59000
diff changeset
   214
lemma dvd_0_left:
348561aa3869 generalized lemmas (particularly concerning dvd) as far as appropriate
haftmann
parents: 59000
diff changeset
   215
  "0 dvd a \<Longrightarrow> a = 0"
348561aa3869 generalized lemmas (particularly concerning dvd) as far as appropriate
haftmann
parents: 59000
diff changeset
   216
  by simp
348561aa3869 generalized lemmas (particularly concerning dvd) as far as appropriate
haftmann
parents: 59000
diff changeset
   217
348561aa3869 generalized lemmas (particularly concerning dvd) as far as appropriate
haftmann
parents: 59000
diff changeset
   218
lemma dvd_add [simp]:
348561aa3869 generalized lemmas (particularly concerning dvd) as far as appropriate
haftmann
parents: 59000
diff changeset
   219
  assumes "a dvd b" and "a dvd c"
348561aa3869 generalized lemmas (particularly concerning dvd) as far as appropriate
haftmann
parents: 59000
diff changeset
   220
  shows "a dvd (b + c)"
27651
16a26996c30e moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
haftmann
parents: 27516
diff changeset
   221
proof -
29925
17d1e32ef867 dvd and setprod lemmas
nipkow
parents: 29915
diff changeset
   222
  from `a dvd b` obtain b' where "b = a * b'" ..
17d1e32ef867 dvd and setprod lemmas
nipkow
parents: 29915
diff changeset
   223
  moreover from `a dvd c` obtain c' where "c = a * c'" ..
49962
a8cc904a6820 Renamed {left,right}_distrib to distrib_{right,left}.
webertj
parents: 44921
diff changeset
   224
  ultimately have "b + c = a * (b' + c')" by (simp add: distrib_left)
27651
16a26996c30e moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
haftmann
parents: 27516
diff changeset
   225
  then show ?thesis ..
16a26996c30e moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
haftmann
parents: 27516
diff changeset
   226
qed
16a26996c30e moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
haftmann
parents: 27516
diff changeset
   227
25152
bfde2f8c0f63 partially localized
haftmann
parents: 25078
diff changeset
   228
end
14421
ee97b6463cb4 new Ring_and_Field hierarchy, eliminating redundant axioms
paulson
parents: 14398
diff changeset
   229
29904
856f16a3b436 add class cancel_comm_monoid_add
huffman
parents: 29833
diff changeset
   230
class semiring_1_cancel = semiring + cancel_comm_monoid_add
856f16a3b436 add class cancel_comm_monoid_add
huffman
parents: 29833
diff changeset
   231
  + zero_neq_one + monoid_mult
25267
1f745c599b5c proper reinitialisation after subclass
haftmann
parents: 25238
diff changeset
   232
begin
14940
b9ab8babd8b3 Further development of matrix theory
obua
parents: 14770
diff changeset
   233
27516
9a5d4a8d4aac by intro_locales -> ..
huffman
parents: 26274
diff changeset
   234
subclass semiring_0_cancel ..
25512
4134f7c782e2 using intro_locales instead of unfold_locales if appropriate
haftmann
parents: 25450
diff changeset
   235
27516
9a5d4a8d4aac by intro_locales -> ..
huffman
parents: 26274
diff changeset
   236
subclass semiring_1 ..
25267
1f745c599b5c proper reinitialisation after subclass
haftmann
parents: 25238
diff changeset
   237
1f745c599b5c proper reinitialisation after subclass
haftmann
parents: 25238
diff changeset
   238
end
21199
2d83f93c3580 * Added annihilation axioms ("x * 0 = 0") to axclass semiring_0.
krauss
parents: 20633
diff changeset
   239
60562
24af00b010cf Amalgamation of the class comm_semiring_1_diff_distrib into comm_semiring_1_cancel. Moving axiom le_add_diff_inverse2 from semiring_numeral_div to linordered_semidom.
paulson <lp15@cam.ac.uk>
parents: 60529
diff changeset
   240
class comm_semiring_1_cancel = comm_semiring + cancel_comm_monoid_add +
24af00b010cf Amalgamation of the class comm_semiring_1_diff_distrib into comm_semiring_1_cancel. Moving axiom le_add_diff_inverse2 from semiring_numeral_div to linordered_semidom.
paulson <lp15@cam.ac.uk>
parents: 60529
diff changeset
   241
                               zero_neq_one + comm_monoid_mult +
24af00b010cf Amalgamation of the class comm_semiring_1_diff_distrib into comm_semiring_1_cancel. Moving axiom le_add_diff_inverse2 from semiring_numeral_div to linordered_semidom.
paulson <lp15@cam.ac.uk>
parents: 60529
diff changeset
   242
  assumes right_diff_distrib' [algebra_simps]: "a * (b - c) = a * b - a * c"
25267
1f745c599b5c proper reinitialisation after subclass
haftmann
parents: 25238
diff changeset
   243
begin
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   244
27516
9a5d4a8d4aac by intro_locales -> ..
huffman
parents: 26274
diff changeset
   245
subclass semiring_1_cancel ..
9a5d4a8d4aac by intro_locales -> ..
huffman
parents: 26274
diff changeset
   246
subclass comm_semiring_0_cancel ..
9a5d4a8d4aac by intro_locales -> ..
huffman
parents: 26274
diff changeset
   247
subclass comm_semiring_1 ..
25267
1f745c599b5c proper reinitialisation after subclass
haftmann
parents: 25238
diff changeset
   248
59816
034b13f4efae distributivity of partial minus establishes desired properties of dvd in semirings
haftmann
parents: 59557
diff changeset
   249
lemma left_diff_distrib' [algebra_simps]:
034b13f4efae distributivity of partial minus establishes desired properties of dvd in semirings
haftmann
parents: 59557
diff changeset
   250
  "(b - c) * a = b * a - c * a"
034b13f4efae distributivity of partial minus establishes desired properties of dvd in semirings
haftmann
parents: 59557
diff changeset
   251
  by (simp add: algebra_simps)
034b13f4efae distributivity of partial minus establishes desired properties of dvd in semirings
haftmann
parents: 59557
diff changeset
   252
034b13f4efae distributivity of partial minus establishes desired properties of dvd in semirings
haftmann
parents: 59557
diff changeset
   253
lemma dvd_add_times_triv_left_iff [simp]:
034b13f4efae distributivity of partial minus establishes desired properties of dvd in semirings
haftmann
parents: 59557
diff changeset
   254
  "a dvd c * a + b \<longleftrightarrow> a dvd b"
034b13f4efae distributivity of partial minus establishes desired properties of dvd in semirings
haftmann
parents: 59557
diff changeset
   255
proof -
034b13f4efae distributivity of partial minus establishes desired properties of dvd in semirings
haftmann
parents: 59557
diff changeset
   256
  have "a dvd a * c + b \<longleftrightarrow> a dvd b" (is "?P \<longleftrightarrow> ?Q")
034b13f4efae distributivity of partial minus establishes desired properties of dvd in semirings
haftmann
parents: 59557
diff changeset
   257
  proof
034b13f4efae distributivity of partial minus establishes desired properties of dvd in semirings
haftmann
parents: 59557
diff changeset
   258
    assume ?Q then show ?P by simp
034b13f4efae distributivity of partial minus establishes desired properties of dvd in semirings
haftmann
parents: 59557
diff changeset
   259
  next
034b13f4efae distributivity of partial minus establishes desired properties of dvd in semirings
haftmann
parents: 59557
diff changeset
   260
    assume ?P
034b13f4efae distributivity of partial minus establishes desired properties of dvd in semirings
haftmann
parents: 59557
diff changeset
   261
    then obtain d where "a * c + b = a * d" ..
034b13f4efae distributivity of partial minus establishes desired properties of dvd in semirings
haftmann
parents: 59557
diff changeset
   262
    then have "a * c + b - a * c = a * d - a * c" by simp
034b13f4efae distributivity of partial minus establishes desired properties of dvd in semirings
haftmann
parents: 59557
diff changeset
   263
    then have "b = a * d - a * c" by simp
60562
24af00b010cf Amalgamation of the class comm_semiring_1_diff_distrib into comm_semiring_1_cancel. Moving axiom le_add_diff_inverse2 from semiring_numeral_div to linordered_semidom.
paulson <lp15@cam.ac.uk>
parents: 60529
diff changeset
   264
    then have "b = a * (d - c)" by (simp add: algebra_simps)
59816
034b13f4efae distributivity of partial minus establishes desired properties of dvd in semirings
haftmann
parents: 59557
diff changeset
   265
    then show ?Q ..
034b13f4efae distributivity of partial minus establishes desired properties of dvd in semirings
haftmann
parents: 59557
diff changeset
   266
  qed
034b13f4efae distributivity of partial minus establishes desired properties of dvd in semirings
haftmann
parents: 59557
diff changeset
   267
  then show "a dvd c * a + b \<longleftrightarrow> a dvd b" by (simp add: ac_simps)
034b13f4efae distributivity of partial minus establishes desired properties of dvd in semirings
haftmann
parents: 59557
diff changeset
   268
qed
034b13f4efae distributivity of partial minus establishes desired properties of dvd in semirings
haftmann
parents: 59557
diff changeset
   269
034b13f4efae distributivity of partial minus establishes desired properties of dvd in semirings
haftmann
parents: 59557
diff changeset
   270
lemma dvd_add_times_triv_right_iff [simp]:
034b13f4efae distributivity of partial minus establishes desired properties of dvd in semirings
haftmann
parents: 59557
diff changeset
   271
  "a dvd b + c * a \<longleftrightarrow> a dvd b"
034b13f4efae distributivity of partial minus establishes desired properties of dvd in semirings
haftmann
parents: 59557
diff changeset
   272
  using dvd_add_times_triv_left_iff [of a c b] by (simp add: ac_simps)
034b13f4efae distributivity of partial minus establishes desired properties of dvd in semirings
haftmann
parents: 59557
diff changeset
   273
034b13f4efae distributivity of partial minus establishes desired properties of dvd in semirings
haftmann
parents: 59557
diff changeset
   274
lemma dvd_add_triv_left_iff [simp]:
034b13f4efae distributivity of partial minus establishes desired properties of dvd in semirings
haftmann
parents: 59557
diff changeset
   275
  "a dvd a + b \<longleftrightarrow> a dvd b"
034b13f4efae distributivity of partial minus establishes desired properties of dvd in semirings
haftmann
parents: 59557
diff changeset
   276
  using dvd_add_times_triv_left_iff [of a 1 b] by simp
034b13f4efae distributivity of partial minus establishes desired properties of dvd in semirings
haftmann
parents: 59557
diff changeset
   277
034b13f4efae distributivity of partial minus establishes desired properties of dvd in semirings
haftmann
parents: 59557
diff changeset
   278
lemma dvd_add_triv_right_iff [simp]:
034b13f4efae distributivity of partial minus establishes desired properties of dvd in semirings
haftmann
parents: 59557
diff changeset
   279
  "a dvd b + a \<longleftrightarrow> a dvd b"
034b13f4efae distributivity of partial minus establishes desired properties of dvd in semirings
haftmann
parents: 59557
diff changeset
   280
  using dvd_add_times_triv_right_iff [of a b 1] by simp
034b13f4efae distributivity of partial minus establishes desired properties of dvd in semirings
haftmann
parents: 59557
diff changeset
   281
034b13f4efae distributivity of partial minus establishes desired properties of dvd in semirings
haftmann
parents: 59557
diff changeset
   282
lemma dvd_add_right_iff:
034b13f4efae distributivity of partial minus establishes desired properties of dvd in semirings
haftmann
parents: 59557
diff changeset
   283
  assumes "a dvd b"
034b13f4efae distributivity of partial minus establishes desired properties of dvd in semirings
haftmann
parents: 59557
diff changeset
   284
  shows "a dvd b + c \<longleftrightarrow> a dvd c" (is "?P \<longleftrightarrow> ?Q")
034b13f4efae distributivity of partial minus establishes desired properties of dvd in semirings
haftmann
parents: 59557
diff changeset
   285
proof
034b13f4efae distributivity of partial minus establishes desired properties of dvd in semirings
haftmann
parents: 59557
diff changeset
   286
  assume ?P then obtain d where "b + c = a * d" ..
034b13f4efae distributivity of partial minus establishes desired properties of dvd in semirings
haftmann
parents: 59557
diff changeset
   287
  moreover from `a dvd b` obtain e where "b = a * e" ..
034b13f4efae distributivity of partial minus establishes desired properties of dvd in semirings
haftmann
parents: 59557
diff changeset
   288
  ultimately have "a * e + c = a * d" by simp
034b13f4efae distributivity of partial minus establishes desired properties of dvd in semirings
haftmann
parents: 59557
diff changeset
   289
  then have "a * e + c - a * e = a * d - a * e" by simp
034b13f4efae distributivity of partial minus establishes desired properties of dvd in semirings
haftmann
parents: 59557
diff changeset
   290
  then have "c = a * d - a * e" by simp
034b13f4efae distributivity of partial minus establishes desired properties of dvd in semirings
haftmann
parents: 59557
diff changeset
   291
  then have "c = a * (d - e)" by (simp add: algebra_simps)
034b13f4efae distributivity of partial minus establishes desired properties of dvd in semirings
haftmann
parents: 59557
diff changeset
   292
  then show ?Q ..
034b13f4efae distributivity of partial minus establishes desired properties of dvd in semirings
haftmann
parents: 59557
diff changeset
   293
next
034b13f4efae distributivity of partial minus establishes desired properties of dvd in semirings
haftmann
parents: 59557
diff changeset
   294
  assume ?Q with assms show ?P by simp
034b13f4efae distributivity of partial minus establishes desired properties of dvd in semirings
haftmann
parents: 59557
diff changeset
   295
qed
034b13f4efae distributivity of partial minus establishes desired properties of dvd in semirings
haftmann
parents: 59557
diff changeset
   296
034b13f4efae distributivity of partial minus establishes desired properties of dvd in semirings
haftmann
parents: 59557
diff changeset
   297
lemma dvd_add_left_iff:
034b13f4efae distributivity of partial minus establishes desired properties of dvd in semirings
haftmann
parents: 59557
diff changeset
   298
  assumes "a dvd c"
034b13f4efae distributivity of partial minus establishes desired properties of dvd in semirings
haftmann
parents: 59557
diff changeset
   299
  shows "a dvd b + c \<longleftrightarrow> a dvd b"
034b13f4efae distributivity of partial minus establishes desired properties of dvd in semirings
haftmann
parents: 59557
diff changeset
   300
  using assms dvd_add_right_iff [of a c b] by (simp add: ac_simps)
034b13f4efae distributivity of partial minus establishes desired properties of dvd in semirings
haftmann
parents: 59557
diff changeset
   301
034b13f4efae distributivity of partial minus establishes desired properties of dvd in semirings
haftmann
parents: 59557
diff changeset
   302
end
034b13f4efae distributivity of partial minus establishes desired properties of dvd in semirings
haftmann
parents: 59557
diff changeset
   303
22390
378f34b1e380 now using "class"
haftmann
parents: 21328
diff changeset
   304
class ring = semiring + ab_group_add
25267
1f745c599b5c proper reinitialisation after subclass
haftmann
parents: 25238
diff changeset
   305
begin
25152
bfde2f8c0f63 partially localized
haftmann
parents: 25078
diff changeset
   306
27516
9a5d4a8d4aac by intro_locales -> ..
huffman
parents: 26274
diff changeset
   307
subclass semiring_0_cancel ..
25152
bfde2f8c0f63 partially localized
haftmann
parents: 25078
diff changeset
   308
bfde2f8c0f63 partially localized
haftmann
parents: 25078
diff changeset
   309
text {* Distribution rules *}
bfde2f8c0f63 partially localized
haftmann
parents: 25078
diff changeset
   310
bfde2f8c0f63 partially localized
haftmann
parents: 25078
diff changeset
   311
lemma minus_mult_left: "- (a * b) = - a * b"
60562
24af00b010cf Amalgamation of the class comm_semiring_1_diff_distrib into comm_semiring_1_cancel. Moving axiom le_add_diff_inverse2 from semiring_numeral_div to linordered_semidom.
paulson <lp15@cam.ac.uk>
parents: 60529
diff changeset
   312
by (rule minus_unique) (simp add: distrib_right [symmetric])
25152
bfde2f8c0f63 partially localized
haftmann
parents: 25078
diff changeset
   313
bfde2f8c0f63 partially localized
haftmann
parents: 25078
diff changeset
   314
lemma minus_mult_right: "- (a * b) = a * - b"
60562
24af00b010cf Amalgamation of the class comm_semiring_1_diff_distrib into comm_semiring_1_cancel. Moving axiom le_add_diff_inverse2 from semiring_numeral_div to linordered_semidom.
paulson <lp15@cam.ac.uk>
parents: 60529
diff changeset
   315
by (rule minus_unique) (simp add: distrib_left [symmetric])
25152
bfde2f8c0f63 partially localized
haftmann
parents: 25078
diff changeset
   316
29407
5ef7e97fd9e4 move lemmas mult_minus{left,right} inside class ring
huffman
parents: 29406
diff changeset
   317
text{*Extract signs from products*}
54147
97a8ff4e4ac9 killed most "no_atp", to make Sledgehammer more complete
blanchet
parents: 52435
diff changeset
   318
lemmas mult_minus_left [simp] = minus_mult_left [symmetric]
97a8ff4e4ac9 killed most "no_atp", to make Sledgehammer more complete
blanchet
parents: 52435
diff changeset
   319
lemmas mult_minus_right [simp] = minus_mult_right [symmetric]
29407
5ef7e97fd9e4 move lemmas mult_minus{left,right} inside class ring
huffman
parents: 29406
diff changeset
   320
25152
bfde2f8c0f63 partially localized
haftmann
parents: 25078
diff changeset
   321
lemma minus_mult_minus [simp]: "- a * - b = a * b"
29667
53103fc8ffa3 Replaced group_ and ring_simps by algebra_simps;
nipkow
parents: 29465
diff changeset
   322
by simp
25152
bfde2f8c0f63 partially localized
haftmann
parents: 25078
diff changeset
   323
bfde2f8c0f63 partially localized
haftmann
parents: 25078
diff changeset
   324
lemma minus_mult_commute: "- a * b = a * - b"
29667
53103fc8ffa3 Replaced group_ and ring_simps by algebra_simps;
nipkow
parents: 29465
diff changeset
   325
by simp
53103fc8ffa3 Replaced group_ and ring_simps by algebra_simps;
nipkow
parents: 29465
diff changeset
   326
58776
95e58e04e534 use NO_MATCH-simproc for distribution rules in field_simps, otherwise field_simps on '(a / (c + d)) * (e + f)' can be non-terminating
hoelzl
parents: 58649
diff changeset
   327
lemma right_diff_distrib [algebra_simps]:
54230
b1d955791529 more simplification rules on unary and binary minus
haftmann
parents: 54225
diff changeset
   328
  "a * (b - c) = a * b - a * c"
b1d955791529 more simplification rules on unary and binary minus
haftmann
parents: 54225
diff changeset
   329
  using distrib_left [of a b "-c "] by simp
29667
53103fc8ffa3 Replaced group_ and ring_simps by algebra_simps;
nipkow
parents: 29465
diff changeset
   330
58776
95e58e04e534 use NO_MATCH-simproc for distribution rules in field_simps, otherwise field_simps on '(a / (c + d)) * (e + f)' can be non-terminating
hoelzl
parents: 58649
diff changeset
   331
lemma left_diff_distrib [algebra_simps]:
54230
b1d955791529 more simplification rules on unary and binary minus
haftmann
parents: 54225
diff changeset
   332
  "(a - b) * c = a * c - b * c"
b1d955791529 more simplification rules on unary and binary minus
haftmann
parents: 54225
diff changeset
   333
  using distrib_right [of a "- b" c] by simp
25152
bfde2f8c0f63 partially localized
haftmann
parents: 25078
diff changeset
   334
54147
97a8ff4e4ac9 killed most "no_atp", to make Sledgehammer more complete
blanchet
parents: 52435
diff changeset
   335
lemmas ring_distribs =
49962
a8cc904a6820 Renamed {left,right}_distrib to distrib_{right,left}.
webertj
parents: 44921
diff changeset
   336
  distrib_left distrib_right left_diff_distrib right_diff_distrib
25152
bfde2f8c0f63 partially localized
haftmann
parents: 25078
diff changeset
   337
25230
022029099a83 continued localization
haftmann
parents: 25193
diff changeset
   338
lemma eq_add_iff1:
022029099a83 continued localization
haftmann
parents: 25193
diff changeset
   339
  "a * e + c = b * e + d \<longleftrightarrow> (a - b) * e + c = d"
29667
53103fc8ffa3 Replaced group_ and ring_simps by algebra_simps;
nipkow
parents: 29465
diff changeset
   340
by (simp add: algebra_simps)
25230
022029099a83 continued localization
haftmann
parents: 25193
diff changeset
   341
022029099a83 continued localization
haftmann
parents: 25193
diff changeset
   342
lemma eq_add_iff2:
022029099a83 continued localization
haftmann
parents: 25193
diff changeset
   343
  "a * e + c = b * e + d \<longleftrightarrow> c = (b - a) * e + d"
29667
53103fc8ffa3 Replaced group_ and ring_simps by algebra_simps;
nipkow
parents: 29465
diff changeset
   344
by (simp add: algebra_simps)
25230
022029099a83 continued localization
haftmann
parents: 25193
diff changeset
   345
25152
bfde2f8c0f63 partially localized
haftmann
parents: 25078
diff changeset
   346
end
bfde2f8c0f63 partially localized
haftmann
parents: 25078
diff changeset
   347
54147
97a8ff4e4ac9 killed most "no_atp", to make Sledgehammer more complete
blanchet
parents: 52435
diff changeset
   348
lemmas ring_distribs =
49962
a8cc904a6820 Renamed {left,right}_distrib to distrib_{right,left}.
webertj
parents: 44921
diff changeset
   349
  distrib_left distrib_right left_diff_distrib right_diff_distrib
25152
bfde2f8c0f63 partially localized
haftmann
parents: 25078
diff changeset
   350
22390
378f34b1e380 now using "class"
haftmann
parents: 21328
diff changeset
   351
class comm_ring = comm_semiring + ab_group_add
25267
1f745c599b5c proper reinitialisation after subclass
haftmann
parents: 25238
diff changeset
   352
begin
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   353
27516
9a5d4a8d4aac by intro_locales -> ..
huffman
parents: 26274
diff changeset
   354
subclass ring ..
28141
193c3ea0f63b instances comm_semiring_0_cancel < comm_semiring_0, comm_ring < comm_semiring_0_cancel
huffman
parents: 27651
diff changeset
   355
subclass comm_semiring_0_cancel ..
25267
1f745c599b5c proper reinitialisation after subclass
haftmann
parents: 25238
diff changeset
   356
44350
63cddfbc5a09 replace lemma realpow_two_diff with new lemma square_diff_square_factored
huffman
parents: 44346
diff changeset
   357
lemma square_diff_square_factored:
63cddfbc5a09 replace lemma realpow_two_diff with new lemma square_diff_square_factored
huffman
parents: 44346
diff changeset
   358
  "x * x - y * y = (x + y) * (x - y)"
63cddfbc5a09 replace lemma realpow_two_diff with new lemma square_diff_square_factored
huffman
parents: 44346
diff changeset
   359
  by (simp add: algebra_simps)
63cddfbc5a09 replace lemma realpow_two_diff with new lemma square_diff_square_factored
huffman
parents: 44346
diff changeset
   360
25267
1f745c599b5c proper reinitialisation after subclass
haftmann
parents: 25238
diff changeset
   361
end
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   362
22390
378f34b1e380 now using "class"
haftmann
parents: 21328
diff changeset
   363
class ring_1 = ring + zero_neq_one + monoid_mult
25267
1f745c599b5c proper reinitialisation after subclass
haftmann
parents: 25238
diff changeset
   364
begin
14265
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
   365
27516
9a5d4a8d4aac by intro_locales -> ..
huffman
parents: 26274
diff changeset
   366
subclass semiring_1_cancel ..
25267
1f745c599b5c proper reinitialisation after subclass
haftmann
parents: 25238
diff changeset
   367
44346
00dd3c4dabe0 rename real_squared_diff_one_factored to square_diff_one_factored and move to Rings.thy
huffman
parents: 44064
diff changeset
   368
lemma square_diff_one_factored:
00dd3c4dabe0 rename real_squared_diff_one_factored to square_diff_one_factored and move to Rings.thy
huffman
parents: 44064
diff changeset
   369
  "x * x - 1 = (x + 1) * (x - 1)"
00dd3c4dabe0 rename real_squared_diff_one_factored to square_diff_one_factored and move to Rings.thy
huffman
parents: 44064
diff changeset
   370
  by (simp add: algebra_simps)
00dd3c4dabe0 rename real_squared_diff_one_factored to square_diff_one_factored and move to Rings.thy
huffman
parents: 44064
diff changeset
   371
25267
1f745c599b5c proper reinitialisation after subclass
haftmann
parents: 25238
diff changeset
   372
end
25152
bfde2f8c0f63 partially localized
haftmann
parents: 25078
diff changeset
   373
22390
378f34b1e380 now using "class"
haftmann
parents: 21328
diff changeset
   374
class comm_ring_1 = comm_ring + zero_neq_one + comm_monoid_mult
25267
1f745c599b5c proper reinitialisation after subclass
haftmann
parents: 25238
diff changeset
   375
begin
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
   376
27516
9a5d4a8d4aac by intro_locales -> ..
huffman
parents: 26274
diff changeset
   377
subclass ring_1 ..
60562
24af00b010cf Amalgamation of the class comm_semiring_1_diff_distrib into comm_semiring_1_cancel. Moving axiom le_add_diff_inverse2 from semiring_numeral_div to linordered_semidom.
paulson <lp15@cam.ac.uk>
parents: 60529
diff changeset
   378
subclass comm_semiring_1_cancel
59816
034b13f4efae distributivity of partial minus establishes desired properties of dvd in semirings
haftmann
parents: 59557
diff changeset
   379
  by unfold_locales (simp add: algebra_simps)
58647
fce800afeec7 more facts about abstract divisibility
haftmann
parents: 58198
diff changeset
   380
29465
b2cfb5d0a59e change dvd_minus_iff, minus_dvd_iff from [iff] to [simp] (due to problems with Library/Primes.thy)
huffman
parents: 29461
diff changeset
   381
lemma dvd_minus_iff [simp]: "x dvd - y \<longleftrightarrow> x dvd y"
29408
6d10cf26b5dc add lemmas dvd_minus_iff and minus_dvd_iff in class comm_ring_1
huffman
parents: 29407
diff changeset
   382
proof
6d10cf26b5dc add lemmas dvd_minus_iff and minus_dvd_iff in class comm_ring_1
huffman
parents: 29407
diff changeset
   383
  assume "x dvd - y"
6d10cf26b5dc add lemmas dvd_minus_iff and minus_dvd_iff in class comm_ring_1
huffman
parents: 29407
diff changeset
   384
  then have "x dvd - 1 * - y" by (rule dvd_mult)
6d10cf26b5dc add lemmas dvd_minus_iff and minus_dvd_iff in class comm_ring_1
huffman
parents: 29407
diff changeset
   385
  then show "x dvd y" by simp
6d10cf26b5dc add lemmas dvd_minus_iff and minus_dvd_iff in class comm_ring_1
huffman
parents: 29407
diff changeset
   386
next
6d10cf26b5dc add lemmas dvd_minus_iff and minus_dvd_iff in class comm_ring_1
huffman
parents: 29407
diff changeset
   387
  assume "x dvd y"
6d10cf26b5dc add lemmas dvd_minus_iff and minus_dvd_iff in class comm_ring_1
huffman
parents: 29407
diff changeset
   388
  then have "x dvd - 1 * y" by (rule dvd_mult)
6d10cf26b5dc add lemmas dvd_minus_iff and minus_dvd_iff in class comm_ring_1
huffman
parents: 29407
diff changeset
   389
  then show "x dvd - y" by simp
6d10cf26b5dc add lemmas dvd_minus_iff and minus_dvd_iff in class comm_ring_1
huffman
parents: 29407
diff changeset
   390
qed
6d10cf26b5dc add lemmas dvd_minus_iff and minus_dvd_iff in class comm_ring_1
huffman
parents: 29407
diff changeset
   391
29465
b2cfb5d0a59e change dvd_minus_iff, minus_dvd_iff from [iff] to [simp] (due to problems with Library/Primes.thy)
huffman
parents: 29461
diff changeset
   392
lemma minus_dvd_iff [simp]: "- x dvd y \<longleftrightarrow> x dvd y"
29408
6d10cf26b5dc add lemmas dvd_minus_iff and minus_dvd_iff in class comm_ring_1
huffman
parents: 29407
diff changeset
   393
proof
6d10cf26b5dc add lemmas dvd_minus_iff and minus_dvd_iff in class comm_ring_1
huffman
parents: 29407
diff changeset
   394
  assume "- x dvd y"
6d10cf26b5dc add lemmas dvd_minus_iff and minus_dvd_iff in class comm_ring_1
huffman
parents: 29407
diff changeset
   395
  then obtain k where "y = - x * k" ..
6d10cf26b5dc add lemmas dvd_minus_iff and minus_dvd_iff in class comm_ring_1
huffman
parents: 29407
diff changeset
   396
  then have "y = x * - k" by simp
6d10cf26b5dc add lemmas dvd_minus_iff and minus_dvd_iff in class comm_ring_1
huffman
parents: 29407
diff changeset
   397
  then show "x dvd y" ..
6d10cf26b5dc add lemmas dvd_minus_iff and minus_dvd_iff in class comm_ring_1
huffman
parents: 29407
diff changeset
   398
next
6d10cf26b5dc add lemmas dvd_minus_iff and minus_dvd_iff in class comm_ring_1
huffman
parents: 29407
diff changeset
   399
  assume "x dvd y"
6d10cf26b5dc add lemmas dvd_minus_iff and minus_dvd_iff in class comm_ring_1
huffman
parents: 29407
diff changeset
   400
  then obtain k where "y = x * k" ..
6d10cf26b5dc add lemmas dvd_minus_iff and minus_dvd_iff in class comm_ring_1
huffman
parents: 29407
diff changeset
   401
  then have "y = - x * - k" by simp
6d10cf26b5dc add lemmas dvd_minus_iff and minus_dvd_iff in class comm_ring_1
huffman
parents: 29407
diff changeset
   402
  then show "- x dvd y" ..
6d10cf26b5dc add lemmas dvd_minus_iff and minus_dvd_iff in class comm_ring_1
huffman
parents: 29407
diff changeset
   403
qed
6d10cf26b5dc add lemmas dvd_minus_iff and minus_dvd_iff in class comm_ring_1
huffman
parents: 29407
diff changeset
   404
54230
b1d955791529 more simplification rules on unary and binary minus
haftmann
parents: 54225
diff changeset
   405
lemma dvd_diff [simp]:
b1d955791529 more simplification rules on unary and binary minus
haftmann
parents: 54225
diff changeset
   406
  "x dvd y \<Longrightarrow> x dvd z \<Longrightarrow> x dvd (y - z)"
b1d955791529 more simplification rules on unary and binary minus
haftmann
parents: 54225
diff changeset
   407
  using dvd_add [of x y "- z"] by simp
29409
f0a8fe83bc07 add lemma dvd_diff to class comm_ring_1
huffman
parents: 29408
diff changeset
   408
25267
1f745c599b5c proper reinitialisation after subclass
haftmann
parents: 25238
diff changeset
   409
end
25152
bfde2f8c0f63 partially localized
haftmann
parents: 25078
diff changeset
   410
59833
ab828c2c5d67 clarified no_zero_devisors: makes only sense in a semiring;
haftmann
parents: 59832
diff changeset
   411
class semiring_no_zero_divisors = semiring_0 +
ab828c2c5d67 clarified no_zero_devisors: makes only sense in a semiring;
haftmann
parents: 59832
diff changeset
   412
  assumes no_zero_divisors: "a \<noteq> 0 \<Longrightarrow> b \<noteq> 0 \<Longrightarrow> a * b \<noteq> 0"
25230
022029099a83 continued localization
haftmann
parents: 25193
diff changeset
   413
begin
022029099a83 continued localization
haftmann
parents: 25193
diff changeset
   414
59833
ab828c2c5d67 clarified no_zero_devisors: makes only sense in a semiring;
haftmann
parents: 59832
diff changeset
   415
lemma divisors_zero:
ab828c2c5d67 clarified no_zero_devisors: makes only sense in a semiring;
haftmann
parents: 59832
diff changeset
   416
  assumes "a * b = 0"
ab828c2c5d67 clarified no_zero_devisors: makes only sense in a semiring;
haftmann
parents: 59832
diff changeset
   417
  shows "a = 0 \<or> b = 0"
ab828c2c5d67 clarified no_zero_devisors: makes only sense in a semiring;
haftmann
parents: 59832
diff changeset
   418
proof (rule classical)
ab828c2c5d67 clarified no_zero_devisors: makes only sense in a semiring;
haftmann
parents: 59832
diff changeset
   419
  assume "\<not> (a = 0 \<or> b = 0)"
ab828c2c5d67 clarified no_zero_devisors: makes only sense in a semiring;
haftmann
parents: 59832
diff changeset
   420
  then have "a \<noteq> 0" and "b \<noteq> 0" by auto
ab828c2c5d67 clarified no_zero_devisors: makes only sense in a semiring;
haftmann
parents: 59832
diff changeset
   421
  with no_zero_divisors have "a * b \<noteq> 0" by blast
ab828c2c5d67 clarified no_zero_devisors: makes only sense in a semiring;
haftmann
parents: 59832
diff changeset
   422
  with assms show ?thesis by simp
ab828c2c5d67 clarified no_zero_devisors: makes only sense in a semiring;
haftmann
parents: 59832
diff changeset
   423
qed
ab828c2c5d67 clarified no_zero_devisors: makes only sense in a semiring;
haftmann
parents: 59832
diff changeset
   424
25230
022029099a83 continued localization
haftmann
parents: 25193
diff changeset
   425
lemma mult_eq_0_iff [simp]:
58952
5d82cdef6c1b equivalence rules for structures without zero divisors
haftmann
parents: 58889
diff changeset
   426
  shows "a * b = 0 \<longleftrightarrow> a = 0 \<or> b = 0"
25230
022029099a83 continued localization
haftmann
parents: 25193
diff changeset
   427
proof (cases "a = 0 \<or> b = 0")
022029099a83 continued localization
haftmann
parents: 25193
diff changeset
   428
  case False then have "a \<noteq> 0" and "b \<noteq> 0" by auto
022029099a83 continued localization
haftmann
parents: 25193
diff changeset
   429
    then show ?thesis using no_zero_divisors by simp
022029099a83 continued localization
haftmann
parents: 25193
diff changeset
   430
next
022029099a83 continued localization
haftmann
parents: 25193
diff changeset
   431
  case True then show ?thesis by auto
022029099a83 continued localization
haftmann
parents: 25193
diff changeset
   432
qed
022029099a83 continued localization
haftmann
parents: 25193
diff changeset
   433
58952
5d82cdef6c1b equivalence rules for structures without zero divisors
haftmann
parents: 58889
diff changeset
   434
end
5d82cdef6c1b equivalence rules for structures without zero divisors
haftmann
parents: 58889
diff changeset
   435
60516
0826b7025d07 generalized some theorems about integral domains and moved to HOL theories
haftmann
parents: 60429
diff changeset
   436
class semiring_no_zero_divisors_cancel = semiring_no_zero_divisors +
0826b7025d07 generalized some theorems about integral domains and moved to HOL theories
haftmann
parents: 60429
diff changeset
   437
  assumes mult_cancel_right [simp]: "a * c = b * c \<longleftrightarrow> c = 0 \<or> a = b"
0826b7025d07 generalized some theorems about integral domains and moved to HOL theories
haftmann
parents: 60429
diff changeset
   438
    and mult_cancel_left [simp]: "c * a = c * b \<longleftrightarrow> c = 0 \<or> a = b"
58952
5d82cdef6c1b equivalence rules for structures without zero divisors
haftmann
parents: 58889
diff changeset
   439
begin
5d82cdef6c1b equivalence rules for structures without zero divisors
haftmann
parents: 58889
diff changeset
   440
5d82cdef6c1b equivalence rules for structures without zero divisors
haftmann
parents: 58889
diff changeset
   441
lemma mult_left_cancel:
5d82cdef6c1b equivalence rules for structures without zero divisors
haftmann
parents: 58889
diff changeset
   442
  "c \<noteq> 0 \<Longrightarrow> c * a = c * b \<longleftrightarrow> a = b"
60562
24af00b010cf Amalgamation of the class comm_semiring_1_diff_distrib into comm_semiring_1_cancel. Moving axiom le_add_diff_inverse2 from semiring_numeral_div to linordered_semidom.
paulson <lp15@cam.ac.uk>
parents: 60529
diff changeset
   443
  by simp
56217
dc429a5b13c4 Some rationalisation of basic lemmas
paulson <lp15@cam.ac.uk>
parents: 55912
diff changeset
   444
58952
5d82cdef6c1b equivalence rules for structures without zero divisors
haftmann
parents: 58889
diff changeset
   445
lemma mult_right_cancel:
5d82cdef6c1b equivalence rules for structures without zero divisors
haftmann
parents: 58889
diff changeset
   446
  "c \<noteq> 0 \<Longrightarrow> a * c = b * c \<longleftrightarrow> a = b"
60562
24af00b010cf Amalgamation of the class comm_semiring_1_diff_distrib into comm_semiring_1_cancel. Moving axiom le_add_diff_inverse2 from semiring_numeral_div to linordered_semidom.
paulson <lp15@cam.ac.uk>
parents: 60529
diff changeset
   447
  by simp
56217
dc429a5b13c4 Some rationalisation of basic lemmas
paulson <lp15@cam.ac.uk>
parents: 55912
diff changeset
   448
25230
022029099a83 continued localization
haftmann
parents: 25193
diff changeset
   449
end
22990
775e9de3db48 added classes ring_no_zero_divisors and dom (non-commutative version of idom);
huffman
parents: 22987
diff changeset
   450
60516
0826b7025d07 generalized some theorems about integral domains and moved to HOL theories
haftmann
parents: 60429
diff changeset
   451
class ring_no_zero_divisors = ring + semiring_no_zero_divisors
0826b7025d07 generalized some theorems about integral domains and moved to HOL theories
haftmann
parents: 60429
diff changeset
   452
begin
0826b7025d07 generalized some theorems about integral domains and moved to HOL theories
haftmann
parents: 60429
diff changeset
   453
0826b7025d07 generalized some theorems about integral domains and moved to HOL theories
haftmann
parents: 60429
diff changeset
   454
subclass semiring_no_zero_divisors_cancel
0826b7025d07 generalized some theorems about integral domains and moved to HOL theories
haftmann
parents: 60429
diff changeset
   455
proof
0826b7025d07 generalized some theorems about integral domains and moved to HOL theories
haftmann
parents: 60429
diff changeset
   456
  fix a b c
0826b7025d07 generalized some theorems about integral domains and moved to HOL theories
haftmann
parents: 60429
diff changeset
   457
  have "a * c = b * c \<longleftrightarrow> (a - b) * c = 0"
0826b7025d07 generalized some theorems about integral domains and moved to HOL theories
haftmann
parents: 60429
diff changeset
   458
    by (simp add: algebra_simps)
0826b7025d07 generalized some theorems about integral domains and moved to HOL theories
haftmann
parents: 60429
diff changeset
   459
  also have "\<dots> \<longleftrightarrow> c = 0 \<or> a = b"
0826b7025d07 generalized some theorems about integral domains and moved to HOL theories
haftmann
parents: 60429
diff changeset
   460
    by auto
0826b7025d07 generalized some theorems about integral domains and moved to HOL theories
haftmann
parents: 60429
diff changeset
   461
  finally show "a * c = b * c \<longleftrightarrow> c = 0 \<or> a = b" .
0826b7025d07 generalized some theorems about integral domains and moved to HOL theories
haftmann
parents: 60429
diff changeset
   462
  have "c * a = c * b \<longleftrightarrow> c * (a - b) = 0"
0826b7025d07 generalized some theorems about integral domains and moved to HOL theories
haftmann
parents: 60429
diff changeset
   463
    by (simp add: algebra_simps)
0826b7025d07 generalized some theorems about integral domains and moved to HOL theories
haftmann
parents: 60429
diff changeset
   464
  also have "\<dots> \<longleftrightarrow> c = 0 \<or> a = b"
0826b7025d07 generalized some theorems about integral domains and moved to HOL theories
haftmann
parents: 60429
diff changeset
   465
    by auto
0826b7025d07 generalized some theorems about integral domains and moved to HOL theories
haftmann
parents: 60429
diff changeset
   466
  finally show "c * a = c * b \<longleftrightarrow> c = 0 \<or> a = b" .
0826b7025d07 generalized some theorems about integral domains and moved to HOL theories
haftmann
parents: 60429
diff changeset
   467
qed
0826b7025d07 generalized some theorems about integral domains and moved to HOL theories
haftmann
parents: 60429
diff changeset
   468
0826b7025d07 generalized some theorems about integral domains and moved to HOL theories
haftmann
parents: 60429
diff changeset
   469
end
0826b7025d07 generalized some theorems about integral domains and moved to HOL theories
haftmann
parents: 60429
diff changeset
   470
23544
4b4165cb3e0d rename class dom to ring_1_no_zero_divisors
huffman
parents: 23527
diff changeset
   471
class ring_1_no_zero_divisors = ring_1 + ring_no_zero_divisors
26274
2bdb61a28971 continued localization
haftmann
parents: 26234
diff changeset
   472
begin
2bdb61a28971 continued localization
haftmann
parents: 26234
diff changeset
   473
36970
fb3fdb4b585e remove simp attribute from square_eq_1_iff
huffman
parents: 36821
diff changeset
   474
lemma square_eq_1_iff:
36821
9207505d1ee5 move lemma real_mult_is_one to Rings.thy, renamed to square_eq_1_iff
huffman
parents: 36719
diff changeset
   475
  "x * x = 1 \<longleftrightarrow> x = 1 \<or> x = - 1"
9207505d1ee5 move lemma real_mult_is_one to Rings.thy, renamed to square_eq_1_iff
huffman
parents: 36719
diff changeset
   476
proof -
9207505d1ee5 move lemma real_mult_is_one to Rings.thy, renamed to square_eq_1_iff
huffman
parents: 36719
diff changeset
   477
  have "(x - 1) * (x + 1) = x * x - 1"
9207505d1ee5 move lemma real_mult_is_one to Rings.thy, renamed to square_eq_1_iff
huffman
parents: 36719
diff changeset
   478
    by (simp add: algebra_simps)
9207505d1ee5 move lemma real_mult_is_one to Rings.thy, renamed to square_eq_1_iff
huffman
parents: 36719
diff changeset
   479
  hence "x * x = 1 \<longleftrightarrow> (x - 1) * (x + 1) = 0"
9207505d1ee5 move lemma real_mult_is_one to Rings.thy, renamed to square_eq_1_iff
huffman
parents: 36719
diff changeset
   480
    by simp
9207505d1ee5 move lemma real_mult_is_one to Rings.thy, renamed to square_eq_1_iff
huffman
parents: 36719
diff changeset
   481
  thus ?thesis
9207505d1ee5 move lemma real_mult_is_one to Rings.thy, renamed to square_eq_1_iff
huffman
parents: 36719
diff changeset
   482
    by (simp add: eq_neg_iff_add_eq_0)
9207505d1ee5 move lemma real_mult_is_one to Rings.thy, renamed to square_eq_1_iff
huffman
parents: 36719
diff changeset
   483
qed
9207505d1ee5 move lemma real_mult_is_one to Rings.thy, renamed to square_eq_1_iff
huffman
parents: 36719
diff changeset
   484
26274
2bdb61a28971 continued localization
haftmann
parents: 26234
diff changeset
   485
lemma mult_cancel_right1 [simp]:
2bdb61a28971 continued localization
haftmann
parents: 26234
diff changeset
   486
  "c = b * c \<longleftrightarrow> c = 0 \<or> b = 1"
29667
53103fc8ffa3 Replaced group_ and ring_simps by algebra_simps;
nipkow
parents: 29465
diff changeset
   487
by (insert mult_cancel_right [of 1 c b], force)
26274
2bdb61a28971 continued localization
haftmann
parents: 26234
diff changeset
   488
2bdb61a28971 continued localization
haftmann
parents: 26234
diff changeset
   489
lemma mult_cancel_right2 [simp]:
2bdb61a28971 continued localization
haftmann
parents: 26234
diff changeset
   490
  "a * c = c \<longleftrightarrow> c = 0 \<or> a = 1"
29667
53103fc8ffa3 Replaced group_ and ring_simps by algebra_simps;
nipkow
parents: 29465
diff changeset
   491
by (insert mult_cancel_right [of a c 1], simp)
60562
24af00b010cf Amalgamation of the class comm_semiring_1_diff_distrib into comm_semiring_1_cancel. Moving axiom le_add_diff_inverse2 from semiring_numeral_div to linordered_semidom.
paulson <lp15@cam.ac.uk>
parents: 60529
diff changeset
   492
26274
2bdb61a28971 continued localization
haftmann
parents: 26234
diff changeset
   493
lemma mult_cancel_left1 [simp]:
2bdb61a28971 continued localization
haftmann
parents: 26234
diff changeset
   494
  "c = c * b \<longleftrightarrow> c = 0 \<or> b = 1"
29667
53103fc8ffa3 Replaced group_ and ring_simps by algebra_simps;
nipkow
parents: 29465
diff changeset
   495
by (insert mult_cancel_left [of c 1 b], force)
26274
2bdb61a28971 continued localization
haftmann
parents: 26234
diff changeset
   496
2bdb61a28971 continued localization
haftmann
parents: 26234
diff changeset
   497
lemma mult_cancel_left2 [simp]:
2bdb61a28971 continued localization
haftmann
parents: 26234
diff changeset
   498
  "c * a = c \<longleftrightarrow> c = 0 \<or> a = 1"
29667
53103fc8ffa3 Replaced group_ and ring_simps by algebra_simps;
nipkow
parents: 29465
diff changeset
   499
by (insert mult_cancel_left [of c a 1], simp)
26274
2bdb61a28971 continued localization
haftmann
parents: 26234
diff changeset
   500
2bdb61a28971 continued localization
haftmann
parents: 26234
diff changeset
   501
end
22990
775e9de3db48 added classes ring_no_zero_divisors and dom (non-commutative version of idom);
huffman
parents: 22987
diff changeset
   502
60562
24af00b010cf Amalgamation of the class comm_semiring_1_diff_distrib into comm_semiring_1_cancel. Moving axiom le_add_diff_inverse2 from semiring_numeral_div to linordered_semidom.
paulson <lp15@cam.ac.uk>
parents: 60529
diff changeset
   503
class semidom = comm_semiring_1_cancel + semiring_no_zero_divisors
59833
ab828c2c5d67 clarified no_zero_devisors: makes only sense in a semiring;
haftmann
parents: 59832
diff changeset
   504
ab828c2c5d67 clarified no_zero_devisors: makes only sense in a semiring;
haftmann
parents: 59832
diff changeset
   505
class idom = comm_ring_1 + semiring_no_zero_divisors
25186
f4d1ebffd025 localized further
haftmann
parents: 25152
diff changeset
   506
begin
14421
ee97b6463cb4 new Ring_and_Field hierarchy, eliminating redundant axioms
paulson
parents: 14398
diff changeset
   507
59833
ab828c2c5d67 clarified no_zero_devisors: makes only sense in a semiring;
haftmann
parents: 59832
diff changeset
   508
subclass semidom ..
ab828c2c5d67 clarified no_zero_devisors: makes only sense in a semiring;
haftmann
parents: 59832
diff changeset
   509
27516
9a5d4a8d4aac by intro_locales -> ..
huffman
parents: 26274
diff changeset
   510
subclass ring_1_no_zero_divisors ..
22990
775e9de3db48 added classes ring_no_zero_divisors and dom (non-commutative version of idom);
huffman
parents: 22987
diff changeset
   511
29981
7d0ed261b712 generalize int_dvd_cancel_factor simproc to idom class
huffman
parents: 29949
diff changeset
   512
lemma dvd_mult_cancel_right [simp]:
7d0ed261b712 generalize int_dvd_cancel_factor simproc to idom class
huffman
parents: 29949
diff changeset
   513
  "a * c dvd b * c \<longleftrightarrow> c = 0 \<or> a dvd b"
7d0ed261b712 generalize int_dvd_cancel_factor simproc to idom class
huffman
parents: 29949
diff changeset
   514
proof -
7d0ed261b712 generalize int_dvd_cancel_factor simproc to idom class
huffman
parents: 29949
diff changeset
   515
  have "a * c dvd b * c \<longleftrightarrow> (\<exists>k. b * c = (a * k) * c)"
57514
bdc2c6b40bf2 prefer ac_simps collections over separate name bindings for add and mult
haftmann
parents: 57512
diff changeset
   516
    unfolding dvd_def by (simp add: ac_simps)
29981
7d0ed261b712 generalize int_dvd_cancel_factor simproc to idom class
huffman
parents: 29949
diff changeset
   517
  also have "(\<exists>k. b * c = (a * k) * c) \<longleftrightarrow> c = 0 \<or> a dvd b"
7d0ed261b712 generalize int_dvd_cancel_factor simproc to idom class
huffman
parents: 29949
diff changeset
   518
    unfolding dvd_def by simp
7d0ed261b712 generalize int_dvd_cancel_factor simproc to idom class
huffman
parents: 29949
diff changeset
   519
  finally show ?thesis .
7d0ed261b712 generalize int_dvd_cancel_factor simproc to idom class
huffman
parents: 29949
diff changeset
   520
qed
7d0ed261b712 generalize int_dvd_cancel_factor simproc to idom class
huffman
parents: 29949
diff changeset
   521
7d0ed261b712 generalize int_dvd_cancel_factor simproc to idom class
huffman
parents: 29949
diff changeset
   522
lemma dvd_mult_cancel_left [simp]:
7d0ed261b712 generalize int_dvd_cancel_factor simproc to idom class
huffman
parents: 29949
diff changeset
   523
  "c * a dvd c * b \<longleftrightarrow> c = 0 \<or> a dvd b"
7d0ed261b712 generalize int_dvd_cancel_factor simproc to idom class
huffman
parents: 29949
diff changeset
   524
proof -
7d0ed261b712 generalize int_dvd_cancel_factor simproc to idom class
huffman
parents: 29949
diff changeset
   525
  have "c * a dvd c * b \<longleftrightarrow> (\<exists>k. b * c = (a * k) * c)"
57514
bdc2c6b40bf2 prefer ac_simps collections over separate name bindings for add and mult
haftmann
parents: 57512
diff changeset
   526
    unfolding dvd_def by (simp add: ac_simps)
29981
7d0ed261b712 generalize int_dvd_cancel_factor simproc to idom class
huffman
parents: 29949
diff changeset
   527
  also have "(\<exists>k. b * c = (a * k) * c) \<longleftrightarrow> c = 0 \<or> a dvd b"
7d0ed261b712 generalize int_dvd_cancel_factor simproc to idom class
huffman
parents: 29949
diff changeset
   528
    unfolding dvd_def by simp
7d0ed261b712 generalize int_dvd_cancel_factor simproc to idom class
huffman
parents: 29949
diff changeset
   529
  finally show ?thesis .
7d0ed261b712 generalize int_dvd_cancel_factor simproc to idom class
huffman
parents: 29949
diff changeset
   530
qed
7d0ed261b712 generalize int_dvd_cancel_factor simproc to idom class
huffman
parents: 29949
diff changeset
   531
60516
0826b7025d07 generalized some theorems about integral domains and moved to HOL theories
haftmann
parents: 60429
diff changeset
   532
lemma square_eq_iff: "a * a = b * b \<longleftrightarrow> a = b \<or> a = - b"
59833
ab828c2c5d67 clarified no_zero_devisors: makes only sense in a semiring;
haftmann
parents: 59832
diff changeset
   533
proof
ab828c2c5d67 clarified no_zero_devisors: makes only sense in a semiring;
haftmann
parents: 59832
diff changeset
   534
  assume "a * a = b * b"
ab828c2c5d67 clarified no_zero_devisors: makes only sense in a semiring;
haftmann
parents: 59832
diff changeset
   535
  then have "(a - b) * (a + b) = 0"
ab828c2c5d67 clarified no_zero_devisors: makes only sense in a semiring;
haftmann
parents: 59832
diff changeset
   536
    by (simp add: algebra_simps)
ab828c2c5d67 clarified no_zero_devisors: makes only sense in a semiring;
haftmann
parents: 59832
diff changeset
   537
  then show "a = b \<or> a = - b"
ab828c2c5d67 clarified no_zero_devisors: makes only sense in a semiring;
haftmann
parents: 59832
diff changeset
   538
    by (simp add: eq_neg_iff_add_eq_0)
ab828c2c5d67 clarified no_zero_devisors: makes only sense in a semiring;
haftmann
parents: 59832
diff changeset
   539
next
ab828c2c5d67 clarified no_zero_devisors: makes only sense in a semiring;
haftmann
parents: 59832
diff changeset
   540
  assume "a = b \<or> a = - b"
ab828c2c5d67 clarified no_zero_devisors: makes only sense in a semiring;
haftmann
parents: 59832
diff changeset
   541
  then show "a * a = b * b" by auto
ab828c2c5d67 clarified no_zero_devisors: makes only sense in a semiring;
haftmann
parents: 59832
diff changeset
   542
qed
ab828c2c5d67 clarified no_zero_devisors: makes only sense in a semiring;
haftmann
parents: 59832
diff changeset
   543
25186
f4d1ebffd025 localized further
haftmann
parents: 25152
diff changeset
   544
end
25152
bfde2f8c0f63 partially localized
haftmann
parents: 25078
diff changeset
   545
35302
4bc6b4d70e08 tuned text
haftmann
parents: 35216
diff changeset
   546
text {*
4bc6b4d70e08 tuned text
haftmann
parents: 35216
diff changeset
   547
  The theory of partially ordered rings is taken from the books:
4bc6b4d70e08 tuned text
haftmann
parents: 35216
diff changeset
   548
  \begin{itemize}
60562
24af00b010cf Amalgamation of the class comm_semiring_1_diff_distrib into comm_semiring_1_cancel. Moving axiom le_add_diff_inverse2 from semiring_numeral_div to linordered_semidom.
paulson <lp15@cam.ac.uk>
parents: 60529
diff changeset
   549
  \item \emph{Lattice Theory} by Garret Birkhoff, American Mathematical Society 1979
35302
4bc6b4d70e08 tuned text
haftmann
parents: 35216
diff changeset
   550
  \item \emph{Partially Ordered Algebraic Systems}, Pergamon Press 1963
4bc6b4d70e08 tuned text
haftmann
parents: 35216
diff changeset
   551
  \end{itemize}
60562
24af00b010cf Amalgamation of the class comm_semiring_1_diff_distrib into comm_semiring_1_cancel. Moving axiom le_add_diff_inverse2 from semiring_numeral_div to linordered_semidom.
paulson <lp15@cam.ac.uk>
parents: 60529
diff changeset
   552
  Most of the used notions can also be looked up in
35302
4bc6b4d70e08 tuned text
haftmann
parents: 35216
diff changeset
   553
  \begin{itemize}
54703
499f92dc6e45 more antiquotations;
wenzelm
parents: 54489
diff changeset
   554
  \item @{url "http://www.mathworld.com"} by Eric Weisstein et. al.
35302
4bc6b4d70e08 tuned text
haftmann
parents: 35216
diff changeset
   555
  \item \emph{Algebra I} by van der Waerden, Springer.
4bc6b4d70e08 tuned text
haftmann
parents: 35216
diff changeset
   556
  \end{itemize}
4bc6b4d70e08 tuned text
haftmann
parents: 35216
diff changeset
   557
*}
4bc6b4d70e08 tuned text
haftmann
parents: 35216
diff changeset
   558
60353
838025c6e278 implicit partial divison operation in integral domains
haftmann
parents: 60352
diff changeset
   559
class divide =
60429
d3d1e185cd63 uniform _ div _ as infix syntax for ring division
haftmann
parents: 60353
diff changeset
   560
  fixes divide :: "'a \<Rightarrow> 'a \<Rightarrow> 'a"  (infixl "div" 70)
60353
838025c6e278 implicit partial divison operation in integral domains
haftmann
parents: 60352
diff changeset
   561
838025c6e278 implicit partial divison operation in integral domains
haftmann
parents: 60352
diff changeset
   562
setup {* Sign.add_const_constraint (@{const_name "divide"}, SOME @{typ "'a \<Rightarrow> 'a \<Rightarrow> 'a"}) *}
838025c6e278 implicit partial divison operation in integral domains
haftmann
parents: 60352
diff changeset
   563
838025c6e278 implicit partial divison operation in integral domains
haftmann
parents: 60352
diff changeset
   564
context semiring
838025c6e278 implicit partial divison operation in integral domains
haftmann
parents: 60352
diff changeset
   565
begin
838025c6e278 implicit partial divison operation in integral domains
haftmann
parents: 60352
diff changeset
   566
838025c6e278 implicit partial divison operation in integral domains
haftmann
parents: 60352
diff changeset
   567
lemma [field_simps]:
60429
d3d1e185cd63 uniform _ div _ as infix syntax for ring division
haftmann
parents: 60353
diff changeset
   568
  shows distrib_left_NO_MATCH: "NO_MATCH (x div y) a \<Longrightarrow> a * (b + c) = a * b + a * c"
d3d1e185cd63 uniform _ div _ as infix syntax for ring division
haftmann
parents: 60353
diff changeset
   569
    and distrib_right_NO_MATCH: "NO_MATCH (x div y) c \<Longrightarrow> (a + b) * c = a * c + b * c"
60353
838025c6e278 implicit partial divison operation in integral domains
haftmann
parents: 60352
diff changeset
   570
  by (rule distrib_left distrib_right)+
838025c6e278 implicit partial divison operation in integral domains
haftmann
parents: 60352
diff changeset
   571
838025c6e278 implicit partial divison operation in integral domains
haftmann
parents: 60352
diff changeset
   572
end
838025c6e278 implicit partial divison operation in integral domains
haftmann
parents: 60352
diff changeset
   573
838025c6e278 implicit partial divison operation in integral domains
haftmann
parents: 60352
diff changeset
   574
context ring
838025c6e278 implicit partial divison operation in integral domains
haftmann
parents: 60352
diff changeset
   575
begin
838025c6e278 implicit partial divison operation in integral domains
haftmann
parents: 60352
diff changeset
   576
838025c6e278 implicit partial divison operation in integral domains
haftmann
parents: 60352
diff changeset
   577
lemma [field_simps]:
60429
d3d1e185cd63 uniform _ div _ as infix syntax for ring division
haftmann
parents: 60353
diff changeset
   578
  shows left_diff_distrib_NO_MATCH: "NO_MATCH (x div y) c \<Longrightarrow> (a - b) * c = a * c - b * c"
d3d1e185cd63 uniform _ div _ as infix syntax for ring division
haftmann
parents: 60353
diff changeset
   579
    and right_diff_distrib_NO_MATCH: "NO_MATCH (x div y) a \<Longrightarrow> a * (b - c) = a * b - a * c"
60353
838025c6e278 implicit partial divison operation in integral domains
haftmann
parents: 60352
diff changeset
   580
  by (rule left_diff_distrib right_diff_distrib)+
838025c6e278 implicit partial divison operation in integral domains
haftmann
parents: 60352
diff changeset
   581
838025c6e278 implicit partial divison operation in integral domains
haftmann
parents: 60352
diff changeset
   582
end
838025c6e278 implicit partial divison operation in integral domains
haftmann
parents: 60352
diff changeset
   583
838025c6e278 implicit partial divison operation in integral domains
haftmann
parents: 60352
diff changeset
   584
setup {* Sign.add_const_constraint (@{const_name "divide"}, SOME @{typ "'a::divide \<Rightarrow> 'a \<Rightarrow> 'a"}) *}
838025c6e278 implicit partial divison operation in integral domains
haftmann
parents: 60352
diff changeset
   585
838025c6e278 implicit partial divison operation in integral domains
haftmann
parents: 60352
diff changeset
   586
class semidom_divide = semidom + divide +
60429
d3d1e185cd63 uniform _ div _ as infix syntax for ring division
haftmann
parents: 60353
diff changeset
   587
  assumes nonzero_mult_divide_cancel_right [simp]: "b \<noteq> 0 \<Longrightarrow> (a * b) div b = a"
d3d1e185cd63 uniform _ div _ as infix syntax for ring division
haftmann
parents: 60353
diff changeset
   588
  assumes divide_zero [simp]: "a div 0 = 0"
60353
838025c6e278 implicit partial divison operation in integral domains
haftmann
parents: 60352
diff changeset
   589
begin
838025c6e278 implicit partial divison operation in integral domains
haftmann
parents: 60352
diff changeset
   590
838025c6e278 implicit partial divison operation in integral domains
haftmann
parents: 60352
diff changeset
   591
lemma nonzero_mult_divide_cancel_left [simp]:
60429
d3d1e185cd63 uniform _ div _ as infix syntax for ring division
haftmann
parents: 60353
diff changeset
   592
  "a \<noteq> 0 \<Longrightarrow> (a * b) div a = b"
60353
838025c6e278 implicit partial divison operation in integral domains
haftmann
parents: 60352
diff changeset
   593
  using nonzero_mult_divide_cancel_right [of a b] by (simp add: ac_simps)
838025c6e278 implicit partial divison operation in integral domains
haftmann
parents: 60352
diff changeset
   594
60516
0826b7025d07 generalized some theorems about integral domains and moved to HOL theories
haftmann
parents: 60429
diff changeset
   595
subclass semiring_no_zero_divisors_cancel
0826b7025d07 generalized some theorems about integral domains and moved to HOL theories
haftmann
parents: 60429
diff changeset
   596
proof
0826b7025d07 generalized some theorems about integral domains and moved to HOL theories
haftmann
parents: 60429
diff changeset
   597
  fix a b c
0826b7025d07 generalized some theorems about integral domains and moved to HOL theories
haftmann
parents: 60429
diff changeset
   598
  { fix a b c
0826b7025d07 generalized some theorems about integral domains and moved to HOL theories
haftmann
parents: 60429
diff changeset
   599
    show "a * c = b * c \<longleftrightarrow> c = 0 \<or> a = b"
0826b7025d07 generalized some theorems about integral domains and moved to HOL theories
haftmann
parents: 60429
diff changeset
   600
    proof (cases "c = 0")
0826b7025d07 generalized some theorems about integral domains and moved to HOL theories
haftmann
parents: 60429
diff changeset
   601
      case True then show ?thesis by simp
0826b7025d07 generalized some theorems about integral domains and moved to HOL theories
haftmann
parents: 60429
diff changeset
   602
    next
0826b7025d07 generalized some theorems about integral domains and moved to HOL theories
haftmann
parents: 60429
diff changeset
   603
      case False
0826b7025d07 generalized some theorems about integral domains and moved to HOL theories
haftmann
parents: 60429
diff changeset
   604
      { assume "a * c = b * c"
0826b7025d07 generalized some theorems about integral domains and moved to HOL theories
haftmann
parents: 60429
diff changeset
   605
        then have "a * c div c = b * c div c"
0826b7025d07 generalized some theorems about integral domains and moved to HOL theories
haftmann
parents: 60429
diff changeset
   606
          by simp
0826b7025d07 generalized some theorems about integral domains and moved to HOL theories
haftmann
parents: 60429
diff changeset
   607
        with False have "a = b"
0826b7025d07 generalized some theorems about integral domains and moved to HOL theories
haftmann
parents: 60429
diff changeset
   608
          by simp
0826b7025d07 generalized some theorems about integral domains and moved to HOL theories
haftmann
parents: 60429
diff changeset
   609
      } then show ?thesis by auto
0826b7025d07 generalized some theorems about integral domains and moved to HOL theories
haftmann
parents: 60429
diff changeset
   610
    qed
0826b7025d07 generalized some theorems about integral domains and moved to HOL theories
haftmann
parents: 60429
diff changeset
   611
  }
0826b7025d07 generalized some theorems about integral domains and moved to HOL theories
haftmann
parents: 60429
diff changeset
   612
  from this [of a c b]
0826b7025d07 generalized some theorems about integral domains and moved to HOL theories
haftmann
parents: 60429
diff changeset
   613
  show "c * a = c * b \<longleftrightarrow> c = 0 \<or> a = b"
0826b7025d07 generalized some theorems about integral domains and moved to HOL theories
haftmann
parents: 60429
diff changeset
   614
    by (simp add: ac_simps)
0826b7025d07 generalized some theorems about integral domains and moved to HOL theories
haftmann
parents: 60429
diff changeset
   615
qed
0826b7025d07 generalized some theorems about integral domains and moved to HOL theories
haftmann
parents: 60429
diff changeset
   616
0826b7025d07 generalized some theorems about integral domains and moved to HOL theories
haftmann
parents: 60429
diff changeset
   617
lemma div_self [simp]:
0826b7025d07 generalized some theorems about integral domains and moved to HOL theories
haftmann
parents: 60429
diff changeset
   618
  assumes "a \<noteq> 0"
0826b7025d07 generalized some theorems about integral domains and moved to HOL theories
haftmann
parents: 60429
diff changeset
   619
  shows "a div a = 1"
0826b7025d07 generalized some theorems about integral domains and moved to HOL theories
haftmann
parents: 60429
diff changeset
   620
  using assms nonzero_mult_divide_cancel_left [of a 1] by simp
0826b7025d07 generalized some theorems about integral domains and moved to HOL theories
haftmann
parents: 60429
diff changeset
   621
60570
7ed2cde6806d more theorems
haftmann
parents: 60562
diff changeset
   622
lemma divide_zero_left [simp]:
7ed2cde6806d more theorems
haftmann
parents: 60562
diff changeset
   623
  "0 div a = 0"
7ed2cde6806d more theorems
haftmann
parents: 60562
diff changeset
   624
proof (cases "a = 0")
7ed2cde6806d more theorems
haftmann
parents: 60562
diff changeset
   625
  case True then show ?thesis by simp
7ed2cde6806d more theorems
haftmann
parents: 60562
diff changeset
   626
next
7ed2cde6806d more theorems
haftmann
parents: 60562
diff changeset
   627
  case False then have "a * 0 div a = 0"
7ed2cde6806d more theorems
haftmann
parents: 60562
diff changeset
   628
    by (rule nonzero_mult_divide_cancel_left)
7ed2cde6806d more theorems
haftmann
parents: 60562
diff changeset
   629
  then show ?thesis by simp
7ed2cde6806d more theorems
haftmann
parents: 60562
diff changeset
   630
qed 
7ed2cde6806d more theorems
haftmann
parents: 60562
diff changeset
   631
60353
838025c6e278 implicit partial divison operation in integral domains
haftmann
parents: 60352
diff changeset
   632
end
838025c6e278 implicit partial divison operation in integral domains
haftmann
parents: 60352
diff changeset
   633
838025c6e278 implicit partial divison operation in integral domains
haftmann
parents: 60352
diff changeset
   634
class idom_divide = idom + semidom_divide
838025c6e278 implicit partial divison operation in integral domains
haftmann
parents: 60352
diff changeset
   635
60517
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   636
class algebraic_semidom = semidom_divide
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   637
begin
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   638
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   639
lemma dvd_div_mult_self [simp]:
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   640
  "a dvd b \<Longrightarrow> b div a * a = b"
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   641
  by (cases "a = 0") (auto elim: dvdE simp add: ac_simps)
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   642
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   643
lemma dvd_mult_div_cancel [simp]:
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   644
  "a dvd b \<Longrightarrow> a * (b div a) = b"
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   645
  using dvd_div_mult_self [of a b] by (simp add: ac_simps)
60562
24af00b010cf Amalgamation of the class comm_semiring_1_diff_distrib into comm_semiring_1_cancel. Moving axiom le_add_diff_inverse2 from semiring_numeral_div to linordered_semidom.
paulson <lp15@cam.ac.uk>
parents: 60529
diff changeset
   646
60517
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   647
lemma div_mult_swap:
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   648
  assumes "c dvd b"
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   649
  shows "a * (b div c) = (a * b) div c"
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   650
proof (cases "c = 0")
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   651
  case True then show ?thesis by simp
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   652
next
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   653
  case False from assms obtain d where "b = c * d" ..
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   654
  moreover from False have "a * divide (d * c) c = ((a * d) * c) div c"
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   655
    by simp
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   656
  ultimately show ?thesis by (simp add: ac_simps)
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   657
qed
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   658
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   659
lemma dvd_div_mult:
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   660
  assumes "c dvd b"
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   661
  shows "b div c * a = (b * a) div c"
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   662
  using assms div_mult_swap [of c b a] by (simp add: ac_simps)
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   663
60570
7ed2cde6806d more theorems
haftmann
parents: 60562
diff changeset
   664
lemma dvd_div_mult2_eq:
7ed2cde6806d more theorems
haftmann
parents: 60562
diff changeset
   665
  assumes "b * c dvd a"
7ed2cde6806d more theorems
haftmann
parents: 60562
diff changeset
   666
  shows "a div (b * c) = a div b div c"
7ed2cde6806d more theorems
haftmann
parents: 60562
diff changeset
   667
using assms proof
7ed2cde6806d more theorems
haftmann
parents: 60562
diff changeset
   668
  fix k
7ed2cde6806d more theorems
haftmann
parents: 60562
diff changeset
   669
  assume "a = b * c * k"
7ed2cde6806d more theorems
haftmann
parents: 60562
diff changeset
   670
  then show ?thesis
7ed2cde6806d more theorems
haftmann
parents: 60562
diff changeset
   671
    by (cases "b = 0 \<or> c = 0") (auto, simp add: ac_simps)
7ed2cde6806d more theorems
haftmann
parents: 60562
diff changeset
   672
qed
7ed2cde6806d more theorems
haftmann
parents: 60562
diff changeset
   673
60562
24af00b010cf Amalgamation of the class comm_semiring_1_diff_distrib into comm_semiring_1_cancel. Moving axiom le_add_diff_inverse2 from semiring_numeral_div to linordered_semidom.
paulson <lp15@cam.ac.uk>
parents: 60529
diff changeset
   674
60517
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   675
text \<open>Units: invertible elements in a ring\<close>
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   676
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   677
abbreviation is_unit :: "'a \<Rightarrow> bool"
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   678
where
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   679
  "is_unit a \<equiv> a dvd 1"
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   680
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   681
lemma not_is_unit_0 [simp]:
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   682
  "\<not> is_unit 0"
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   683
  by simp
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   684
60562
24af00b010cf Amalgamation of the class comm_semiring_1_diff_distrib into comm_semiring_1_cancel. Moving axiom le_add_diff_inverse2 from semiring_numeral_div to linordered_semidom.
paulson <lp15@cam.ac.uk>
parents: 60529
diff changeset
   685
lemma unit_imp_dvd [dest]:
60517
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   686
  "is_unit b \<Longrightarrow> b dvd a"
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   687
  by (rule dvd_trans [of _ 1]) simp_all
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   688
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   689
lemma unit_dvdE:
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   690
  assumes "is_unit a"
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   691
  obtains c where "a \<noteq> 0" and "b = a * c"
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   692
proof -
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   693
  from assms have "a dvd b" by auto
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   694
  then obtain c where "b = a * c" ..
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   695
  moreover from assms have "a \<noteq> 0" by auto
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   696
  ultimately show thesis using that by blast
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   697
qed
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   698
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   699
lemma dvd_unit_imp_unit:
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   700
  "a dvd b \<Longrightarrow> is_unit b \<Longrightarrow> is_unit a"
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   701
  by (rule dvd_trans)
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   702
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   703
lemma unit_div_1_unit [simp, intro]:
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   704
  assumes "is_unit a"
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   705
  shows "is_unit (1 div a)"
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   706
proof -
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   707
  from assms have "1 = 1 div a * a" by simp
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   708
  then show "is_unit (1 div a)" by (rule dvdI)
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   709
qed
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   710
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   711
lemma is_unitE [elim?]:
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   712
  assumes "is_unit a"
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   713
  obtains b where "a \<noteq> 0" and "b \<noteq> 0"
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   714
    and "is_unit b" and "1 div a = b" and "1 div b = a"
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   715
    and "a * b = 1" and "c div a = c * b"
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   716
proof (rule that)
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   717
  def b \<equiv> "1 div a"
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   718
  then show "1 div a = b" by simp
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   719
  from b_def `is_unit a` show "is_unit b" by simp
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   720
  from `is_unit a` and `is_unit b` show "a \<noteq> 0" and "b \<noteq> 0" by auto
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   721
  from b_def `is_unit a` show "a * b = 1" by simp
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   722
  then have "1 = a * b" ..
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   723
  with b_def `b \<noteq> 0` show "1 div b = a" by simp
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   724
  from `is_unit a` have "a dvd c" ..
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   725
  then obtain d where "c = a * d" ..
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   726
  with `a \<noteq> 0` `a * b = 1` show "c div a = c * b"
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   727
    by (simp add: mult.assoc mult.left_commute [of a])
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   728
qed
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   729
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   730
lemma unit_prod [intro]:
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   731
  "is_unit a \<Longrightarrow> is_unit b \<Longrightarrow> is_unit (a * b)"
60562
24af00b010cf Amalgamation of the class comm_semiring_1_diff_distrib into comm_semiring_1_cancel. Moving axiom le_add_diff_inverse2 from semiring_numeral_div to linordered_semidom.
paulson <lp15@cam.ac.uk>
parents: 60529
diff changeset
   732
  by (subst mult_1_left [of 1, symmetric]) (rule mult_dvd_mono)
24af00b010cf Amalgamation of the class comm_semiring_1_diff_distrib into comm_semiring_1_cancel. Moving axiom le_add_diff_inverse2 from semiring_numeral_div to linordered_semidom.
paulson <lp15@cam.ac.uk>
parents: 60529
diff changeset
   733
60517
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   734
lemma unit_div [intro]:
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   735
  "is_unit a \<Longrightarrow> is_unit b \<Longrightarrow> is_unit (a div b)"
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   736
  by (erule is_unitE [of b a]) (simp add: ac_simps unit_prod)
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   737
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   738
lemma mult_unit_dvd_iff:
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   739
  assumes "is_unit b"
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   740
  shows "a * b dvd c \<longleftrightarrow> a dvd c"
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   741
proof
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   742
  assume "a * b dvd c"
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   743
  with assms show "a dvd c"
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   744
    by (simp add: dvd_mult_left)
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   745
next
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   746
  assume "a dvd c"
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   747
  then obtain k where "c = a * k" ..
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   748
  with assms have "c = (a * b) * (1 div b * k)"
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   749
    by (simp add: mult_ac)
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   750
  then show "a * b dvd c" by (rule dvdI)
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   751
qed
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   752
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   753
lemma dvd_mult_unit_iff:
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   754
  assumes "is_unit b"
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   755
  shows "a dvd c * b \<longleftrightarrow> a dvd c"
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   756
proof
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   757
  assume "a dvd c * b"
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   758
  with assms have "c * b dvd c * (b * (1 div b))"
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   759
    by (subst mult_assoc [symmetric]) simp
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   760
  also from `is_unit b` have "b * (1 div b) = 1" by (rule is_unitE) simp
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   761
  finally have "c * b dvd c" by simp
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   762
  with `a dvd c * b` show "a dvd c" by (rule dvd_trans)
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   763
next
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   764
  assume "a dvd c"
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   765
  then show "a dvd c * b" by simp
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   766
qed
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   767
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   768
lemma div_unit_dvd_iff:
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   769
  "is_unit b \<Longrightarrow> a div b dvd c \<longleftrightarrow> a dvd c"
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   770
  by (erule is_unitE [of _ a]) (auto simp add: mult_unit_dvd_iff)
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   771
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   772
lemma dvd_div_unit_iff:
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   773
  "is_unit b \<Longrightarrow> a dvd c div b \<longleftrightarrow> a dvd c"
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   774
  by (erule is_unitE [of _ c]) (simp add: dvd_mult_unit_iff)
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   775
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   776
lemmas unit_dvd_iff = mult_unit_dvd_iff div_unit_dvd_iff
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   777
  dvd_mult_unit_iff dvd_div_unit_iff -- \<open>FIXME consider fact collection\<close>
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   778
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   779
lemma unit_mult_div_div [simp]:
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   780
  "is_unit a \<Longrightarrow> b * (1 div a) = b div a"
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   781
  by (erule is_unitE [of _ b]) simp
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   782
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   783
lemma unit_div_mult_self [simp]:
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   784
  "is_unit a \<Longrightarrow> b div a * a = b"
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   785
  by (rule dvd_div_mult_self) auto
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   786
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   787
lemma unit_div_1_div_1 [simp]:
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   788
  "is_unit a \<Longrightarrow> 1 div (1 div a) = a"
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   789
  by (erule is_unitE) simp
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   790
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   791
lemma unit_div_mult_swap:
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   792
  "is_unit c \<Longrightarrow> a * (b div c) = (a * b) div c"
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   793
  by (erule unit_dvdE [of _ b]) (simp add: mult.left_commute [of _ c])
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   794
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   795
lemma unit_div_commute:
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   796
  "is_unit b \<Longrightarrow> (a div b) * c = (a * c) div b"
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   797
  using unit_div_mult_swap [of b c a] by (simp add: ac_simps)
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   798
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   799
lemma unit_eq_div1:
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   800
  "is_unit b \<Longrightarrow> a div b = c \<longleftrightarrow> a = c * b"
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   801
  by (auto elim: is_unitE)
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   802
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   803
lemma unit_eq_div2:
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   804
  "is_unit b \<Longrightarrow> a = c div b \<longleftrightarrow> a * b = c"
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   805
  using unit_eq_div1 [of b c a] by auto
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   806
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   807
lemma unit_mult_left_cancel:
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   808
  assumes "is_unit a"
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   809
  shows "a * b = a * c \<longleftrightarrow> b = c" (is "?P \<longleftrightarrow> ?Q")
60562
24af00b010cf Amalgamation of the class comm_semiring_1_diff_distrib into comm_semiring_1_cancel. Moving axiom le_add_diff_inverse2 from semiring_numeral_div to linordered_semidom.
paulson <lp15@cam.ac.uk>
parents: 60529
diff changeset
   810
  using assms mult_cancel_left [of a b c] by auto
60517
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   811
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   812
lemma unit_mult_right_cancel:
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   813
  "is_unit a \<Longrightarrow> b * a = c * a \<longleftrightarrow> b = c"
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   814
  using unit_mult_left_cancel [of a b c] by (auto simp add: ac_simps)
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   815
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   816
lemma unit_div_cancel:
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   817
  assumes "is_unit a"
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   818
  shows "b div a = c div a \<longleftrightarrow> b = c"
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   819
proof -
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   820
  from assms have "is_unit (1 div a)" by simp
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   821
  then have "b * (1 div a) = c * (1 div a) \<longleftrightarrow> b = c"
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   822
    by (rule unit_mult_right_cancel)
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   823
  with assms show ?thesis by simp
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   824
qed
60562
24af00b010cf Amalgamation of the class comm_semiring_1_diff_distrib into comm_semiring_1_cancel. Moving axiom le_add_diff_inverse2 from semiring_numeral_div to linordered_semidom.
paulson <lp15@cam.ac.uk>
parents: 60529
diff changeset
   825
60570
7ed2cde6806d more theorems
haftmann
parents: 60562
diff changeset
   826
lemma is_unit_div_mult2_eq:
7ed2cde6806d more theorems
haftmann
parents: 60562
diff changeset
   827
  assumes "is_unit b" and "is_unit c"
7ed2cde6806d more theorems
haftmann
parents: 60562
diff changeset
   828
  shows "a div (b * c) = a div b div c"
7ed2cde6806d more theorems
haftmann
parents: 60562
diff changeset
   829
proof -
7ed2cde6806d more theorems
haftmann
parents: 60562
diff changeset
   830
  from assms have "is_unit (b * c)" by (simp add: unit_prod)
7ed2cde6806d more theorems
haftmann
parents: 60562
diff changeset
   831
  then have "b * c dvd a"
7ed2cde6806d more theorems
haftmann
parents: 60562
diff changeset
   832
    by (rule unit_imp_dvd)
7ed2cde6806d more theorems
haftmann
parents: 60562
diff changeset
   833
  then show ?thesis
7ed2cde6806d more theorems
haftmann
parents: 60562
diff changeset
   834
    by (rule dvd_div_mult2_eq)
7ed2cde6806d more theorems
haftmann
parents: 60562
diff changeset
   835
qed
7ed2cde6806d more theorems
haftmann
parents: 60562
diff changeset
   836
60517
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   837
60529
24c2ef12318b avoid suspicious Unicode;
wenzelm
parents: 60517
diff changeset
   838
text \<open>Associated elements in a ring --- an equivalence relation induced
24c2ef12318b avoid suspicious Unicode;
wenzelm
parents: 60517
diff changeset
   839
  by the quasi-order divisibility.\<close>
60517
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   840
60562
24af00b010cf Amalgamation of the class comm_semiring_1_diff_distrib into comm_semiring_1_cancel. Moving axiom le_add_diff_inverse2 from semiring_numeral_div to linordered_semidom.
paulson <lp15@cam.ac.uk>
parents: 60529
diff changeset
   841
definition associated :: "'a \<Rightarrow> 'a \<Rightarrow> bool"
60517
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   842
where
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   843
  "associated a b \<longleftrightarrow> a dvd b \<and> b dvd a"
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   844
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   845
lemma associatedI:
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   846
  "a dvd b \<Longrightarrow> b dvd a \<Longrightarrow> associated a b"
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   847
  by (simp add: associated_def)
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   848
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   849
lemma associatedD1:
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   850
  "associated a b \<Longrightarrow> a dvd b"
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   851
  by (simp add: associated_def)
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   852
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   853
lemma associatedD2:
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   854
  "associated a b \<Longrightarrow> b dvd a"
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   855
  by (simp add: associated_def)
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   856
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   857
lemma associated_refl [simp]:
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   858
  "associated a a"
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   859
  by (auto intro: associatedI)
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   860
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   861
lemma associated_sym:
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   862
  "associated b a \<longleftrightarrow> associated a b"
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   863
  by (auto intro: associatedI dest: associatedD1 associatedD2)
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   864
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   865
lemma associated_trans:
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   866
  "associated a b \<Longrightarrow> associated b c \<Longrightarrow> associated a c"
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   867
  by (auto intro: associatedI dvd_trans dest: associatedD1 associatedD2)
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   868
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   869
lemma associated_0 [simp]:
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   870
  "associated 0 b \<longleftrightarrow> b = 0"
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   871
  "associated a 0 \<longleftrightarrow> a = 0"
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   872
  by (auto dest: associatedD1 associatedD2)
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   873
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   874
lemma associated_unit:
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   875
  "associated a b \<Longrightarrow> is_unit a \<Longrightarrow> is_unit b"
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   876
  using dvd_unit_imp_unit by (auto dest!: associatedD1 associatedD2)
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   877
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   878
lemma is_unit_associatedI:
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   879
  assumes "is_unit c" and "a = c * b"
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   880
  shows "associated a b"
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   881
proof (rule associatedI)
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   882
  from `a = c * b` show "b dvd a" by auto
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   883
  from `is_unit c` obtain d where "c * d = 1" by (rule is_unitE)
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   884
  moreover from `a = c * b` have "d * a = d * (c * b)" by simp
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   885
  ultimately have "b = a * d" by (simp add: ac_simps)
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   886
  then show "a dvd b" ..
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   887
qed
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   888
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   889
lemma associated_is_unitE:
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   890
  assumes "associated a b"
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   891
  obtains c where "is_unit c" and "a = c * b"
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   892
proof (cases "b = 0")
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   893
  case True with assms have "is_unit 1" and "a = 1 * b" by simp_all
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   894
  with that show thesis .
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   895
next
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   896
  case False
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   897
  from assms have "a dvd b" and "b dvd a" by (auto dest: associatedD1 associatedD2)
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   898
  then obtain c d where "b = a * d" and "a = b * c" by (blast elim: dvdE)
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   899
  then have "a = c * b" and "(c * d) * b = 1 * b" by (simp_all add: ac_simps)
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   900
  with False have "c * d = 1" using mult_cancel_right [of "c * d" b 1] by simp
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   901
  then have "is_unit c" by auto
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   902
  with `a = c * b` that show thesis by blast
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   903
qed
60562
24af00b010cf Amalgamation of the class comm_semiring_1_diff_distrib into comm_semiring_1_cancel. Moving axiom le_add_diff_inverse2 from semiring_numeral_div to linordered_semidom.
paulson <lp15@cam.ac.uk>
parents: 60529
diff changeset
   904
24af00b010cf Amalgamation of the class comm_semiring_1_diff_distrib into comm_semiring_1_cancel. Moving axiom le_add_diff_inverse2 from semiring_numeral_div to linordered_semidom.
paulson <lp15@cam.ac.uk>
parents: 60529
diff changeset
   905
lemmas unit_simps = mult_unit_dvd_iff div_unit_dvd_iff dvd_mult_unit_iff
60517
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   906
  dvd_div_unit_iff unit_div_mult_swap unit_div_commute
60562
24af00b010cf Amalgamation of the class comm_semiring_1_diff_distrib into comm_semiring_1_cancel. Moving axiom le_add_diff_inverse2 from semiring_numeral_div to linordered_semidom.
paulson <lp15@cam.ac.uk>
parents: 60529
diff changeset
   907
  unit_mult_left_cancel unit_mult_right_cancel unit_div_cancel
60517
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   908
  unit_eq_div1 unit_eq_div2
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   909
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   910
end
f16e4fb20652 separate class for notions specific for integral (semi)domains, in contrast to fields where these are trivial
haftmann
parents: 60516
diff changeset
   911
38642
8fa437809c67 dropped type classes mult_mono and mult_mono1; tuned names of technical rule duplicates
haftmann
parents: 37767
diff changeset
   912
class ordered_semiring = semiring + comm_monoid_add + ordered_ab_semigroup_add +
8fa437809c67 dropped type classes mult_mono and mult_mono1; tuned names of technical rule duplicates
haftmann
parents: 37767
diff changeset
   913
  assumes mult_left_mono: "a \<le> b \<Longrightarrow> 0 \<le> c \<Longrightarrow> c * a \<le> c * b"
8fa437809c67 dropped type classes mult_mono and mult_mono1; tuned names of technical rule duplicates
haftmann
parents: 37767
diff changeset
   914
  assumes mult_right_mono: "a \<le> b \<Longrightarrow> 0 \<le> c \<Longrightarrow> a * c \<le> b * c"
25230
022029099a83 continued localization
haftmann
parents: 25193
diff changeset
   915
begin
022029099a83 continued localization
haftmann
parents: 25193
diff changeset
   916
022029099a83 continued localization
haftmann
parents: 25193
diff changeset
   917
lemma mult_mono:
38642
8fa437809c67 dropped type classes mult_mono and mult_mono1; tuned names of technical rule duplicates
haftmann
parents: 37767
diff changeset
   918
  "a \<le> b \<Longrightarrow> c \<le> d \<Longrightarrow> 0 \<le> b \<Longrightarrow> 0 \<le> c \<Longrightarrow> a * c \<le> b * d"
25230
022029099a83 continued localization
haftmann
parents: 25193
diff changeset
   919
apply (erule mult_right_mono [THEN order_trans], assumption)
022029099a83 continued localization
haftmann
parents: 25193
diff changeset
   920
apply (erule mult_left_mono, assumption)
022029099a83 continued localization
haftmann
parents: 25193
diff changeset
   921
done
022029099a83 continued localization
haftmann
parents: 25193
diff changeset
   922
022029099a83 continued localization
haftmann
parents: 25193
diff changeset
   923
lemma mult_mono':
38642
8fa437809c67 dropped type classes mult_mono and mult_mono1; tuned names of technical rule duplicates
haftmann
parents: 37767
diff changeset
   924
  "a \<le> b \<Longrightarrow> c \<le> d \<Longrightarrow> 0 \<le> a \<Longrightarrow> 0 \<le> c \<Longrightarrow> a * c \<le> b * d"
25230
022029099a83 continued localization
haftmann
parents: 25193
diff changeset
   925
apply (rule mult_mono)
022029099a83 continued localization
haftmann
parents: 25193
diff changeset
   926
apply (fast intro: order_trans)+
022029099a83 continued localization
haftmann
parents: 25193
diff changeset
   927
done
022029099a83 continued localization
haftmann
parents: 25193
diff changeset
   928
022029099a83 continued localization
haftmann
parents: 25193
diff changeset
   929
end
21199
2d83f93c3580 * Added annihilation axioms ("x * 0 = 0") to axclass semiring_0.
krauss
parents: 20633
diff changeset
   930
38642
8fa437809c67 dropped type classes mult_mono and mult_mono1; tuned names of technical rule duplicates
haftmann
parents: 37767
diff changeset
   931
class ordered_cancel_semiring = ordered_semiring + cancel_comm_monoid_add
25267
1f745c599b5c proper reinitialisation after subclass
haftmann
parents: 25238
diff changeset
   932
begin
14268
5cf13e80be0e Removal of Hyperreal/ExtraThms2.ML, sending the material to the correct files.
paulson
parents: 14267
diff changeset
   933
27516
9a5d4a8d4aac by intro_locales -> ..
huffman
parents: 26274
diff changeset
   934
subclass semiring_0_cancel ..
23521
195fe3fe2831 added ordered_ring and ordered_semiring
obua
parents: 23496
diff changeset
   935
56536
aefb4a8da31f made mult_nonneg_nonneg a simp rule
nipkow
parents: 56480
diff changeset
   936
lemma mult_nonneg_nonneg[simp]: "0 \<le> a \<Longrightarrow> 0 \<le> b \<Longrightarrow> 0 \<le> a * b"
36301
72f4d079ebf8 more localization; factored out lemmas for division_ring
haftmann
parents: 35828
diff changeset
   937
using mult_left_mono [of 0 b a] by simp
25230
022029099a83 continued localization
haftmann
parents: 25193
diff changeset
   938
022029099a83 continued localization
haftmann
parents: 25193
diff changeset
   939
lemma mult_nonneg_nonpos: "0 \<le> a \<Longrightarrow> b \<le> 0 \<Longrightarrow> a * b \<le> 0"
36301
72f4d079ebf8 more localization; factored out lemmas for division_ring
haftmann
parents: 35828
diff changeset
   940
using mult_left_mono [of b 0 a] by simp
30692
44ea10bc07a7 clean up proofs of sign rules for multiplication; add list of lemmas mult_sign_intros
huffman
parents: 30650
diff changeset
   941
44ea10bc07a7 clean up proofs of sign rules for multiplication; add list of lemmas mult_sign_intros
huffman
parents: 30650
diff changeset
   942
lemma mult_nonpos_nonneg: "a \<le> 0 \<Longrightarrow> 0 \<le> b \<Longrightarrow> a * b \<le> 0"
36301
72f4d079ebf8 more localization; factored out lemmas for division_ring
haftmann
parents: 35828
diff changeset
   943
using mult_right_mono [of a 0 b] by simp
30692
44ea10bc07a7 clean up proofs of sign rules for multiplication; add list of lemmas mult_sign_intros
huffman
parents: 30650
diff changeset
   944
44ea10bc07a7 clean up proofs of sign rules for multiplication; add list of lemmas mult_sign_intros
huffman
parents: 30650
diff changeset
   945
text {* Legacy - use @{text mult_nonpos_nonneg} *}
60562
24af00b010cf Amalgamation of the class comm_semiring_1_diff_distrib into comm_semiring_1_cancel. Moving axiom le_add_diff_inverse2 from semiring_numeral_div to linordered_semidom.
paulson <lp15@cam.ac.uk>
parents: 60529
diff changeset
   946
lemma mult_nonneg_nonpos2: "0 \<le> a \<Longrightarrow> b \<le> 0 \<Longrightarrow> b * a \<le> 0"
36301
72f4d079ebf8 more localization; factored out lemmas for division_ring
haftmann
parents: 35828
diff changeset
   947
by (drule mult_right_mono [of b 0], auto)
25230
022029099a83 continued localization
haftmann
parents: 25193
diff changeset
   948
60562
24af00b010cf Amalgamation of the class comm_semiring_1_diff_distrib into comm_semiring_1_cancel. Moving axiom le_add_diff_inverse2 from semiring_numeral_div to linordered_semidom.
paulson <lp15@cam.ac.uk>
parents: 60529
diff changeset
   949
lemma split_mult_neg_le: "(0 \<le> a & b \<le> 0) | (a \<le> 0 & 0 \<le> b) \<Longrightarrow> a * b \<le> 0"
29667
53103fc8ffa3 Replaced group_ and ring_simps by algebra_simps;
nipkow
parents: 29465
diff changeset
   950
by (auto simp add: mult_nonneg_nonpos mult_nonneg_nonpos2)
25230
022029099a83 continued localization
haftmann
parents: 25193
diff changeset
   951
022029099a83 continued localization
haftmann
parents: 25193
diff changeset
   952
end
022029099a83 continued localization
haftmann
parents: 25193
diff changeset
   953
38642
8fa437809c67 dropped type classes mult_mono and mult_mono1; tuned names of technical rule duplicates
haftmann
parents: 37767
diff changeset
   954
class linordered_semiring = ordered_semiring + linordered_cancel_ab_semigroup_add
25267
1f745c599b5c proper reinitialisation after subclass
haftmann
parents: 25238
diff changeset
   955
begin
25230
022029099a83 continued localization
haftmann
parents: 25193
diff changeset
   956
35028
108662d50512 more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents: 34146
diff changeset
   957
subclass ordered_cancel_semiring ..
108662d50512 more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents: 34146
diff changeset
   958
108662d50512 more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents: 34146
diff changeset
   959
subclass ordered_comm_monoid_add ..
25304
7491c00f0915 removed subclass edge ordered_ring < lordered_ring
haftmann
parents: 25267
diff changeset
   960
25230
022029099a83 continued localization
haftmann
parents: 25193
diff changeset
   961
lemma mult_left_less_imp_less:
022029099a83 continued localization
haftmann
parents: 25193
diff changeset
   962
  "c * a < c * b \<Longrightarrow> 0 \<le> c \<Longrightarrow> a < b"
29667
53103fc8ffa3 Replaced group_ and ring_simps by algebra_simps;
nipkow
parents: 29465
diff changeset
   963
by (force simp add: mult_left_mono not_le [symmetric])
60562
24af00b010cf Amalgamation of the class comm_semiring_1_diff_distrib into comm_semiring_1_cancel. Moving axiom le_add_diff_inverse2 from semiring_numeral_div to linordered_semidom.
paulson <lp15@cam.ac.uk>
parents: 60529
diff changeset
   964
25230
022029099a83 continued localization
haftmann
parents: 25193
diff changeset
   965
lemma mult_right_less_imp_less:
022029099a83 continued localization
haftmann
parents: 25193
diff changeset
   966
  "a * c < b * c \<Longrightarrow> 0 \<le> c \<Longrightarrow> a < b"
29667
53103fc8ffa3 Replaced group_ and ring_simps by algebra_simps;
nipkow
parents: 29465
diff changeset
   967
by (force simp add: mult_right_mono not_le [symmetric])
23521
195fe3fe2831 added ordered_ring and ordered_semiring
obua
parents: 23496
diff changeset
   968
25186
f4d1ebffd025 localized further
haftmann
parents: 25152
diff changeset
   969
end
25152
bfde2f8c0f63 partially localized
haftmann
parents: 25078
diff changeset
   970
35043
07dbdf60d5ad dropped accidental duplication of "lin" prefix from cs. 108662d50512
haftmann
parents: 35032
diff changeset
   971
class linordered_semiring_1 = linordered_semiring + semiring_1
36622
e393a91f86df Generalize swap_inj_on; add simps for Times; add Ex_list_of_length, log_inj; Added missing locale edges for linordered semiring with 1.
hoelzl
parents: 36348
diff changeset
   972
begin
e393a91f86df Generalize swap_inj_on; add simps for Times; add Ex_list_of_length, log_inj; Added missing locale edges for linordered semiring with 1.
hoelzl
parents: 36348
diff changeset
   973
e393a91f86df Generalize swap_inj_on; add simps for Times; add Ex_list_of_length, log_inj; Added missing locale edges for linordered semiring with 1.
hoelzl
parents: 36348
diff changeset
   974
lemma convex_bound_le:
e393a91f86df Generalize swap_inj_on; add simps for Times; add Ex_list_of_length, log_inj; Added missing locale edges for linordered semiring with 1.
hoelzl
parents: 36348
diff changeset
   975
  assumes "x \<le> a" "y \<le> a" "0 \<le> u" "0 \<le> v" "u + v = 1"
e393a91f86df Generalize swap_inj_on; add simps for Times; add Ex_list_of_length, log_inj; Added missing locale edges for linordered semiring with 1.
hoelzl
parents: 36348
diff changeset
   976
  shows "u * x + v * y \<le> a"
e393a91f86df Generalize swap_inj_on; add simps for Times; add Ex_list_of_length, log_inj; Added missing locale edges for linordered semiring with 1.
hoelzl
parents: 36348
diff changeset
   977
proof-
e393a91f86df Generalize swap_inj_on; add simps for Times; add Ex_list_of_length, log_inj; Added missing locale edges for linordered semiring with 1.
hoelzl
parents: 36348
diff changeset
   978
  from assms have "u * x + v * y \<le> u * a + v * a"
e393a91f86df Generalize swap_inj_on; add simps for Times; add Ex_list_of_length, log_inj; Added missing locale edges for linordered semiring with 1.
hoelzl
parents: 36348
diff changeset
   979
    by (simp add: add_mono mult_left_mono)
49962
a8cc904a6820 Renamed {left,right}_distrib to distrib_{right,left}.
webertj
parents: 44921
diff changeset
   980
  thus ?thesis using assms unfolding distrib_right[symmetric] by simp
36622
e393a91f86df Generalize swap_inj_on; add simps for Times; add Ex_list_of_length, log_inj; Added missing locale edges for linordered semiring with 1.
hoelzl
parents: 36348
diff changeset
   981
qed
e393a91f86df Generalize swap_inj_on; add simps for Times; add Ex_list_of_length, log_inj; Added missing locale edges for linordered semiring with 1.
hoelzl
parents: 36348
diff changeset
   982
e393a91f86df Generalize swap_inj_on; add simps for Times; add Ex_list_of_length, log_inj; Added missing locale edges for linordered semiring with 1.
hoelzl
parents: 36348
diff changeset
   983
end
35043
07dbdf60d5ad dropped accidental duplication of "lin" prefix from cs. 108662d50512
haftmann
parents: 35032
diff changeset
   984
07dbdf60d5ad dropped accidental duplication of "lin" prefix from cs. 108662d50512
haftmann
parents: 35032
diff changeset
   985
class linordered_semiring_strict = semiring + comm_monoid_add + linordered_cancel_ab_semigroup_add +
25062
af5ef0d4d655 global class syntax
haftmann
parents: 24748
diff changeset
   986
  assumes mult_strict_left_mono: "a < b \<Longrightarrow> 0 < c \<Longrightarrow> c * a < c * b"
af5ef0d4d655 global class syntax
haftmann
parents: 24748
diff changeset
   987
  assumes mult_strict_right_mono: "a < b \<Longrightarrow> 0 < c \<Longrightarrow> a * c < b * c"
25267
1f745c599b5c proper reinitialisation after subclass
haftmann
parents: 25238
diff changeset
   988
begin
14341
a09441bd4f1e Ring_and_Field now requires axiom add_left_imp_eq for semirings.
paulson
parents: 14334
diff changeset
   989
27516
9a5d4a8d4aac by intro_locales -> ..
huffman
parents: 26274
diff changeset
   990
subclass semiring_0_cancel ..
14940
b9ab8babd8b3 Further development of matrix theory
obua
parents: 14770
diff changeset
   991
35028
108662d50512 more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents: 34146
diff changeset
   992
subclass linordered_semiring
28823
dcbef866c9e2 tuned unfold_locales invocation
haftmann
parents: 28559
diff changeset
   993
proof
23550
d4f1d6ef119c convert instance proofs to Isar style
huffman
parents: 23544
diff changeset
   994
  fix a b c :: 'a
d4f1d6ef119c convert instance proofs to Isar style
huffman
parents: 23544
diff changeset
   995
  assume A: "a \<le> b" "0 \<le> c"
d4f1d6ef119c convert instance proofs to Isar style
huffman
parents: 23544
diff changeset
   996
  from A show "c * a \<le> c * b"
25186
f4d1ebffd025 localized further
haftmann
parents: 25152
diff changeset
   997
    unfolding le_less
f4d1ebffd025 localized further
haftmann
parents: 25152
diff changeset
   998
    using mult_strict_left_mono by (cases "c = 0") auto
23550
d4f1d6ef119c convert instance proofs to Isar style
huffman
parents: 23544
diff changeset
   999
  from A show "a * c \<le> b * c"
25152
bfde2f8c0f63 partially localized
haftmann
parents: 25078
diff changeset
  1000
    unfolding le_less
25186
f4d1ebffd025 localized further
haftmann
parents: 25152
diff changeset
  1001
    using mult_strict_right_mono by (cases "c = 0") auto
25152
bfde2f8c0f63 partially localized
haftmann
parents: 25078
diff changeset
  1002
qed
bfde2f8c0f63 partially localized
haftmann
parents: 25078
diff changeset
  1003
25230
022029099a83 continued localization
haftmann
parents: 25193
diff changeset
  1004
lemma mult_left_le_imp_le:
022029099a83 continued localization
haftmann
parents: 25193
diff changeset
  1005
  "c * a \<le> c * b \<Longrightarrow> 0 < c \<Longrightarrow> a \<le> b"
29667
53103fc8ffa3 Replaced group_ and ring_simps by algebra_simps;
nipkow
parents: 29465
diff changeset
  1006
by (force simp add: mult_strict_left_mono _not_less [symmetric])
60562
24af00b010cf Amalgamation of the class comm_semiring_1_diff_distrib into comm_semiring_1_cancel. Moving axiom le_add_diff_inverse2 from semiring_numeral_div to linordered_semidom.
paulson <lp15@cam.ac.uk>
parents: 60529
diff changeset
  1007
25230
022029099a83 continued localization
haftmann
parents: 25193
diff changeset
  1008
lemma mult_right_le_imp_le:
022029099a83 continued localization
haftmann
parents: 25193
diff changeset
  1009
  "a * c \<le> b * c \<Longrightarrow> 0 < c \<Longrightarrow> a \<le> b"
29667
53103fc8ffa3 Replaced group_ and ring_simps by algebra_simps;
nipkow
parents: 29465
diff changeset
  1010
by (force simp add: mult_strict_right_mono not_less [symmetric])
25230
022029099a83 continued localization
haftmann
parents: 25193
diff changeset
  1011
56544
b60d5d119489 made mult_pos_pos a simp rule
nipkow
parents: 56536
diff changeset
  1012
lemma mult_pos_pos[simp]: "0 < a \<Longrightarrow> 0 < b \<Longrightarrow> 0 < a * b"
36301
72f4d079ebf8 more localization; factored out lemmas for division_ring
haftmann
parents: 35828
diff changeset
  1013
using mult_strict_left_mono [of 0 b a] by simp
30692
44ea10bc07a7 clean up proofs of sign rules for multiplication; add list of lemmas mult_sign_intros
huffman
parents: 30650
diff changeset
  1014
44ea10bc07a7 clean up proofs of sign rules for multiplication; add list of lemmas mult_sign_intros
huffman
parents: 30650
diff changeset
  1015
lemma mult_pos_neg: "0 < a \<Longrightarrow> b < 0 \<Longrightarrow> a * b < 0"
36301
72f4d079ebf8 more localization; factored out lemmas for division_ring
haftmann
parents: 35828
diff changeset
  1016
using mult_strict_left_mono [of b 0 a] by simp
30692
44ea10bc07a7 clean up proofs of sign rules for multiplication; add list of lemmas mult_sign_intros
huffman
parents: 30650
diff changeset
  1017
44ea10bc07a7 clean up proofs of sign rules for multiplication; add list of lemmas mult_sign_intros
huffman
parents: 30650
diff changeset
  1018
lemma mult_neg_pos: "a < 0 \<Longrightarrow> 0 < b \<Longrightarrow> a * b < 0"
36301
72f4d079ebf8 more localization; factored out lemmas for division_ring
haftmann
parents: 35828
diff changeset
  1019
using mult_strict_right_mono [of a 0 b] by simp
30692
44ea10bc07a7 clean up proofs of sign rules for multiplication; add list of lemmas mult_sign_intros
huffman
parents: 30650
diff changeset
  1020
44ea10bc07a7 clean up proofs of sign rules for multiplication; add list of lemmas mult_sign_intros
huffman
parents: 30650
diff changeset
  1021
text {* Legacy - use @{text mult_neg_pos} *}
60562
24af00b010cf Amalgamation of the class comm_semiring_1_diff_distrib into comm_semiring_1_cancel. Moving axiom le_add_diff_inverse2 from semiring_numeral_div to linordered_semidom.
paulson <lp15@cam.ac.uk>
parents: 60529
diff changeset
  1022
lemma mult_pos_neg2: "0 < a \<Longrightarrow> b < 0 \<Longrightarrow> b * a < 0"
36301
72f4d079ebf8 more localization; factored out lemmas for division_ring
haftmann
parents: 35828
diff changeset
  1023
by (drule mult_strict_right_mono [of b 0], auto)
25230
022029099a83 continued localization
haftmann
parents: 25193
diff changeset
  1024
022029099a83 continued localization
haftmann
parents: 25193
diff changeset
  1025
lemma zero_less_mult_pos:
022029099a83 continued localization
haftmann
parents: 25193
diff changeset
  1026
  "0 < a * b \<Longrightarrow> 0 < a \<Longrightarrow> 0 < b"
30692
44ea10bc07a7 clean up proofs of sign rules for multiplication; add list of lemmas mult_sign_intros
huffman
parents: 30650
diff changeset
  1027
apply (cases "b\<le>0")
25230
022029099a83 continued localization
haftmann
parents: 25193
diff changeset
  1028
 apply (auto simp add: le_less not_less)
30692
44ea10bc07a7 clean up proofs of sign rules for multiplication; add list of lemmas mult_sign_intros
huffman
parents: 30650
diff changeset
  1029
apply (drule_tac mult_pos_neg [of a b])
25230
022029099a83 continued localization
haftmann
parents: 25193
diff changeset
  1030
 apply (auto dest: less_not_sym)
022029099a83 continued localization
haftmann
parents: 25193
diff changeset
  1031
done
022029099a83 continued localization
haftmann
parents: 25193
diff changeset
  1032
022029099a83 continued localization
haftmann
parents: 25193
diff changeset
  1033
lemma zero_less_mult_pos2:
022029099a83 continued localization
haftmann
parents: 25193
diff changeset
  1034
  "0 < b * a \<Longrightarrow> 0 < a \<Longrightarrow> 0 < b"
30692
44ea10bc07a7 clean up proofs of sign rules for multiplication; add list of lemmas mult_sign_intros
huffman
parents: 30650
diff changeset
  1035
apply (cases "b\<le>0")
25230
022029099a83 continued localization
haftmann
parents: 25193
diff changeset
  1036
 apply (auto simp add: le_less not_less)
30692
44ea10bc07a7 clean up proofs of sign rules for multiplication; add list of lemmas mult_sign_intros
huffman
parents: 30650
diff changeset
  1037
apply (drule_tac mult_pos_neg2 [of a b])
25230
022029099a83 continued localization
haftmann
parents: 25193
diff changeset
  1038
 apply (auto dest: less_not_sym)
022029099a83 continued localization
haftmann
parents: 25193
diff changeset
  1039
done
022029099a83 continued localization
haftmann
parents: 25193
diff changeset
  1040
26193
37a7eb7fd5f7 continued localization
haftmann
parents: 25917
diff changeset
  1041
text{*Strict monotonicity in both arguments*}
37a7eb7fd5f7 continued localization
haftmann
parents: 25917
diff changeset
  1042
lemma mult_strict_mono:
37a7eb7fd5f7 continued localization
haftmann
parents: 25917
diff changeset
  1043
  assumes "a < b" and "c < d" and "0 < b" and "0 \<le> c"
37a7eb7fd5f7 continued localization
haftmann
parents: 25917
diff changeset
  1044
  shows "a * c < b * d"
37a7eb7fd5f7 continued localization
haftmann
parents: 25917
diff changeset
  1045
  using assms apply (cases "c=0")
56544
b60d5d119489 made mult_pos_pos a simp rule
nipkow
parents: 56536
diff changeset
  1046
  apply (simp)
26193
37a7eb7fd5f7 continued localization
haftmann
parents: 25917
diff changeset
  1047
  apply (erule mult_strict_right_mono [THEN less_trans])
30692
44ea10bc07a7 clean up proofs of sign rules for multiplication; add list of lemmas mult_sign_intros
huffman
parents: 30650
diff changeset
  1048
  apply (force simp add: le_less)
26193
37a7eb7fd5f7 continued localization
haftmann
parents: 25917
diff changeset
  1049
  apply (erule mult_strict_left_mono, assumption)
37a7eb7fd5f7 continued localization
haftmann
parents: 25917
diff changeset
  1050
  done
37a7eb7fd5f7 continued localization
haftmann
parents: 25917
diff changeset
  1051
37a7eb7fd5f7 continued localization
haftmann
parents: 25917
diff changeset
  1052
text{*This weaker variant has more natural premises*}
37a7eb7fd5f7 continued localization
haftmann
parents: 25917
diff changeset
  1053
lemma mult_strict_mono':
37a7eb7fd5f7 continued localization
haftmann
parents: 25917
diff changeset
  1054
  assumes "a < b" and "c < d" and "0 \<le> a" and "0 \<le> c"
37a7eb7fd5f7 continued localization
haftmann
parents: 25917
diff changeset
  1055
  shows "a * c < b * d"
29667
53103fc8ffa3 Replaced group_ and ring_simps by algebra_simps;
nipkow
parents: 29465
diff changeset
  1056
by (rule mult_strict_mono) (insert assms, auto)
26193
37a7eb7fd5f7 continued localization
haftmann
parents: 25917
diff changeset
  1057
37a7eb7fd5f7 continued localization
haftmann
parents: 25917
diff changeset
  1058
lemma mult_less_le_imp_less:
37a7eb7fd5f7 continued localization
haftmann
parents: 25917
diff changeset
  1059
  assumes "a < b" and "c \<le> d" and "0 \<le> a" and "0 < c"
37a7eb7fd5f7 continued localization
haftmann
parents: 25917
diff changeset
  1060
  shows "a * c < b * d"
37a7eb7fd5f7 continued localization
haftmann
parents: 25917
diff changeset
  1061
  using assms apply (subgoal_tac "a * c < b * c")
37a7eb7fd5f7 continued localization
haftmann
parents: 25917
diff changeset
  1062
  apply (erule less_le_trans)
37a7eb7fd5f7 continued localization
haftmann
parents: 25917
diff changeset
  1063
  apply (erule mult_left_mono)
37a7eb7fd5f7 continued localization
haftmann
parents: 25917
diff changeset
  1064
  apply simp
37a7eb7fd5f7 continued localization
haftmann
parents: 25917
diff changeset
  1065
  apply (erule mult_strict_right_mono)
37a7eb7fd5f7 continued localization
haftmann
parents: 25917
diff changeset
  1066
  apply assumption
37a7eb7fd5f7 continued localization
haftmann
parents: 25917
diff changeset
  1067
  done
37a7eb7fd5f7 continued localization
haftmann
parents: 25917
diff changeset
  1068
37a7eb7fd5f7 continued localization
haftmann
parents: 25917
diff changeset
  1069
lemma mult_le_less_imp_less:
37a7eb7fd5f7 continued localization
haftmann
parents: 25917
diff changeset
  1070
  assumes "a \<le> b" and "c < d" and "0 < a" and "0 \<le> c"
37a7eb7fd5f7 continued localization
haftmann
parents: 25917
diff changeset
  1071
  shows "a * c < b * d"
37a7eb7fd5f7 continued localization
haftmann
parents: 25917
diff changeset
  1072
  using assms apply (subgoal_tac "a * c \<le> b * c")
37a7eb7fd5f7 continued localization
haftmann
parents: 25917
diff changeset
  1073
  apply (erule le_less_trans)
37a7eb7fd5f7 continued localization
haftmann
parents: 25917
diff changeset
  1074
  apply (erule mult_strict_left_mono)
37a7eb7fd5f7 continued localization
haftmann
parents: 25917
diff changeset
  1075
  apply simp
37a7eb7fd5f7 continued localization
haftmann
parents: 25917
diff changeset
  1076
  apply (erule mult_right_mono)
37a7eb7fd5f7 continued localization
haftmann
parents: 25917
diff changeset
  1077
  apply simp
37a7eb7fd5f7 continued localization
haftmann
parents: 25917
diff changeset
  1078
  done
37a7eb7fd5f7 continued localization
haftmann
parents: 25917
diff changeset
  1079
25230
022029099a83 continued localization
haftmann
parents: 25193
diff changeset
  1080
end
022029099a83 continued localization
haftmann
parents: 25193
diff changeset
  1081
35097
4554bb2abfa3 dropped last occurence of the linlinordered accident
haftmann
parents: 35092
diff changeset
  1082
class linordered_semiring_1_strict = linordered_semiring_strict + semiring_1
36622
e393a91f86df Generalize swap_inj_on; add simps for Times; add Ex_list_of_length, log_inj; Added missing locale edges for linordered semiring with 1.
hoelzl
parents: 36348
diff changeset
  1083
begin
e393a91f86df Generalize swap_inj_on; add simps for Times; add Ex_list_of_length, log_inj; Added missing locale edges for linordered semiring with 1.
hoelzl
parents: 36348
diff changeset
  1084
e393a91f86df Generalize swap_inj_on; add simps for Times; add Ex_list_of_length, log_inj; Added missing locale edges for linordered semiring with 1.
hoelzl
parents: 36348
diff changeset
  1085
subclass linordered_semiring_1 ..
e393a91f86df Generalize swap_inj_on; add simps for Times; add Ex_list_of_length, log_inj; Added missing locale edges for linordered semiring with 1.
hoelzl
parents: 36348
diff changeset
  1086
e393a91f86df Generalize swap_inj_on; add simps for Times; add Ex_list_of_length, log_inj; Added missing locale edges for linordered semiring with 1.
hoelzl
parents: 36348
diff changeset
  1087
lemma convex_bound_lt:
e393a91f86df Generalize swap_inj_on; add simps for Times; add Ex_list_of_length, log_inj; Added missing locale edges for linordered semiring with 1.
hoelzl
parents: 36348
diff changeset
  1088
  assumes "x < a" "y < a" "0 \<le> u" "0 \<le> v" "u + v = 1"
e393a91f86df Generalize swap_inj_on; add simps for Times; add Ex_list_of_length, log_inj; Added missing locale edges for linordered semiring with 1.
hoelzl
parents: 36348
diff changeset
  1089
  shows "u * x + v * y < a"
e393a91f86df Generalize swap_inj_on; add simps for Times; add Ex_list_of_length, log_inj; Added missing locale edges for linordered semiring with 1.
hoelzl
parents: 36348
diff changeset
  1090
proof -
e393a91f86df Generalize swap_inj_on; add simps for Times; add Ex_list_of_length, log_inj; Added missing locale edges for linordered semiring with 1.
hoelzl
parents: 36348
diff changeset
  1091
  from assms have "u * x + v * y < u * a + v * a"
e393a91f86df Generalize swap_inj_on; add simps for Times; add Ex_list_of_length, log_inj; Added missing locale edges for linordered semiring with 1.
hoelzl
parents: 36348
diff changeset
  1092
    by (cases "u = 0")
e393a91f86df Generalize swap_inj_on; add simps for Times; add Ex_list_of_length, log_inj; Added missing locale edges for linordered semiring with 1.
hoelzl
parents: 36348
diff changeset
  1093
       (auto intro!: add_less_le_mono mult_strict_left_mono mult_left_mono)
49962
a8cc904a6820 Renamed {left,right}_distrib to distrib_{right,left}.
webertj
parents: 44921
diff changeset
  1094
  thus ?thesis using assms unfolding distrib_right[symmetric] by simp
36622
e393a91f86df Generalize swap_inj_on; add simps for Times; add Ex_list_of_length, log_inj; Added missing locale edges for linordered semiring with 1.
hoelzl
parents: 36348
diff changeset
  1095
qed
e393a91f86df Generalize swap_inj_on; add simps for Times; add Ex_list_of_length, log_inj; Added missing locale edges for linordered semiring with 1.
hoelzl
parents: 36348
diff changeset
  1096
e393a91f86df Generalize swap_inj_on; add simps for Times; add Ex_list_of_length, log_inj; Added missing locale edges for linordered semiring with 1.
hoelzl
parents: 36348
diff changeset
  1097
end
33319
74f0dcc0b5fb moved algebraic classes to Ring_and_Field
haftmann
parents: 32960
diff changeset
  1098
60562
24af00b010cf Amalgamation of the class comm_semiring_1_diff_distrib into comm_semiring_1_cancel. Moving axiom le_add_diff_inverse2 from semiring_numeral_div to linordered_semidom.
paulson <lp15@cam.ac.uk>
parents: 60529
diff changeset
  1099
class ordered_comm_semiring = comm_semiring_0 + ordered_ab_semigroup_add +
38642
8fa437809c67 dropped type classes mult_mono and mult_mono1; tuned names of technical rule duplicates
haftmann
parents: 37767
diff changeset
  1100
  assumes comm_mult_left_mono: "a \<le> b \<Longrightarrow> 0 \<le> c \<Longrightarrow> c * a \<le> c * b"
25186
f4d1ebffd025 localized further
haftmann
parents: 25152
diff changeset
  1101
begin
25152
bfde2f8c0f63 partially localized
haftmann
parents: 25078
diff changeset
  1102
35028
108662d50512 more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents: 34146
diff changeset
  1103
subclass ordered_semiring
28823
dcbef866c9e2 tuned unfold_locales invocation
haftmann
parents: 28559
diff changeset
  1104
proof
21199
2d83f93c3580 * Added annihilation axioms ("x * 0 = 0") to axclass semiring_0.
krauss
parents: 20633
diff changeset
  1105
  fix a b c :: 'a
23550
d4f1d6ef119c convert instance proofs to Isar style
huffman
parents: 23544
diff changeset
  1106
  assume "a \<le> b" "0 \<le> c"
38642
8fa437809c67 dropped type classes mult_mono and mult_mono1; tuned names of technical rule duplicates
haftmann
parents: 37767
diff changeset
  1107
  thus "c * a \<le> c * b" by (rule comm_mult_left_mono)
57512
cc97b347b301 reduced name variants for assoc and commute on plus and mult
haftmann
parents: 56544
diff changeset
  1108
  thus "a * c \<le> b * c" by (simp only: mult.commute)
21199
2d83f93c3580 * Added annihilation axioms ("x * 0 = 0") to axclass semiring_0.
krauss
parents: 20633
diff changeset
  1109
qed
14265
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
  1110
25267
1f745c599b5c proper reinitialisation after subclass
haftmann
parents: 25238
diff changeset
  1111
end
1f745c599b5c proper reinitialisation after subclass
haftmann
parents: 25238
diff changeset
  1112
38642
8fa437809c67 dropped type classes mult_mono and mult_mono1; tuned names of technical rule duplicates
haftmann
parents: 37767
diff changeset
  1113
class ordered_cancel_comm_semiring = ordered_comm_semiring + cancel_comm_monoid_add
25267
1f745c599b5c proper reinitialisation after subclass
haftmann
parents: 25238
diff changeset
  1114
begin
14265
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
  1115
38642
8fa437809c67 dropped type classes mult_mono and mult_mono1; tuned names of technical rule duplicates
haftmann
parents: 37767
diff changeset
  1116
subclass comm_semiring_0_cancel ..
35028
108662d50512 more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents: 34146
diff changeset
  1117
subclass ordered_comm_semiring ..
108662d50512 more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents: 34146
diff changeset
  1118
subclass ordered_cancel_semiring ..
25267
1f745c599b5c proper reinitialisation after subclass
haftmann
parents: 25238
diff changeset
  1119
1f745c599b5c proper reinitialisation after subclass
haftmann
parents: 25238
diff changeset
  1120
end
1f745c599b5c proper reinitialisation after subclass
haftmann
parents: 25238
diff changeset
  1121
35028
108662d50512 more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents: 34146
diff changeset
  1122
class linordered_comm_semiring_strict = comm_semiring_0 + linordered_cancel_ab_semigroup_add +
38642
8fa437809c67 dropped type classes mult_mono and mult_mono1; tuned names of technical rule duplicates
haftmann
parents: 37767
diff changeset
  1123
  assumes comm_mult_strict_left_mono: "a < b \<Longrightarrow> 0 < c \<Longrightarrow> c * a < c * b"
25267
1f745c599b5c proper reinitialisation after subclass
haftmann
parents: 25238
diff changeset
  1124
begin
1f745c599b5c proper reinitialisation after subclass
haftmann
parents: 25238
diff changeset
  1125
35043
07dbdf60d5ad dropped accidental duplication of "lin" prefix from cs. 108662d50512
haftmann
parents: 35032
diff changeset
  1126
subclass linordered_semiring_strict
28823
dcbef866c9e2 tuned unfold_locales invocation
haftmann
parents: 28559
diff changeset
  1127
proof
23550
d4f1d6ef119c convert instance proofs to Isar style
huffman
parents: 23544
diff changeset
  1128
  fix a b c :: 'a
d4f1d6ef119c convert instance proofs to Isar style
huffman
parents: 23544
diff changeset
  1129
  assume "a < b" "0 < c"
38642
8fa437809c67 dropped type classes mult_mono and mult_mono1; tuned names of technical rule duplicates
haftmann
parents: 37767
diff changeset
  1130
  thus "c * a < c * b" by (rule comm_mult_strict_left_mono)
57512
cc97b347b301 reduced name variants for assoc and commute on plus and mult
haftmann
parents: 56544
diff changeset
  1131
  thus "a * c < b * c" by (simp only: mult.commute)
23550
d4f1d6ef119c convert instance proofs to Isar style
huffman
parents: 23544
diff changeset
  1132
qed
14272
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
  1133
35028
108662d50512 more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents: 34146
diff changeset
  1134
subclass ordered_cancel_comm_semiring
28823
dcbef866c9e2 tuned unfold_locales invocation
haftmann
parents: 28559
diff changeset
  1135
proof
23550
d4f1d6ef119c convert instance proofs to Isar style
huffman
parents: 23544
diff changeset
  1136
  fix a b c :: 'a
d4f1d6ef119c convert instance proofs to Isar style
huffman
parents: 23544
diff changeset
  1137
  assume "a \<le> b" "0 \<le> c"
d4f1d6ef119c convert instance proofs to Isar style
huffman
parents: 23544
diff changeset
  1138
  thus "c * a \<le> c * b"
25186
f4d1ebffd025 localized further
haftmann
parents: 25152
diff changeset
  1139
    unfolding le_less
26193
37a7eb7fd5f7 continued localization
haftmann
parents: 25917
diff changeset
  1140
    using mult_strict_left_mono by (cases "c = 0") auto
23550
d4f1d6ef119c convert instance proofs to Isar style
huffman
parents: 23544
diff changeset
  1141
qed
14272
5efbb548107d Tidying of the integer development; towards removing the
paulson
parents: 14270
diff changeset
  1142
25267
1f745c599b5c proper reinitialisation after subclass
haftmann
parents: 25238
diff changeset
  1143
end
25230
022029099a83 continued localization
haftmann
parents: 25193
diff changeset
  1144
60562
24af00b010cf Amalgamation of the class comm_semiring_1_diff_distrib into comm_semiring_1_cancel. Moving axiom le_add_diff_inverse2 from semiring_numeral_div to linordered_semidom.
paulson <lp15@cam.ac.uk>
parents: 60529
diff changeset
  1145
class ordered_ring = ring + ordered_cancel_semiring
25267
1f745c599b5c proper reinitialisation after subclass
haftmann
parents: 25238
diff changeset
  1146
begin
25230
022029099a83 continued localization
haftmann
parents: 25193
diff changeset
  1147
35028
108662d50512 more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents: 34146
diff changeset
  1148
subclass ordered_ab_group_add ..
14270
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
  1149
25230
022029099a83 continued localization
haftmann
parents: 25193
diff changeset
  1150
lemma less_add_iff1:
022029099a83 continued localization
haftmann
parents: 25193
diff changeset
  1151
  "a * e + c < b * e + d \<longleftrightarrow> (a - b) * e + c < d"
29667
53103fc8ffa3 Replaced group_ and ring_simps by algebra_simps;
nipkow
parents: 29465
diff changeset
  1152
by (simp add: algebra_simps)
25230
022029099a83 continued localization
haftmann
parents: 25193
diff changeset
  1153
022029099a83 continued localization
haftmann
parents: 25193
diff changeset
  1154
lemma less_add_iff2:
022029099a83 continued localization
haftmann
parents: 25193
diff changeset
  1155
  "a * e + c < b * e + d \<longleftrightarrow> c < (b - a) * e + d"
29667
53103fc8ffa3 Replaced group_ and ring_simps by algebra_simps;
nipkow
parents: 29465
diff changeset
  1156
by (simp add: algebra_simps)
25230
022029099a83 continued localization
haftmann
parents: 25193
diff changeset
  1157
022029099a83 continued localization
haftmann
parents: 25193
diff changeset
  1158
lemma le_add_iff1:
022029099a83 continued localization
haftmann
parents: 25193
diff changeset
  1159
  "a * e + c \<le> b * e + d \<longleftrightarrow> (a - b) * e + c \<le> d"
29667
53103fc8ffa3 Replaced group_ and ring_simps by algebra_simps;
nipkow
parents: 29465
diff changeset
  1160
by (simp add: algebra_simps)
25230
022029099a83 continued localization
haftmann
parents: 25193
diff changeset
  1161
022029099a83 continued localization
haftmann
parents: 25193
diff changeset
  1162
lemma le_add_iff2:
022029099a83 continued localization
haftmann
parents: 25193
diff changeset
  1163
  "a * e + c \<le> b * e + d \<longleftrightarrow> c \<le> (b - a) * e + d"
29667
53103fc8ffa3 Replaced group_ and ring_simps by algebra_simps;
nipkow
parents: 29465
diff changeset
  1164
by (simp add: algebra_simps)
25230
022029099a83 continued localization
haftmann
parents: 25193
diff changeset
  1165
022029099a83 continued localization
haftmann
parents: 25193
diff changeset
  1166
lemma mult_left_mono_neg:
022029099a83 continued localization
haftmann
parents: 25193
diff changeset
  1167
  "b \<le> a \<Longrightarrow> c \<le> 0 \<Longrightarrow> c * a \<le> c * b"
36301
72f4d079ebf8 more localization; factored out lemmas for division_ring
haftmann
parents: 35828
diff changeset
  1168
  apply (drule mult_left_mono [of _ _ "- c"])
35216
7641e8d831d2 get rid of many duplicate simp rule warnings
huffman
parents: 35097
diff changeset
  1169
  apply simp_all
25230
022029099a83 continued localization
haftmann
parents: 25193
diff changeset
  1170
  done
022029099a83 continued localization
haftmann
parents: 25193
diff changeset
  1171
022029099a83 continued localization
haftmann
parents: 25193
diff changeset
  1172
lemma mult_right_mono_neg:
022029099a83 continued localization
haftmann
parents: 25193
diff changeset
  1173
  "b \<le> a \<Longrightarrow> c \<le> 0 \<Longrightarrow> a * c \<le> b * c"
36301
72f4d079ebf8 more localization; factored out lemmas for division_ring
haftmann
parents: 35828
diff changeset
  1174
  apply (drule mult_right_mono [of _ _ "- c"])
35216
7641e8d831d2 get rid of many duplicate simp rule warnings
huffman
parents: 35097
diff changeset
  1175
  apply simp_all
25230
022029099a83 continued localization
haftmann
parents: 25193
diff changeset
  1176
  done
022029099a83 continued localization
haftmann
parents: 25193
diff changeset
  1177
30692
44ea10bc07a7 clean up proofs of sign rules for multiplication; add list of lemmas mult_sign_intros
huffman
parents: 30650
diff changeset
  1178
lemma mult_nonpos_nonpos: "a \<le> 0 \<Longrightarrow> b \<le> 0 \<Longrightarrow> 0 \<le> a * b"
36301
72f4d079ebf8 more localization; factored out lemmas for division_ring
haftmann
parents: 35828
diff changeset
  1179
using mult_right_mono_neg [of a 0 b] by simp
25230
022029099a83 continued localization
haftmann
parents: 25193
diff changeset
  1180
022029099a83 continued localization
haftmann
parents: 25193
diff changeset
  1181
lemma split_mult_pos_le:
022029099a83 continued localization
haftmann
parents: 25193
diff changeset
  1182
  "(0 \<le> a \<and> 0 \<le> b) \<or> (a \<le> 0 \<and> b \<le> 0) \<Longrightarrow> 0 \<le> a * b"
56536
aefb4a8da31f made mult_nonneg_nonneg a simp rule
nipkow
parents: 56480
diff changeset
  1183
by (auto simp add: mult_nonpos_nonpos)
25186
f4d1ebffd025 localized further
haftmann
parents: 25152
diff changeset
  1184
f4d1ebffd025 localized further
haftmann
parents: 25152
diff changeset
  1185
end
14270
342451d763f9 More re-organising of numerical theorems
paulson
parents: 14269
diff changeset
  1186
35028
108662d50512 more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents: 34146
diff changeset
  1187
class linordered_ring = ring + linordered_semiring + linordered_ab_group_add + abs_if
25304
7491c00f0915 removed subclass edge ordered_ring < lordered_ring
haftmann
parents: 25267
diff changeset
  1188
begin
7491c00f0915 removed subclass edge ordered_ring < lordered_ring
haftmann
parents: 25267
diff changeset
  1189
35028
108662d50512 more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents: 34146
diff changeset
  1190
subclass ordered_ring ..
108662d50512 more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents: 34146
diff changeset
  1191
108662d50512 more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents: 34146
diff changeset
  1192
subclass ordered_ab_group_add_abs
28823
dcbef866c9e2 tuned unfold_locales invocation
haftmann
parents: 28559
diff changeset
  1193
proof
25304
7491c00f0915 removed subclass edge ordered_ring < lordered_ring
haftmann
parents: 25267
diff changeset
  1194
  fix a b
7491c00f0915 removed subclass edge ordered_ring < lordered_ring
haftmann
parents: 25267
diff changeset
  1195
  show "\<bar>a + b\<bar> \<le> \<bar>a\<bar> + \<bar>b\<bar>"
54230
b1d955791529 more simplification rules on unary and binary minus
haftmann
parents: 54225
diff changeset
  1196
    by (auto simp add: abs_if not_le not_less algebra_simps simp del: add.commute dest: add_neg_neg add_nonneg_nonneg)
35216
7641e8d831d2 get rid of many duplicate simp rule warnings
huffman
parents: 35097
diff changeset
  1197
qed (auto simp add: abs_if)
25304
7491c00f0915 removed subclass edge ordered_ring < lordered_ring
haftmann
parents: 25267
diff changeset
  1198
35631
0b8a5fd339ab generalize some lemmas from class linordered_ring_strict to linordered_ring
huffman
parents: 35302
diff changeset
  1199
lemma zero_le_square [simp]: "0 \<le> a * a"
0b8a5fd339ab generalize some lemmas from class linordered_ring_strict to linordered_ring
huffman
parents: 35302
diff changeset
  1200
  using linear [of 0 a]
56536
aefb4a8da31f made mult_nonneg_nonneg a simp rule
nipkow
parents: 56480
diff changeset
  1201
  by (auto simp add: mult_nonpos_nonpos)
35631
0b8a5fd339ab generalize some lemmas from class linordered_ring_strict to linordered_ring
huffman
parents: 35302
diff changeset
  1202
0b8a5fd339ab generalize some lemmas from class linordered_ring_strict to linordered_ring
huffman
parents: 35302
diff changeset
  1203
lemma not_square_less_zero [simp]: "\<not> (a * a < 0)"
0b8a5fd339ab generalize some lemmas from class linordered_ring_strict to linordered_ring
huffman
parents: 35302
diff changeset
  1204
  by (simp add: not_less)
0b8a5fd339ab generalize some lemmas from class linordered_ring_strict to linordered_ring
huffman
parents: 35302
diff changeset
  1205
25304
7491c00f0915 removed subclass edge ordered_ring < lordered_ring
haftmann
parents: 25267
diff changeset
  1206
end
23521
195fe3fe2831 added ordered_ring and ordered_semiring
obua
parents: 23496
diff changeset
  1207
35043
07dbdf60d5ad dropped accidental duplication of "lin" prefix from cs. 108662d50512
haftmann
parents: 35032
diff changeset
  1208
class linordered_ring_strict = ring + linordered_semiring_strict
25304
7491c00f0915 removed subclass edge ordered_ring < lordered_ring
haftmann
parents: 25267
diff changeset
  1209
  + ordered_ab_group_add + abs_if
25230
022029099a83 continued localization
haftmann
parents: 25193
diff changeset
  1210
begin
14348
744c868ee0b7 Defining the type class "ringpower" and deleting superseded theorems for
paulson
parents: 14341
diff changeset
  1211
35028
108662d50512 more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents: 34146
diff changeset
  1212
subclass linordered_ring ..
25304
7491c00f0915 removed subclass edge ordered_ring < lordered_ring
haftmann
parents: 25267
diff changeset
  1213
30692
44ea10bc07a7 clean up proofs of sign rules for multiplication; add list of lemmas mult_sign_intros
huffman
parents: 30650
diff changeset
  1214
lemma mult_strict_left_mono_neg: "b < a \<Longrightarrow> c < 0 \<Longrightarrow> c * a < c * b"
44ea10bc07a7 clean up proofs of sign rules for multiplication; add list of lemmas mult_sign_intros
huffman
parents: 30650
diff changeset
  1215
using mult_strict_left_mono [of b a "- c"] by simp
44ea10bc07a7 clean up proofs of sign rules for multiplication; add list of lemmas mult_sign_intros
huffman
parents: 30650
diff changeset
  1216
44ea10bc07a7 clean up proofs of sign rules for multiplication; add list of lemmas mult_sign_intros
huffman
parents: 30650
diff changeset
  1217
lemma mult_strict_right_mono_neg: "b < a \<Longrightarrow> c < 0 \<Longrightarrow> a * c < b * c"
44ea10bc07a7 clean up proofs of sign rules for multiplication; add list of lemmas mult_sign_intros
huffman
parents: 30650
diff changeset
  1218
using mult_strict_right_mono [of b a "- c"] by simp
44ea10bc07a7 clean up proofs of sign rules for multiplication; add list of lemmas mult_sign_intros
huffman
parents: 30650
diff changeset
  1219
44ea10bc07a7 clean up proofs of sign rules for multiplication; add list of lemmas mult_sign_intros
huffman
parents: 30650
diff changeset
  1220
lemma mult_neg_neg: "a < 0 \<Longrightarrow> b < 0 \<Longrightarrow> 0 < a * b"
36301
72f4d079ebf8 more localization; factored out lemmas for division_ring
haftmann
parents: 35828
diff changeset
  1221
using mult_strict_right_mono_neg [of a 0 b] by simp
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1222
25917
d6c920623afc further localization
haftmann
parents: 25762
diff changeset
  1223
subclass ring_no_zero_divisors
28823
dcbef866c9e2 tuned unfold_locales invocation
haftmann
parents: 28559
diff changeset
  1224
proof
25917
d6c920623afc further localization
haftmann
parents: 25762
diff changeset
  1225
  fix a b
d6c920623afc further localization
haftmann
parents: 25762
diff changeset
  1226
  assume "a \<noteq> 0" then have A: "a < 0 \<or> 0 < a" by (simp add: neq_iff)
d6c920623afc further localization
haftmann
parents: 25762
diff changeset
  1227
  assume "b \<noteq> 0" then have B: "b < 0 \<or> 0 < b" by (simp add: neq_iff)
d6c920623afc further localization
haftmann
parents: 25762
diff changeset
  1228
  have "a * b < 0 \<or> 0 < a * b"
d6c920623afc further localization
haftmann
parents: 25762
diff changeset
  1229
  proof (cases "a < 0")
d6c920623afc further localization
haftmann
parents: 25762
diff changeset
  1230
    case True note A' = this
d6c920623afc further localization
haftmann
parents: 25762
diff changeset
  1231
    show ?thesis proof (cases "b < 0")
d6c920623afc further localization
haftmann
parents: 25762
diff changeset
  1232
      case True with A'
d6c920623afc further localization
haftmann
parents: 25762
diff changeset
  1233
      show ?thesis by (auto dest: mult_neg_neg)
d6c920623afc further localization
haftmann
parents: 25762
diff changeset
  1234
    next
d6c920623afc further localization
haftmann
parents: 25762
diff changeset
  1235
      case False with B have "0 < b" by auto
d6c920623afc further localization
haftmann
parents: 25762
diff changeset
  1236
      with A' show ?thesis by (auto dest: mult_strict_right_mono)
d6c920623afc further localization
haftmann
parents: 25762
diff changeset
  1237
    qed
d6c920623afc further localization
haftmann
parents: 25762
diff changeset
  1238
  next
d6c920623afc further localization
haftmann
parents: 25762
diff changeset
  1239
    case False with A have A': "0 < a" by auto
d6c920623afc further localization
haftmann
parents: 25762
diff changeset
  1240
    show ?thesis proof (cases "b < 0")
d6c920623afc further localization
haftmann
parents: 25762
diff changeset
  1241
      case True with A'
d6c920623afc further localization
haftmann
parents: 25762
diff changeset
  1242
      show ?thesis by (auto dest: mult_strict_right_mono_neg)
d6c920623afc further localization
haftmann
parents: 25762
diff changeset
  1243
    next
d6c920623afc further localization
haftmann
parents: 25762
diff changeset
  1244
      case False with B have "0 < b" by auto
56544
b60d5d119489 made mult_pos_pos a simp rule
nipkow
parents: 56536
diff changeset
  1245
      with A' show ?thesis by auto
25917
d6c920623afc further localization
haftmann
parents: 25762
diff changeset
  1246
    qed
d6c920623afc further localization
haftmann
parents: 25762
diff changeset
  1247
  qed
d6c920623afc further localization
haftmann
parents: 25762
diff changeset
  1248
  then show "a * b \<noteq> 0" by (simp add: neq_iff)
d6c920623afc further localization
haftmann
parents: 25762
diff changeset
  1249
qed
25304
7491c00f0915 removed subclass edge ordered_ring < lordered_ring
haftmann
parents: 25267
diff changeset
  1250
56480
093ea91498e6 field_simps: better support for negation and division, and power
hoelzl
parents: 56217
diff changeset
  1251
lemma zero_less_mult_iff: "0 < a * b \<longleftrightarrow> 0 < a \<and> 0 < b \<or> a < 0 \<and> b < 0"
093ea91498e6 field_simps: better support for negation and division, and power
hoelzl
parents: 56217
diff changeset
  1252
  by (cases a 0 b 0 rule: linorder_cases[case_product linorder_cases])
56544
b60d5d119489 made mult_pos_pos a simp rule
nipkow
parents: 56536
diff changeset
  1253
     (auto simp add: mult_neg_neg not_less le_less dest: zero_less_mult_pos zero_less_mult_pos2)
22990
775e9de3db48 added classes ring_no_zero_divisors and dom (non-commutative version of idom);
huffman
parents: 22987
diff changeset
  1254
56480
093ea91498e6 field_simps: better support for negation and division, and power
hoelzl
parents: 56217
diff changeset
  1255
lemma zero_le_mult_iff: "0 \<le> a * b \<longleftrightarrow> 0 \<le> a \<and> 0 \<le> b \<or> a \<le> 0 \<and> b \<le> 0"
093ea91498e6 field_simps: better support for negation and division, and power
hoelzl
parents: 56217
diff changeset
  1256
  by (auto simp add: eq_commute [of 0] le_less not_less zero_less_mult_iff)
14265
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
  1257
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
  1258
lemma mult_less_0_iff:
25917
d6c920623afc further localization
haftmann
parents: 25762
diff changeset
  1259
  "a * b < 0 \<longleftrightarrow> 0 < a \<and> b < 0 \<or> a < 0 \<and> 0 < b"
35216
7641e8d831d2 get rid of many duplicate simp rule warnings
huffman
parents: 35097
diff changeset
  1260
  apply (insert zero_less_mult_iff [of "-a" b])
7641e8d831d2 get rid of many duplicate simp rule warnings
huffman
parents: 35097
diff changeset
  1261
  apply force
25917
d6c920623afc further localization
haftmann
parents: 25762
diff changeset
  1262
  done
14265
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
  1263
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
  1264
lemma mult_le_0_iff:
25917
d6c920623afc further localization
haftmann
parents: 25762
diff changeset
  1265
  "a * b \<le> 0 \<longleftrightarrow> 0 \<le> a \<and> b \<le> 0 \<or> a \<le> 0 \<and> 0 \<le> b"
60562
24af00b010cf Amalgamation of the class comm_semiring_1_diff_distrib into comm_semiring_1_cancel. Moving axiom le_add_diff_inverse2 from semiring_numeral_div to linordered_semidom.
paulson <lp15@cam.ac.uk>
parents: 60529
diff changeset
  1266
  apply (insert zero_le_mult_iff [of "-a" b])
35216
7641e8d831d2 get rid of many duplicate simp rule warnings
huffman
parents: 35097
diff changeset
  1267
  apply force
25917
d6c920623afc further localization
haftmann
parents: 25762
diff changeset
  1268
  done
d6c920623afc further localization
haftmann
parents: 25762
diff changeset
  1269
26193
37a7eb7fd5f7 continued localization
haftmann
parents: 25917
diff changeset
  1270
text{*Cancellation laws for @{term "c*a < c*b"} and @{term "a*c < b*c"},
37a7eb7fd5f7 continued localization
haftmann
parents: 25917
diff changeset
  1271
   also with the relations @{text "\<le>"} and equality.*}
37a7eb7fd5f7 continued localization
haftmann
parents: 25917
diff changeset
  1272
37a7eb7fd5f7 continued localization
haftmann
parents: 25917
diff changeset
  1273
text{*These ``disjunction'' versions produce two cases when the comparison is
37a7eb7fd5f7 continued localization
haftmann
parents: 25917
diff changeset
  1274
 an assumption, but effectively four when the comparison is a goal.*}
37a7eb7fd5f7 continued localization
haftmann
parents: 25917
diff changeset
  1275
37a7eb7fd5f7 continued localization
haftmann
parents: 25917
diff changeset
  1276
lemma mult_less_cancel_right_disj:
37a7eb7fd5f7 continued localization
haftmann
parents: 25917
diff changeset
  1277
  "a * c < b * c \<longleftrightarrow> 0 < c \<and> a < b \<or> c < 0 \<and>  b < a"
37a7eb7fd5f7 continued localization
haftmann
parents: 25917
diff changeset
  1278
  apply (cases "c = 0")
60562
24af00b010cf Amalgamation of the class comm_semiring_1_diff_distrib into comm_semiring_1_cancel. Moving axiom le_add_diff_inverse2 from semiring_numeral_div to linordered_semidom.
paulson <lp15@cam.ac.uk>
parents: 60529
diff changeset
  1279
  apply (auto simp add: neq_iff mult_strict_right_mono
26193
37a7eb7fd5f7 continued localization
haftmann
parents: 25917
diff changeset
  1280
                      mult_strict_right_mono_neg)
60562
24af00b010cf Amalgamation of the class comm_semiring_1_diff_distrib into comm_semiring_1_cancel. Moving axiom le_add_diff_inverse2 from semiring_numeral_div to linordered_semidom.
paulson <lp15@cam.ac.uk>
parents: 60529
diff changeset
  1281
  apply (auto simp add: not_less
26193
37a7eb7fd5f7 continued localization
haftmann
parents: 25917
diff changeset
  1282
                      not_le [symmetric, of "a*c"]
37a7eb7fd5f7 continued localization
haftmann
parents: 25917
diff changeset
  1283
                      not_le [symmetric, of a])
37a7eb7fd5f7 continued localization
haftmann
parents: 25917
diff changeset
  1284
  apply (erule_tac [!] notE)
60562
24af00b010cf Amalgamation of the class comm_semiring_1_diff_distrib into comm_semiring_1_cancel. Moving axiom le_add_diff_inverse2 from semiring_numeral_div to linordered_semidom.
paulson <lp15@cam.ac.uk>
parents: 60529
diff changeset
  1285
  apply (auto simp add: less_imp_le mult_right_mono
26193
37a7eb7fd5f7 continued localization
haftmann
parents: 25917
diff changeset
  1286
                      mult_right_mono_neg)
37a7eb7fd5f7 continued localization
haftmann
parents: 25917
diff changeset
  1287
  done
37a7eb7fd5f7 continued localization
haftmann
parents: 25917
diff changeset
  1288
37a7eb7fd5f7 continued localization
haftmann
parents: 25917
diff changeset
  1289
lemma mult_less_cancel_left_disj:
37a7eb7fd5f7 continued localization
haftmann
parents: 25917
diff changeset
  1290
  "c * a < c * b \<longleftrightarrow> 0 < c \<and> a < b \<or> c < 0 \<and>  b < a"
37a7eb7fd5f7 continued localization
haftmann
parents: 25917
diff changeset
  1291
  apply (cases "c = 0")
60562
24af00b010cf Amalgamation of the class comm_semiring_1_diff_distrib into comm_semiring_1_cancel. Moving axiom le_add_diff_inverse2 from semiring_numeral_div to linordered_semidom.
paulson <lp15@cam.ac.uk>
parents: 60529
diff changeset
  1292
  apply (auto simp add: neq_iff mult_strict_left_mono
26193
37a7eb7fd5f7 continued localization
haftmann
parents: 25917
diff changeset
  1293
                      mult_strict_left_mono_neg)
60562
24af00b010cf Amalgamation of the class comm_semiring_1_diff_distrib into comm_semiring_1_cancel. Moving axiom le_add_diff_inverse2 from semiring_numeral_div to linordered_semidom.
paulson <lp15@cam.ac.uk>
parents: 60529
diff changeset
  1294
  apply (auto simp add: not_less
26193
37a7eb7fd5f7 continued localization
haftmann
parents: 25917
diff changeset
  1295
                      not_le [symmetric, of "c*a"]
37a7eb7fd5f7 continued localization
haftmann
parents: 25917
diff changeset
  1296
                      not_le [symmetric, of a])
37a7eb7fd5f7 continued localization
haftmann
parents: 25917
diff changeset
  1297
  apply (erule_tac [!] notE)
60562
24af00b010cf Amalgamation of the class comm_semiring_1_diff_distrib into comm_semiring_1_cancel. Moving axiom le_add_diff_inverse2 from semiring_numeral_div to linordered_semidom.
paulson <lp15@cam.ac.uk>
parents: 60529
diff changeset
  1298
  apply (auto simp add: less_imp_le mult_left_mono
26193
37a7eb7fd5f7 continued localization
haftmann
parents: 25917
diff changeset
  1299
                      mult_left_mono_neg)
37a7eb7fd5f7 continued localization
haftmann
parents: 25917
diff changeset
  1300
  done
37a7eb7fd5f7 continued localization
haftmann
parents: 25917
diff changeset
  1301
37a7eb7fd5f7 continued localization
haftmann
parents: 25917
diff changeset
  1302
text{*The ``conjunction of implication'' lemmas produce two cases when the
37a7eb7fd5f7 continued localization
haftmann
parents: 25917
diff changeset
  1303
comparison is a goal, but give four when the comparison is an assumption.*}
37a7eb7fd5f7 continued localization
haftmann
parents: 25917
diff changeset
  1304
37a7eb7fd5f7 continued localization
haftmann
parents: 25917
diff changeset
  1305
lemma mult_less_cancel_right:
37a7eb7fd5f7 continued localization
haftmann
parents: 25917
diff changeset
  1306
  "a * c < b * c \<longleftrightarrow> (0 \<le> c \<longrightarrow> a < b) \<and> (c \<le> 0 \<longrightarrow> b < a)"
37a7eb7fd5f7 continued localization
haftmann
parents: 25917
diff changeset
  1307
  using mult_less_cancel_right_disj [of a c b] by auto
37a7eb7fd5f7 continued localization
haftmann
parents: 25917
diff changeset
  1308
37a7eb7fd5f7 continued localization
haftmann
parents: 25917
diff changeset
  1309
lemma mult_less_cancel_left:
37a7eb7fd5f7 continued localization
haftmann
parents: 25917
diff changeset
  1310
  "c * a < c * b \<longleftrightarrow> (0 \<le> c \<longrightarrow> a < b) \<and> (c \<le> 0 \<longrightarrow> b < a)"
37a7eb7fd5f7 continued localization
haftmann
parents: 25917
diff changeset
  1311
  using mult_less_cancel_left_disj [of c a b] by auto
37a7eb7fd5f7 continued localization
haftmann
parents: 25917
diff changeset
  1312
37a7eb7fd5f7 continued localization
haftmann
parents: 25917
diff changeset
  1313
lemma mult_le_cancel_right:
37a7eb7fd5f7 continued localization
haftmann
parents: 25917
diff changeset
  1314
   "a * c \<le> b * c \<longleftrightarrow> (0 < c \<longrightarrow> a \<le> b) \<and> (c < 0 \<longrightarrow> b \<le> a)"
29667
53103fc8ffa3 Replaced group_ and ring_simps by algebra_simps;
nipkow
parents: 29465
diff changeset
  1315
by (simp add: not_less [symmetric] mult_less_cancel_right_disj)
26193
37a7eb7fd5f7 continued localization
haftmann
parents: 25917
diff changeset
  1316
37a7eb7fd5f7 continued localization
haftmann
parents: 25917
diff changeset
  1317
lemma mult_le_cancel_left:
37a7eb7fd5f7 continued localization
haftmann
parents: 25917
diff changeset
  1318
  "c * a \<le> c * b \<longleftrightarrow> (0 < c \<longrightarrow> a \<le> b) \<and> (c < 0 \<longrightarrow> b \<le> a)"
29667
53103fc8ffa3 Replaced group_ and ring_simps by algebra_simps;
nipkow
parents: 29465
diff changeset
  1319
by (simp add: not_less [symmetric] mult_less_cancel_left_disj)
26193
37a7eb7fd5f7 continued localization
haftmann
parents: 25917
diff changeset
  1320
30649
57753e0ec1d4 1. New cancellation simprocs for common factors in inequations
nipkow
parents: 30242
diff changeset
  1321
lemma mult_le_cancel_left_pos:
57753e0ec1d4 1. New cancellation simprocs for common factors in inequations
nipkow
parents: 30242
diff changeset
  1322
  "0 < c \<Longrightarrow> c * a \<le> c * b \<longleftrightarrow> a \<le> b"
57753e0ec1d4 1. New cancellation simprocs for common factors in inequations
nipkow
parents: 30242
diff changeset
  1323
by (auto simp: mult_le_cancel_left)
57753e0ec1d4 1. New cancellation simprocs for common factors in inequations
nipkow
parents: 30242
diff changeset
  1324
57753e0ec1d4 1. New cancellation simprocs for common factors in inequations
nipkow
parents: 30242
diff changeset
  1325
lemma mult_le_cancel_left_neg:
57753e0ec1d4 1. New cancellation simprocs for common factors in inequations
nipkow
parents: 30242
diff changeset
  1326
  "c < 0 \<Longrightarrow> c * a \<le> c * b \<longleftrightarrow> b \<le> a"
57753e0ec1d4 1. New cancellation simprocs for common factors in inequations
nipkow
parents: 30242
diff changeset
  1327
by (auto simp: mult_le_cancel_left)
57753e0ec1d4 1. New cancellation simprocs for common factors in inequations
nipkow
parents: 30242
diff changeset
  1328
57753e0ec1d4 1. New cancellation simprocs for common factors in inequations
nipkow
parents: 30242
diff changeset
  1329
lemma mult_less_cancel_left_pos:
57753e0ec1d4 1. New cancellation simprocs for common factors in inequations
nipkow
parents: 30242
diff changeset
  1330
  "0 < c \<Longrightarrow> c * a < c * b \<longleftrightarrow> a < b"
57753e0ec1d4 1. New cancellation simprocs for common factors in inequations
nipkow
parents: 30242
diff changeset
  1331
by (auto simp: mult_less_cancel_left)
57753e0ec1d4 1. New cancellation simprocs for common factors in inequations
nipkow
parents: 30242
diff changeset
  1332
57753e0ec1d4 1. New cancellation simprocs for common factors in inequations
nipkow
parents: 30242
diff changeset
  1333
lemma mult_less_cancel_left_neg:
57753e0ec1d4 1. New cancellation simprocs for common factors in inequations
nipkow
parents: 30242
diff changeset
  1334
  "c < 0 \<Longrightarrow> c * a < c * b \<longleftrightarrow> b < a"
57753e0ec1d4 1. New cancellation simprocs for common factors in inequations
nipkow
parents: 30242
diff changeset
  1335
by (auto simp: mult_less_cancel_left)
57753e0ec1d4 1. New cancellation simprocs for common factors in inequations
nipkow
parents: 30242
diff changeset
  1336
25917
d6c920623afc further localization
haftmann
parents: 25762
diff changeset
  1337
end
14265
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
  1338
30692
44ea10bc07a7 clean up proofs of sign rules for multiplication; add list of lemmas mult_sign_intros
huffman
parents: 30650
diff changeset
  1339
lemmas mult_sign_intros =
44ea10bc07a7 clean up proofs of sign rules for multiplication; add list of lemmas mult_sign_intros
huffman
parents: 30650
diff changeset
  1340
  mult_nonneg_nonneg mult_nonneg_nonpos
44ea10bc07a7 clean up proofs of sign rules for multiplication; add list of lemmas mult_sign_intros
huffman
parents: 30650
diff changeset
  1341
  mult_nonpos_nonneg mult_nonpos_nonpos
44ea10bc07a7 clean up proofs of sign rules for multiplication; add list of lemmas mult_sign_intros
huffman
parents: 30650
diff changeset
  1342
  mult_pos_pos mult_pos_neg
44ea10bc07a7 clean up proofs of sign rules for multiplication; add list of lemmas mult_sign_intros
huffman
parents: 30650
diff changeset
  1343
  mult_neg_pos mult_neg_neg
25230
022029099a83 continued localization
haftmann
parents: 25193
diff changeset
  1344
35028
108662d50512 more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents: 34146
diff changeset
  1345
class ordered_comm_ring = comm_ring + ordered_comm_semiring
25267
1f745c599b5c proper reinitialisation after subclass
haftmann
parents: 25238
diff changeset
  1346
begin
25230
022029099a83 continued localization
haftmann
parents: 25193
diff changeset
  1347
35028
108662d50512 more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents: 34146
diff changeset
  1348
subclass ordered_ring ..
108662d50512 more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents: 34146
diff changeset
  1349
subclass ordered_cancel_comm_semiring ..
25230
022029099a83 continued localization
haftmann
parents: 25193
diff changeset
  1350
25267
1f745c599b5c proper reinitialisation after subclass
haftmann
parents: 25238
diff changeset
  1351
end
25230
022029099a83 continued localization
haftmann
parents: 25193
diff changeset
  1352
59833
ab828c2c5d67 clarified no_zero_devisors: makes only sense in a semiring;
haftmann
parents: 59832
diff changeset
  1353
class linordered_semidom = semidom + linordered_comm_semiring_strict +
25230
022029099a83 continued localization
haftmann
parents: 25193
diff changeset
  1354
  assumes zero_less_one [simp]: "0 < 1"
60562
24af00b010cf Amalgamation of the class comm_semiring_1_diff_distrib into comm_semiring_1_cancel. Moving axiom le_add_diff_inverse2 from semiring_numeral_div to linordered_semidom.
paulson <lp15@cam.ac.uk>
parents: 60529
diff changeset
  1355
  assumes le_add_diff_inverse2 [simp]: "b \<le> a \<Longrightarrow> a - b + b = a"
25230
022029099a83 continued localization
haftmann
parents: 25193
diff changeset
  1356
begin
022029099a83 continued localization
haftmann
parents: 25193
diff changeset
  1357
60562
24af00b010cf Amalgamation of the class comm_semiring_1_diff_distrib into comm_semiring_1_cancel. Moving axiom le_add_diff_inverse2 from semiring_numeral_div to linordered_semidom.
paulson <lp15@cam.ac.uk>
parents: 60529
diff changeset
  1358
text {* Addition is the inverse of subtraction. *}
24af00b010cf Amalgamation of the class comm_semiring_1_diff_distrib into comm_semiring_1_cancel. Moving axiom le_add_diff_inverse2 from semiring_numeral_div to linordered_semidom.
paulson <lp15@cam.ac.uk>
parents: 60529
diff changeset
  1359
24af00b010cf Amalgamation of the class comm_semiring_1_diff_distrib into comm_semiring_1_cancel. Moving axiom le_add_diff_inverse2 from semiring_numeral_div to linordered_semidom.
paulson <lp15@cam.ac.uk>
parents: 60529
diff changeset
  1360
lemma le_add_diff_inverse [simp]: "b \<le> a \<Longrightarrow> b + (a - b) = a"
24af00b010cf Amalgamation of the class comm_semiring_1_diff_distrib into comm_semiring_1_cancel. Moving axiom le_add_diff_inverse2 from semiring_numeral_div to linordered_semidom.
paulson <lp15@cam.ac.uk>
parents: 60529
diff changeset
  1361
  by (frule le_add_diff_inverse2) (simp add: add.commute)
24af00b010cf Amalgamation of the class comm_semiring_1_diff_distrib into comm_semiring_1_cancel. Moving axiom le_add_diff_inverse2 from semiring_numeral_div to linordered_semidom.
paulson <lp15@cam.ac.uk>
parents: 60529
diff changeset
  1362
24af00b010cf Amalgamation of the class comm_semiring_1_diff_distrib into comm_semiring_1_cancel. Moving axiom le_add_diff_inverse2 from semiring_numeral_div to linordered_semidom.
paulson <lp15@cam.ac.uk>
parents: 60529
diff changeset
  1363
lemma add_diff_inverse: "~ a<b \<Longrightarrow> b + (a - b) = a"
24af00b010cf Amalgamation of the class comm_semiring_1_diff_distrib into comm_semiring_1_cancel. Moving axiom le_add_diff_inverse2 from semiring_numeral_div to linordered_semidom.
paulson <lp15@cam.ac.uk>
parents: 60529
diff changeset
  1364
  by simp
60615
e5fa1d5d3952 Useful lemmas. The theorem concerning swapping the variables in a double integral.
paulson <lp15@cam.ac.uk>
parents: 60570
diff changeset
  1365
e5fa1d5d3952 Useful lemmas. The theorem concerning swapping the variables in a double integral.
paulson <lp15@cam.ac.uk>
parents: 60570
diff changeset
  1366
lemma add_le_imp_le_diff: 
e5fa1d5d3952 Useful lemmas. The theorem concerning swapping the variables in a double integral.
paulson <lp15@cam.ac.uk>
parents: 60570
diff changeset
  1367
  shows "i + k \<le> n \<Longrightarrow> i \<le> n - k"
e5fa1d5d3952 Useful lemmas. The theorem concerning swapping the variables in a double integral.
paulson <lp15@cam.ac.uk>
parents: 60570
diff changeset
  1368
  apply (subst add_le_cancel_right [where c=k, symmetric])
e5fa1d5d3952 Useful lemmas. The theorem concerning swapping the variables in a double integral.
paulson <lp15@cam.ac.uk>
parents: 60570
diff changeset
  1369
  apply (frule le_add_diff_inverse2)
e5fa1d5d3952 Useful lemmas. The theorem concerning swapping the variables in a double integral.
paulson <lp15@cam.ac.uk>
parents: 60570
diff changeset
  1370
  apply (simp only: add.assoc [symmetric])
e5fa1d5d3952 Useful lemmas. The theorem concerning swapping the variables in a double integral.
paulson <lp15@cam.ac.uk>
parents: 60570
diff changeset
  1371
  using add_implies_diff by fastforce
e5fa1d5d3952 Useful lemmas. The theorem concerning swapping the variables in a double integral.
paulson <lp15@cam.ac.uk>
parents: 60570
diff changeset
  1372
e5fa1d5d3952 Useful lemmas. The theorem concerning swapping the variables in a double integral.
paulson <lp15@cam.ac.uk>
parents: 60570
diff changeset
  1373
lemma add_le_add_imp_diff_le: 
e5fa1d5d3952 Useful lemmas. The theorem concerning swapping the variables in a double integral.
paulson <lp15@cam.ac.uk>
parents: 60570
diff changeset
  1374
  assumes a1: "i + k \<le> n"
e5fa1d5d3952 Useful lemmas. The theorem concerning swapping the variables in a double integral.
paulson <lp15@cam.ac.uk>
parents: 60570
diff changeset
  1375
      and a2: "n \<le> j + k"
e5fa1d5d3952 Useful lemmas. The theorem concerning swapping the variables in a double integral.
paulson <lp15@cam.ac.uk>
parents: 60570
diff changeset
  1376
  shows "\<lbrakk>i + k \<le> n; n \<le> j + k\<rbrakk> \<Longrightarrow> n - k \<le> j"
e5fa1d5d3952 Useful lemmas. The theorem concerning swapping the variables in a double integral.
paulson <lp15@cam.ac.uk>
parents: 60570
diff changeset
  1377
proof -
e5fa1d5d3952 Useful lemmas. The theorem concerning swapping the variables in a double integral.
paulson <lp15@cam.ac.uk>
parents: 60570
diff changeset
  1378
  have "n - (i + k) + (i + k) = n"
e5fa1d5d3952 Useful lemmas. The theorem concerning swapping the variables in a double integral.
paulson <lp15@cam.ac.uk>
parents: 60570
diff changeset
  1379
    using a1 by simp
e5fa1d5d3952 Useful lemmas. The theorem concerning swapping the variables in a double integral.
paulson <lp15@cam.ac.uk>
parents: 60570
diff changeset
  1380
  moreover have "n - k = n - k - i + i"
e5fa1d5d3952 Useful lemmas. The theorem concerning swapping the variables in a double integral.
paulson <lp15@cam.ac.uk>
parents: 60570
diff changeset
  1381
    using a1 by (simp add: add_le_imp_le_diff)
e5fa1d5d3952 Useful lemmas. The theorem concerning swapping the variables in a double integral.
paulson <lp15@cam.ac.uk>
parents: 60570
diff changeset
  1382
  ultimately show ?thesis
e5fa1d5d3952 Useful lemmas. The theorem concerning swapping the variables in a double integral.
paulson <lp15@cam.ac.uk>
parents: 60570
diff changeset
  1383
    using a2
e5fa1d5d3952 Useful lemmas. The theorem concerning swapping the variables in a double integral.
paulson <lp15@cam.ac.uk>
parents: 60570
diff changeset
  1384
    apply (simp add: add.assoc [symmetric])
e5fa1d5d3952 Useful lemmas. The theorem concerning swapping the variables in a double integral.
paulson <lp15@cam.ac.uk>
parents: 60570
diff changeset
  1385
    apply (rule add_le_imp_le_diff [of _ k "j+k", simplified add_diff_cancel_right'])
e5fa1d5d3952 Useful lemmas. The theorem concerning swapping the variables in a double integral.
paulson <lp15@cam.ac.uk>
parents: 60570
diff changeset
  1386
    by (simp add: add.commute diff_diff_add)
e5fa1d5d3952 Useful lemmas. The theorem concerning swapping the variables in a double integral.
paulson <lp15@cam.ac.uk>
parents: 60570
diff changeset
  1387
qed
e5fa1d5d3952 Useful lemmas. The theorem concerning swapping the variables in a double integral.
paulson <lp15@cam.ac.uk>
parents: 60570
diff changeset
  1388
25230
022029099a83 continued localization
haftmann
parents: 25193
diff changeset
  1389
lemma pos_add_strict:
022029099a83 continued localization
haftmann
parents: 25193
diff changeset
  1390
  shows "0 < a \<Longrightarrow> b < c \<Longrightarrow> b < a + c"
36301
72f4d079ebf8 more localization; factored out lemmas for division_ring
haftmann
parents: 35828
diff changeset
  1391
  using add_strict_mono [of 0 a b c] by simp
25230
022029099a83 continued localization
haftmann
parents: 25193
diff changeset
  1392
26193
37a7eb7fd5f7 continued localization
haftmann
parents: 25917
diff changeset
  1393
lemma zero_le_one [simp]: "0 \<le> 1"
60562
24af00b010cf Amalgamation of the class comm_semiring_1_diff_distrib into comm_semiring_1_cancel. Moving axiom le_add_diff_inverse2 from semiring_numeral_div to linordered_semidom.
paulson <lp15@cam.ac.uk>
parents: 60529
diff changeset
  1394
by (rule zero_less_one [THEN less_imp_le])
26193
37a7eb7fd5f7 continued localization
haftmann
parents: 25917
diff changeset
  1395
37a7eb7fd5f7 continued localization
haftmann
parents: 25917
diff changeset
  1396
lemma not_one_le_zero [simp]: "\<not> 1 \<le> 0"
60562
24af00b010cf Amalgamation of the class comm_semiring_1_diff_distrib into comm_semiring_1_cancel. Moving axiom le_add_diff_inverse2 from semiring_numeral_div to linordered_semidom.
paulson <lp15@cam.ac.uk>
parents: 60529
diff changeset
  1397
by (simp add: not_le)
26193
37a7eb7fd5f7 continued localization
haftmann
parents: 25917
diff changeset
  1398
37a7eb7fd5f7 continued localization
haftmann
parents: 25917
diff changeset
  1399
lemma not_one_less_zero [simp]: "\<not> 1 < 0"
60562
24af00b010cf Amalgamation of the class comm_semiring_1_diff_distrib into comm_semiring_1_cancel. Moving axiom le_add_diff_inverse2 from semiring_numeral_div to linordered_semidom.
paulson <lp15@cam.ac.uk>
parents: 60529
diff changeset
  1400
by (simp add: not_less)
26193
37a7eb7fd5f7 continued localization
haftmann
parents: 25917
diff changeset
  1401
37a7eb7fd5f7 continued localization
haftmann
parents: 25917
diff changeset
  1402
lemma less_1_mult:
37a7eb7fd5f7 continued localization
haftmann
parents: 25917
diff changeset
  1403
  assumes "1 < m" and "1 < n"
37a7eb7fd5f7 continued localization
haftmann
parents: 25917
diff changeset
  1404
  shows "1 < m * n"
37a7eb7fd5f7 continued localization
haftmann
parents: 25917
diff changeset
  1405
  using assms mult_strict_mono [of 1 m 1 n]
60562
24af00b010cf Amalgamation of the class comm_semiring_1_diff_distrib into comm_semiring_1_cancel. Moving axiom le_add_diff_inverse2 from semiring_numeral_div to linordered_semidom.
paulson <lp15@cam.ac.uk>
parents: 60529
diff changeset
  1406
    by (simp add:  less_trans [OF zero_less_one])
26193
37a7eb7fd5f7 continued localization
haftmann
parents: 25917
diff changeset
  1407
59000
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58952
diff changeset
  1408
lemma mult_left_le: "c \<le> 1 \<Longrightarrow> 0 \<le> a \<Longrightarrow> a * c \<le> a"
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58952
diff changeset
  1409
  using mult_left_mono[of c 1 a] by simp
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58952
diff changeset
  1410
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58952
diff changeset
  1411
lemma mult_le_one: "a \<le> 1 \<Longrightarrow> 0 \<le> b \<Longrightarrow> b \<le> 1 \<Longrightarrow> a * b \<le> 1"
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58952
diff changeset
  1412
  using mult_mono[of a 1 b 1] by simp
6eb0725503fc import general theorems from AFP/Markov_Models
hoelzl
parents: 58952
diff changeset
  1413
25230
022029099a83 continued localization
haftmann
parents: 25193
diff changeset
  1414
end
022029099a83 continued localization
haftmann
parents: 25193
diff changeset
  1415
35028
108662d50512 more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents: 34146
diff changeset
  1416
class linordered_idom = comm_ring_1 +
108662d50512 more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents: 34146
diff changeset
  1417
  linordered_comm_semiring_strict + ordered_ab_group_add +
25230
022029099a83 continued localization
haftmann
parents: 25193
diff changeset
  1418
  abs_if + sgn_if
25917
d6c920623afc further localization
haftmann
parents: 25762
diff changeset
  1419
begin
d6c920623afc further localization
haftmann
parents: 25762
diff changeset
  1420
36622
e393a91f86df Generalize swap_inj_on; add simps for Times; add Ex_list_of_length, log_inj; Added missing locale edges for linordered semiring with 1.
hoelzl
parents: 36348
diff changeset
  1421
subclass linordered_semiring_1_strict ..
35043
07dbdf60d5ad dropped accidental duplication of "lin" prefix from cs. 108662d50512
haftmann
parents: 35032
diff changeset
  1422
subclass linordered_ring_strict ..
35028
108662d50512 more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents: 34146
diff changeset
  1423
subclass ordered_comm_ring ..
27516
9a5d4a8d4aac by intro_locales -> ..
huffman
parents: 26274
diff changeset
  1424
subclass idom ..
25917
d6c920623afc further localization
haftmann
parents: 25762
diff changeset
  1425
35028
108662d50512 more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents: 34146
diff changeset
  1426
subclass linordered_semidom
28823
dcbef866c9e2 tuned unfold_locales invocation
haftmann
parents: 28559
diff changeset
  1427
proof
26193
37a7eb7fd5f7 continued localization
haftmann
parents: 25917
diff changeset
  1428
  have "0 \<le> 1 * 1" by (rule zero_le_square)
37a7eb7fd5f7 continued localization
haftmann
parents: 25917
diff changeset
  1429
  thus "0 < 1" by (simp add: le_less)
60562
24af00b010cf Amalgamation of the class comm_semiring_1_diff_distrib into comm_semiring_1_cancel. Moving axiom le_add_diff_inverse2 from semiring_numeral_div to linordered_semidom.
paulson <lp15@cam.ac.uk>
parents: 60529
diff changeset
  1430
  show "\<And>b a. b \<le> a \<Longrightarrow> a - b + b = a"
24af00b010cf Amalgamation of the class comm_semiring_1_diff_distrib into comm_semiring_1_cancel. Moving axiom le_add_diff_inverse2 from semiring_numeral_div to linordered_semidom.
paulson <lp15@cam.ac.uk>
parents: 60529
diff changeset
  1431
    by simp
24af00b010cf Amalgamation of the class comm_semiring_1_diff_distrib into comm_semiring_1_cancel. Moving axiom le_add_diff_inverse2 from semiring_numeral_div to linordered_semidom.
paulson <lp15@cam.ac.uk>
parents: 60529
diff changeset
  1432
qed
25917
d6c920623afc further localization
haftmann
parents: 25762
diff changeset
  1433
35028
108662d50512 more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents: 34146
diff changeset
  1434
lemma linorder_neqE_linordered_idom:
26193
37a7eb7fd5f7 continued localization
haftmann
parents: 25917
diff changeset
  1435
  assumes "x \<noteq> y" obtains "x < y" | "y < x"
37a7eb7fd5f7 continued localization
haftmann
parents: 25917
diff changeset
  1436
  using assms by (rule neqE)
37a7eb7fd5f7 continued localization
haftmann
parents: 25917
diff changeset
  1437
26274
2bdb61a28971 continued localization
haftmann
parents: 26234
diff changeset
  1438
text {* These cancellation simprules also produce two cases when the comparison is a goal. *}
2bdb61a28971 continued localization
haftmann
parents: 26234
diff changeset
  1439
2bdb61a28971 continued localization
haftmann
parents: 26234
diff changeset
  1440
lemma mult_le_cancel_right1:
2bdb61a28971 continued localization
haftmann
parents: 26234
diff changeset
  1441
  "c \<le> b * c \<longleftrightarrow> (0 < c \<longrightarrow> 1 \<le> b) \<and> (c < 0 \<longrightarrow> b \<le> 1)"
29667
53103fc8ffa3 Replaced group_ and ring_simps by algebra_simps;
nipkow
parents: 29465
diff changeset
  1442
by (insert mult_le_cancel_right [of 1 c b], simp)
26274
2bdb61a28971 continued localization
haftmann
parents: 26234
diff changeset
  1443
2bdb61a28971 continued localization
haftmann
parents: 26234
diff changeset
  1444
lemma mult_le_cancel_right2:
2bdb61a28971 continued localization
haftmann
parents: 26234
diff changeset
  1445
  "a * c \<le> c \<longleftrightarrow> (0 < c \<longrightarrow> a \<le> 1) \<and> (c < 0 \<longrightarrow> 1 \<le> a)"
29667
53103fc8ffa3 Replaced group_ and ring_simps by algebra_simps;
nipkow
parents: 29465
diff changeset
  1446
by (insert mult_le_cancel_right [of a c 1], simp)
26274
2bdb61a28971 continued localization
haftmann
parents: 26234
diff changeset
  1447
2bdb61a28971 continued localization
haftmann
parents: 26234
diff changeset
  1448
lemma mult_le_cancel_left1:
2bdb61a28971 continued localization
haftmann
parents: 26234
diff changeset
  1449
  "c \<le> c * b \<longleftrightarrow> (0 < c \<longrightarrow> 1 \<le> b) \<and> (c < 0 \<longrightarrow> b \<le> 1)"
29667
53103fc8ffa3 Replaced group_ and ring_simps by algebra_simps;
nipkow
parents: 29465
diff changeset
  1450
by (insert mult_le_cancel_left [of c 1 b], simp)
26274
2bdb61a28971 continued localization
haftmann
parents: 26234
diff changeset
  1451
2bdb61a28971 continued localization
haftmann
parents: 26234
diff changeset
  1452
lemma mult_le_cancel_left2:
2bdb61a28971 continued localization
haftmann
parents: 26234
diff changeset
  1453
  "c * a \<le> c \<longleftrightarrow> (0 < c \<longrightarrow> a \<le> 1) \<and> (c < 0 \<longrightarrow> 1 \<le> a)"
29667
53103fc8ffa3 Replaced group_ and ring_simps by algebra_simps;
nipkow
parents: 29465
diff changeset
  1454
by (insert mult_le_cancel_left [of c a 1], simp)
26274
2bdb61a28971 continued localization
haftmann
parents: 26234
diff changeset
  1455
2bdb61a28971 continued localization
haftmann
parents: 26234
diff changeset
  1456
lemma mult_less_cancel_right1:
2bdb61a28971 continued localization
haftmann
parents: 26234
diff changeset
  1457
  "c < b * c \<longleftrightarrow> (0 \<le> c \<longrightarrow> 1 < b) \<and> (c \<le> 0 \<longrightarrow> b < 1)"
29667
53103fc8ffa3 Replaced group_ and ring_simps by algebra_simps;
nipkow
parents: 29465
diff changeset
  1458
by (insert mult_less_cancel_right [of 1 c b], simp)
26274
2bdb61a28971 continued localization
haftmann
parents: 26234
diff changeset
  1459
2bdb61a28971 continued localization
haftmann
parents: 26234
diff changeset
  1460
lemma mult_less_cancel_right2:
2bdb61a28971 continued localization
haftmann
parents: 26234
diff changeset
  1461
  "a * c < c \<longleftrightarrow> (0 \<le> c \<longrightarrow> a < 1) \<and> (c \<le> 0 \<longrightarrow> 1 < a)"
29667
53103fc8ffa3 Replaced group_ and ring_simps by algebra_simps;
nipkow
parents: 29465
diff changeset
  1462
by (insert mult_less_cancel_right [of a c 1], simp)
26274
2bdb61a28971 continued localization
haftmann
parents: 26234
diff changeset
  1463
2bdb61a28971 continued localization
haftmann
parents: 26234
diff changeset
  1464
lemma mult_less_cancel_left1:
2bdb61a28971 continued localization
haftmann
parents: 26234
diff changeset
  1465
  "c < c * b \<longleftrightarrow> (0 \<le> c \<longrightarrow> 1 < b) \<and> (c \<le> 0 \<longrightarrow> b < 1)"
29667
53103fc8ffa3 Replaced group_ and ring_simps by algebra_simps;
nipkow
parents: 29465
diff changeset
  1466
by (insert mult_less_cancel_left [of c 1 b], simp)
26274
2bdb61a28971 continued localization
haftmann
parents: 26234
diff changeset
  1467
2bdb61a28971 continued localization
haftmann
parents: 26234
diff changeset
  1468
lemma mult_less_cancel_left2:
2bdb61a28971 continued localization
haftmann
parents: 26234
diff changeset
  1469
  "c * a < c \<longleftrightarrow> (0 \<le> c \<longrightarrow> a < 1) \<and> (c \<le> 0 \<longrightarrow> 1 < a)"
29667
53103fc8ffa3 Replaced group_ and ring_simps by algebra_simps;
nipkow
parents: 29465
diff changeset
  1470
by (insert mult_less_cancel_left [of c a 1], simp)
26274
2bdb61a28971 continued localization
haftmann
parents: 26234
diff changeset
  1471
27651
16a26996c30e moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
haftmann
parents: 27516
diff changeset
  1472
lemma sgn_sgn [simp]:
16a26996c30e moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
haftmann
parents: 27516
diff changeset
  1473
  "sgn (sgn a) = sgn a"
29700
22faf21db3df added some simp rules
nipkow
parents: 29668
diff changeset
  1474
unfolding sgn_if by simp
27651
16a26996c30e moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
haftmann
parents: 27516
diff changeset
  1475
16a26996c30e moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
haftmann
parents: 27516
diff changeset
  1476
lemma sgn_0_0:
16a26996c30e moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
haftmann
parents: 27516
diff changeset
  1477
  "sgn a = 0 \<longleftrightarrow> a = 0"
29700
22faf21db3df added some simp rules
nipkow
parents: 29668
diff changeset
  1478
unfolding sgn_if by simp
27651
16a26996c30e moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
haftmann
parents: 27516
diff changeset
  1479
16a26996c30e moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
haftmann
parents: 27516
diff changeset
  1480
lemma sgn_1_pos:
16a26996c30e moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
haftmann
parents: 27516
diff changeset
  1481
  "sgn a = 1 \<longleftrightarrow> a > 0"
35216
7641e8d831d2 get rid of many duplicate simp rule warnings
huffman
parents: 35097
diff changeset
  1482
unfolding sgn_if by simp
27651
16a26996c30e moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
haftmann
parents: 27516
diff changeset
  1483
16a26996c30e moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
haftmann
parents: 27516
diff changeset
  1484
lemma sgn_1_neg:
16a26996c30e moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
haftmann
parents: 27516
diff changeset
  1485
  "sgn a = - 1 \<longleftrightarrow> a < 0"
35216
7641e8d831d2 get rid of many duplicate simp rule warnings
huffman
parents: 35097
diff changeset
  1486
unfolding sgn_if by auto
27651
16a26996c30e moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
haftmann
parents: 27516
diff changeset
  1487
29940
83b373f61d41 more default simp rules for sgn
haftmann
parents: 29925
diff changeset
  1488
lemma sgn_pos [simp]:
83b373f61d41 more default simp rules for sgn
haftmann
parents: 29925
diff changeset
  1489
  "0 < a \<Longrightarrow> sgn a = 1"
83b373f61d41 more default simp rules for sgn
haftmann
parents: 29925
diff changeset
  1490
unfolding sgn_1_pos .
83b373f61d41 more default simp rules for sgn
haftmann
parents: 29925
diff changeset
  1491
83b373f61d41 more default simp rules for sgn
haftmann
parents: 29925
diff changeset
  1492
lemma sgn_neg [simp]:
83b373f61d41 more default simp rules for sgn
haftmann
parents: 29925
diff changeset
  1493
  "a < 0 \<Longrightarrow> sgn a = - 1"
83b373f61d41 more default simp rules for sgn
haftmann
parents: 29925
diff changeset
  1494
unfolding sgn_1_neg .
83b373f61d41 more default simp rules for sgn
haftmann
parents: 29925
diff changeset
  1495
27651
16a26996c30e moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
haftmann
parents: 27516
diff changeset
  1496
lemma sgn_times:
16a26996c30e moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
haftmann
parents: 27516
diff changeset
  1497
  "sgn (a * b) = sgn a * sgn b"
29667
53103fc8ffa3 Replaced group_ and ring_simps by algebra_simps;
nipkow
parents: 29465
diff changeset
  1498
by (auto simp add: sgn_if zero_less_mult_iff)
27651
16a26996c30e moved op dvd to theory Ring_and_Field; generalized a couple of lemmas
haftmann
parents: 27516
diff changeset
  1499
36301
72f4d079ebf8 more localization; factored out lemmas for division_ring
haftmann
parents: 35828
diff changeset
  1500
lemma abs_sgn: "\<bar>k\<bar> = k * sgn k"
29700
22faf21db3df added some simp rules
nipkow
parents: 29668
diff changeset
  1501
unfolding sgn_if abs_if by auto
22faf21db3df added some simp rules
nipkow
parents: 29668
diff changeset
  1502
29940
83b373f61d41 more default simp rules for sgn
haftmann
parents: 29925
diff changeset
  1503
lemma sgn_greater [simp]:
83b373f61d41 more default simp rules for sgn
haftmann
parents: 29925
diff changeset
  1504
  "0 < sgn a \<longleftrightarrow> 0 < a"
83b373f61d41 more default simp rules for sgn
haftmann
parents: 29925
diff changeset
  1505
  unfolding sgn_if by auto
83b373f61d41 more default simp rules for sgn
haftmann
parents: 29925
diff changeset
  1506
83b373f61d41 more default simp rules for sgn
haftmann
parents: 29925
diff changeset
  1507
lemma sgn_less [simp]:
83b373f61d41 more default simp rules for sgn
haftmann
parents: 29925
diff changeset
  1508
  "sgn a < 0 \<longleftrightarrow> a < 0"
83b373f61d41 more default simp rules for sgn
haftmann
parents: 29925
diff changeset
  1509
  unfolding sgn_if by auto
83b373f61d41 more default simp rules for sgn
haftmann
parents: 29925
diff changeset
  1510
36301
72f4d079ebf8 more localization; factored out lemmas for division_ring
haftmann
parents: 35828
diff changeset
  1511
lemma abs_dvd_iff [simp]: "\<bar>m\<bar> dvd k \<longleftrightarrow> m dvd k"
29949
20a506b8256d lemmas abs_dvd_iff, dvd_abs_iff
huffman
parents: 29940
diff changeset
  1512
  by (simp add: abs_if)
20a506b8256d lemmas abs_dvd_iff, dvd_abs_iff
huffman
parents: 29940
diff changeset
  1513
36301
72f4d079ebf8 more localization; factored out lemmas for division_ring
haftmann
parents: 35828
diff changeset
  1514
lemma dvd_abs_iff [simp]: "m dvd \<bar>k\<bar> \<longleftrightarrow> m dvd k"
29949
20a506b8256d lemmas abs_dvd_iff, dvd_abs_iff
huffman
parents: 29940
diff changeset
  1515
  by (simp add: abs_if)
29653
ece6a0e9f8af added lemma abs_sng
haftmann
parents: 29465
diff changeset
  1516
33676
802f5e233e48 moved lemma from Algebra/IntRing to Ring_and_Field
nipkow
parents: 33364
diff changeset
  1517
lemma dvd_if_abs_eq:
36301
72f4d079ebf8 more localization; factored out lemmas for division_ring
haftmann
parents: 35828
diff changeset
  1518
  "\<bar>l\<bar> = \<bar>k\<bar> \<Longrightarrow> l dvd k"
33676
802f5e233e48 moved lemma from Algebra/IntRing to Ring_and_Field
nipkow
parents: 33364
diff changeset
  1519
by(subst abs_dvd_iff[symmetric]) simp
802f5e233e48 moved lemma from Algebra/IntRing to Ring_and_Field
nipkow
parents: 33364
diff changeset
  1520
55912
e12a0ab9917c fix typo
huffman
parents: 55187
diff changeset
  1521
text {* The following lemmas can be proven in more general structures, but
60562
24af00b010cf Amalgamation of the class comm_semiring_1_diff_distrib into comm_semiring_1_cancel. Moving axiom le_add_diff_inverse2 from semiring_numeral_div to linordered_semidom.
paulson <lp15@cam.ac.uk>
parents: 60529
diff changeset
  1522
are dangerous as simp rules in absence of @{thm neg_equal_zero},
54489
03ff4d1e6784 eliminiated neg_numeral in favour of - (numeral _)
haftmann
parents: 54250
diff changeset
  1523
@{thm neg_less_pos}, @{thm neg_less_eq_nonneg}. *}
03ff4d1e6784 eliminiated neg_numeral in favour of - (numeral _)
haftmann
parents: 54250
diff changeset
  1524
03ff4d1e6784 eliminiated neg_numeral in favour of - (numeral _)
haftmann
parents: 54250
diff changeset
  1525
lemma equation_minus_iff_1 [simp, no_atp]:
03ff4d1e6784 eliminiated neg_numeral in favour of - (numeral _)
haftmann
parents: 54250
diff changeset
  1526
  "1 = - a \<longleftrightarrow> a = - 1"
03ff4d1e6784 eliminiated neg_numeral in favour of - (numeral _)
haftmann
parents: 54250
diff changeset
  1527
  by (fact equation_minus_iff)
03ff4d1e6784 eliminiated neg_numeral in favour of - (numeral _)
haftmann
parents: 54250
diff changeset
  1528
03ff4d1e6784 eliminiated neg_numeral in favour of - (numeral _)
haftmann
parents: 54250
diff changeset
  1529
lemma minus_equation_iff_1 [simp, no_atp]:
03ff4d1e6784 eliminiated neg_numeral in favour of - (numeral _)
haftmann
parents: 54250
diff changeset
  1530
  "- a = 1 \<longleftrightarrow> a = - 1"
03ff4d1e6784 eliminiated neg_numeral in favour of - (numeral _)
haftmann
parents: 54250
diff changeset
  1531
  by (subst minus_equation_iff, auto)
03ff4d1e6784 eliminiated neg_numeral in favour of - (numeral _)
haftmann
parents: 54250
diff changeset
  1532
03ff4d1e6784 eliminiated neg_numeral in favour of - (numeral _)
haftmann
parents: 54250
diff changeset
  1533
lemma le_minus_iff_1 [simp, no_atp]:
03ff4d1e6784 eliminiated neg_numeral in favour of - (numeral _)
haftmann
parents: 54250
diff changeset
  1534
  "1 \<le> - b \<longleftrightarrow> b \<le> - 1"
03ff4d1e6784 eliminiated neg_numeral in favour of - (numeral _)
haftmann
parents: 54250
diff changeset
  1535
  by (fact le_minus_iff)
03ff4d1e6784 eliminiated neg_numeral in favour of - (numeral _)
haftmann
parents: 54250
diff changeset
  1536
03ff4d1e6784 eliminiated neg_numeral in favour of - (numeral _)
haftmann
parents: 54250
diff changeset
  1537
lemma minus_le_iff_1 [simp, no_atp]:
03ff4d1e6784 eliminiated neg_numeral in favour of - (numeral _)
haftmann
parents: 54250
diff changeset
  1538
  "- a \<le> 1 \<longleftrightarrow> - 1 \<le> a"
03ff4d1e6784 eliminiated neg_numeral in favour of - (numeral _)
haftmann
parents: 54250
diff changeset
  1539
  by (fact minus_le_iff)
03ff4d1e6784 eliminiated neg_numeral in favour of - (numeral _)
haftmann
parents: 54250
diff changeset
  1540
03ff4d1e6784 eliminiated neg_numeral in favour of - (numeral _)
haftmann
parents: 54250
diff changeset
  1541
lemma less_minus_iff_1 [simp, no_atp]:
03ff4d1e6784 eliminiated neg_numeral in favour of - (numeral _)
haftmann
parents: 54250
diff changeset
  1542
  "1 < - b \<longleftrightarrow> b < - 1"
03ff4d1e6784 eliminiated neg_numeral in favour of - (numeral _)
haftmann
parents: 54250
diff changeset
  1543
  by (fact less_minus_iff)
03ff4d1e6784 eliminiated neg_numeral in favour of - (numeral _)
haftmann
parents: 54250
diff changeset
  1544
03ff4d1e6784 eliminiated neg_numeral in favour of - (numeral _)
haftmann
parents: 54250
diff changeset
  1545
lemma minus_less_iff_1 [simp, no_atp]:
03ff4d1e6784 eliminiated neg_numeral in favour of - (numeral _)
haftmann
parents: 54250
diff changeset
  1546
  "- a < 1 \<longleftrightarrow> - 1 < a"
03ff4d1e6784 eliminiated neg_numeral in favour of - (numeral _)
haftmann
parents: 54250
diff changeset
  1547
  by (fact minus_less_iff)
03ff4d1e6784 eliminiated neg_numeral in favour of - (numeral _)
haftmann
parents: 54250
diff changeset
  1548
25917
d6c920623afc further localization
haftmann
parents: 25762
diff changeset
  1549
end
25230
022029099a83 continued localization
haftmann
parents: 25193
diff changeset
  1550
26274
2bdb61a28971 continued localization
haftmann
parents: 26234
diff changeset
  1551
text {* Simprules for comparisons where common factors can be cancelled. *}
15234
ec91a90c604e simplification tweaks for better arithmetic reasoning
paulson
parents: 15229
diff changeset
  1552
54147
97a8ff4e4ac9 killed most "no_atp", to make Sledgehammer more complete
blanchet
parents: 52435
diff changeset
  1553
lemmas mult_compare_simps =
15234
ec91a90c604e simplification tweaks for better arithmetic reasoning
paulson
parents: 15229
diff changeset
  1554
    mult_le_cancel_right mult_le_cancel_left
ec91a90c604e simplification tweaks for better arithmetic reasoning
paulson
parents: 15229
diff changeset
  1555
    mult_le_cancel_right1 mult_le_cancel_right2
ec91a90c604e simplification tweaks for better arithmetic reasoning
paulson
parents: 15229
diff changeset
  1556
    mult_le_cancel_left1 mult_le_cancel_left2
ec91a90c604e simplification tweaks for better arithmetic reasoning
paulson
parents: 15229
diff changeset
  1557
    mult_less_cancel_right mult_less_cancel_left
ec91a90c604e simplification tweaks for better arithmetic reasoning
paulson
parents: 15229
diff changeset
  1558
    mult_less_cancel_right1 mult_less_cancel_right2
ec91a90c604e simplification tweaks for better arithmetic reasoning
paulson
parents: 15229
diff changeset
  1559
    mult_less_cancel_left1 mult_less_cancel_left2
ec91a90c604e simplification tweaks for better arithmetic reasoning
paulson
parents: 15229
diff changeset
  1560
    mult_cancel_right mult_cancel_left
ec91a90c604e simplification tweaks for better arithmetic reasoning
paulson
parents: 15229
diff changeset
  1561
    mult_cancel_right1 mult_cancel_right2
ec91a90c604e simplification tweaks for better arithmetic reasoning
paulson
parents: 15229
diff changeset
  1562
    mult_cancel_left1 mult_cancel_left2
ec91a90c604e simplification tweaks for better arithmetic reasoning
paulson
parents: 15229
diff changeset
  1563
36301
72f4d079ebf8 more localization; factored out lemmas for division_ring
haftmann
parents: 35828
diff changeset
  1564
text {* Reasoning about inequalities with division *}
16775
c1b87ef4a1c3 added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
avigad
parents: 16568
diff changeset
  1565
35028
108662d50512 more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents: 34146
diff changeset
  1566
context linordered_semidom
25193
e2e1a4b00de3 various localizations
haftmann
parents: 25186
diff changeset
  1567
begin
e2e1a4b00de3 various localizations
haftmann
parents: 25186
diff changeset
  1568
e2e1a4b00de3 various localizations
haftmann
parents: 25186
diff changeset
  1569
lemma less_add_one: "a < a + 1"
14293
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1570
proof -
25193
e2e1a4b00de3 various localizations
haftmann
parents: 25186
diff changeset
  1571
  have "a + 0 < a + 1"
23482
2f4be6844f7c tuned and used field_simps
nipkow
parents: 23477
diff changeset
  1572
    by (blast intro: zero_less_one add_strict_left_mono)
14293
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1573
  thus ?thesis by simp
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1574
qed
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1575
25193
e2e1a4b00de3 various localizations
haftmann
parents: 25186
diff changeset
  1576
lemma zero_less_two: "0 < 1 + 1"
29667
53103fc8ffa3 Replaced group_ and ring_simps by algebra_simps;
nipkow
parents: 29465
diff changeset
  1577
by (blast intro: less_trans zero_less_one less_add_one)
25193
e2e1a4b00de3 various localizations
haftmann
parents: 25186
diff changeset
  1578
e2e1a4b00de3 various localizations
haftmann
parents: 25186
diff changeset
  1579
end
14365
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 14353
diff changeset
  1580
36301
72f4d079ebf8 more localization; factored out lemmas for division_ring
haftmann
parents: 35828
diff changeset
  1581
context linordered_idom
72f4d079ebf8 more localization; factored out lemmas for division_ring
haftmann
parents: 35828
diff changeset
  1582
begin
15234
ec91a90c604e simplification tweaks for better arithmetic reasoning
paulson
parents: 15229
diff changeset
  1583
36301
72f4d079ebf8 more localization; factored out lemmas for division_ring
haftmann
parents: 35828
diff changeset
  1584
lemma mult_right_le_one_le:
72f4d079ebf8 more localization; factored out lemmas for division_ring
haftmann
parents: 35828
diff changeset
  1585
  "0 \<le> x \<Longrightarrow> 0 \<le> y \<Longrightarrow> y \<le> 1 \<Longrightarrow> x * y \<le> x"
59833
ab828c2c5d67 clarified no_zero_devisors: makes only sense in a semiring;
haftmann
parents: 59832
diff changeset
  1586
  by (rule mult_left_le)
36301
72f4d079ebf8 more localization; factored out lemmas for division_ring
haftmann
parents: 35828
diff changeset
  1587
72f4d079ebf8 more localization; factored out lemmas for division_ring
haftmann
parents: 35828
diff changeset
  1588
lemma mult_left_le_one_le:
72f4d079ebf8 more localization; factored out lemmas for division_ring
haftmann
parents: 35828
diff changeset
  1589
  "0 \<le> x \<Longrightarrow> 0 \<le> y \<Longrightarrow> y \<le> 1 \<Longrightarrow> y * x \<le> x"
72f4d079ebf8 more localization; factored out lemmas for division_ring
haftmann
parents: 35828
diff changeset
  1590
  by (auto simp add: mult_le_cancel_right2)
72f4d079ebf8 more localization; factored out lemmas for division_ring
haftmann
parents: 35828
diff changeset
  1591
72f4d079ebf8 more localization; factored out lemmas for division_ring
haftmann
parents: 35828
diff changeset
  1592
end
72f4d079ebf8 more localization; factored out lemmas for division_ring
haftmann
parents: 35828
diff changeset
  1593
72f4d079ebf8 more localization; factored out lemmas for division_ring
haftmann
parents: 35828
diff changeset
  1594
text {* Absolute Value *}
14293
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1595
35028
108662d50512 more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents: 34146
diff changeset
  1596
context linordered_idom
25304
7491c00f0915 removed subclass edge ordered_ring < lordered_ring
haftmann
parents: 25267
diff changeset
  1597
begin
7491c00f0915 removed subclass edge ordered_ring < lordered_ring
haftmann
parents: 25267
diff changeset
  1598
36301
72f4d079ebf8 more localization; factored out lemmas for division_ring
haftmann
parents: 35828
diff changeset
  1599
lemma mult_sgn_abs:
72f4d079ebf8 more localization; factored out lemmas for division_ring
haftmann
parents: 35828
diff changeset
  1600
  "sgn x * \<bar>x\<bar> = x"
25304
7491c00f0915 removed subclass edge ordered_ring < lordered_ring
haftmann
parents: 25267
diff changeset
  1601
  unfolding abs_if sgn_if by auto
7491c00f0915 removed subclass edge ordered_ring < lordered_ring
haftmann
parents: 25267
diff changeset
  1602
36301
72f4d079ebf8 more localization; factored out lemmas for division_ring
haftmann
parents: 35828
diff changeset
  1603
lemma abs_one [simp]:
72f4d079ebf8 more localization; factored out lemmas for division_ring
haftmann
parents: 35828
diff changeset
  1604
  "\<bar>1\<bar> = 1"
44921
58eef4843641 tuned proofs
huffman
parents: 44350
diff changeset
  1605
  by (simp add: abs_if)
36301
72f4d079ebf8 more localization; factored out lemmas for division_ring
haftmann
parents: 35828
diff changeset
  1606
25304
7491c00f0915 removed subclass edge ordered_ring < lordered_ring
haftmann
parents: 25267
diff changeset
  1607
end
24491
8d194c9198ae added constant sgn
nipkow
parents: 24427
diff changeset
  1608
35028
108662d50512 more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents: 34146
diff changeset
  1609
class ordered_ring_abs = ordered_ring + ordered_ab_group_add_abs +
25304
7491c00f0915 removed subclass edge ordered_ring < lordered_ring
haftmann
parents: 25267
diff changeset
  1610
  assumes abs_eq_mult:
7491c00f0915 removed subclass edge ordered_ring < lordered_ring
haftmann
parents: 25267
diff changeset
  1611
    "(0 \<le> a \<or> a \<le> 0) \<and> (0 \<le> b \<or> b \<le> 0) \<Longrightarrow> \<bar>a * b\<bar> = \<bar>a\<bar> * \<bar>b\<bar>"
7491c00f0915 removed subclass edge ordered_ring < lordered_ring
haftmann
parents: 25267
diff changeset
  1612
35028
108662d50512 more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents: 34146
diff changeset
  1613
context linordered_idom
30961
541bfff659af more localisation
haftmann
parents: 30692
diff changeset
  1614
begin
541bfff659af more localisation
haftmann
parents: 30692
diff changeset
  1615
35028
108662d50512 more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents: 34146
diff changeset
  1616
subclass ordered_ring_abs proof
35216
7641e8d831d2 get rid of many duplicate simp rule warnings
huffman
parents: 35097
diff changeset
  1617
qed (auto simp add: abs_if not_less mult_less_0_iff)
30961
541bfff659af more localisation
haftmann
parents: 30692
diff changeset
  1618
541bfff659af more localisation
haftmann
parents: 30692
diff changeset
  1619
lemma abs_mult:
60562
24af00b010cf Amalgamation of the class comm_semiring_1_diff_distrib into comm_semiring_1_cancel. Moving axiom le_add_diff_inverse2 from semiring_numeral_div to linordered_semidom.
paulson <lp15@cam.ac.uk>
parents: 60529
diff changeset
  1620
  "\<bar>a * b\<bar> = \<bar>a\<bar> * \<bar>b\<bar>"
30961
541bfff659af more localisation
haftmann
parents: 30692
diff changeset
  1621
  by (rule abs_eq_mult) auto
541bfff659af more localisation
haftmann
parents: 30692
diff changeset
  1622
541bfff659af more localisation
haftmann
parents: 30692
diff changeset
  1623
lemma abs_mult_self:
36301
72f4d079ebf8 more localization; factored out lemmas for division_ring
haftmann
parents: 35828
diff changeset
  1624
  "\<bar>a\<bar> * \<bar>a\<bar> = a * a"
60562
24af00b010cf Amalgamation of the class comm_semiring_1_diff_distrib into comm_semiring_1_cancel. Moving axiom le_add_diff_inverse2 from semiring_numeral_div to linordered_semidom.
paulson <lp15@cam.ac.uk>
parents: 60529
diff changeset
  1625
  by (simp add: abs_if)
30961
541bfff659af more localisation
haftmann
parents: 30692
diff changeset
  1626
14294
f4d806fd72ce absolute value theorems moved to HOL/Ring_and_Field
paulson
parents: 14293
diff changeset
  1627
lemma abs_mult_less:
36301
72f4d079ebf8 more localization; factored out lemmas for division_ring
haftmann
parents: 35828
diff changeset
  1628
  "\<bar>a\<bar> < c \<Longrightarrow> \<bar>b\<bar> < d \<Longrightarrow> \<bar>a\<bar> * \<bar>b\<bar> < c * d"
14294
f4d806fd72ce absolute value theorems moved to HOL/Ring_and_Field
paulson
parents: 14293
diff changeset
  1629
proof -
36301
72f4d079ebf8 more localization; factored out lemmas for division_ring
haftmann
parents: 35828
diff changeset
  1630
  assume ac: "\<bar>a\<bar> < c"
72f4d079ebf8 more localization; factored out lemmas for division_ring
haftmann
parents: 35828
diff changeset
  1631
  hence cpos: "0<c" by (blast intro: le_less_trans abs_ge_zero)
72f4d079ebf8 more localization; factored out lemmas for division_ring
haftmann
parents: 35828
diff changeset
  1632
  assume "\<bar>b\<bar> < d"
60562
24af00b010cf Amalgamation of the class comm_semiring_1_diff_distrib into comm_semiring_1_cancel. Moving axiom le_add_diff_inverse2 from semiring_numeral_div to linordered_semidom.
paulson <lp15@cam.ac.uk>
parents: 60529
diff changeset
  1633
  thus ?thesis by (simp add: ac cpos mult_strict_mono)
14294
f4d806fd72ce absolute value theorems moved to HOL/Ring_and_Field
paulson
parents: 14293
diff changeset
  1634
qed
14293
22542982bffd moving some division theorems to Ring_and_Field
paulson
parents: 14288
diff changeset
  1635
36301
72f4d079ebf8 more localization; factored out lemmas for division_ring
haftmann
parents: 35828
diff changeset
  1636
lemma abs_less_iff:
60562
24af00b010cf Amalgamation of the class comm_semiring_1_diff_distrib into comm_semiring_1_cancel. Moving axiom le_add_diff_inverse2 from semiring_numeral_div to linordered_semidom.
paulson <lp15@cam.ac.uk>
parents: 60529
diff changeset
  1637
  "\<bar>a\<bar> < b \<longleftrightarrow> a < b \<and> - a < b"
36301
72f4d079ebf8 more localization; factored out lemmas for division_ring
haftmann
parents: 35828
diff changeset
  1638
  by (simp add: less_le abs_le_iff) (auto simp add: abs_if)
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 14603
diff changeset
  1639
36301
72f4d079ebf8 more localization; factored out lemmas for division_ring
haftmann
parents: 35828
diff changeset
  1640
lemma abs_mult_pos:
72f4d079ebf8 more localization; factored out lemmas for division_ring
haftmann
parents: 35828
diff changeset
  1641
  "0 \<le> x \<Longrightarrow> \<bar>y\<bar> * x = \<bar>y * x\<bar>"
72f4d079ebf8 more localization; factored out lemmas for division_ring
haftmann
parents: 35828
diff changeset
  1642
  by (simp add: abs_mult)
72f4d079ebf8 more localization; factored out lemmas for division_ring
haftmann
parents: 35828
diff changeset
  1643
51520
e9b361845809 move real_isLub_unique to isLub_unique in Lubs; real_sum_of_halves to RealDef; abs_diff_less_iff to Rings
hoelzl
parents: 50420
diff changeset
  1644
lemma abs_diff_less_iff:
e9b361845809 move real_isLub_unique to isLub_unique in Lubs; real_sum_of_halves to RealDef; abs_diff_less_iff to Rings
hoelzl
parents: 50420
diff changeset
  1645
  "\<bar>x - a\<bar> < r \<longleftrightarrow> a - r < x \<and> x < a + r"
e9b361845809 move real_isLub_unique to isLub_unique in Lubs; real_sum_of_halves to RealDef; abs_diff_less_iff to Rings
hoelzl
parents: 50420
diff changeset
  1646
  by (auto simp add: diff_less_eq ac_simps abs_less_iff)
e9b361845809 move real_isLub_unique to isLub_unique in Lubs; real_sum_of_halves to RealDef; abs_diff_less_iff to Rings
hoelzl
parents: 50420
diff changeset
  1647
59865
8a20dd967385 rationalised and generalised some theorems concerning abs and x^2.
paulson <lp15@cam.ac.uk>
parents: 59833
diff changeset
  1648
lemma abs_diff_le_iff:
8a20dd967385 rationalised and generalised some theorems concerning abs and x^2.
paulson <lp15@cam.ac.uk>
parents: 59833
diff changeset
  1649
   "\<bar>x - a\<bar> \<le> r \<longleftrightarrow> a - r \<le> x \<and> x \<le> a + r"
8a20dd967385 rationalised and generalised some theorems concerning abs and x^2.
paulson <lp15@cam.ac.uk>
parents: 59833
diff changeset
  1650
  by (auto simp add: diff_le_eq ac_simps abs_le_iff)
8a20dd967385 rationalised and generalised some theorems concerning abs and x^2.
paulson <lp15@cam.ac.uk>
parents: 59833
diff changeset
  1651
36301
72f4d079ebf8 more localization; factored out lemmas for division_ring
haftmann
parents: 35828
diff changeset
  1652
end
16775
c1b87ef4a1c3 added lemmas to OrderedGroup.thy (reasoning about signs, absolute value, triangle inequalities)
avigad
parents: 16568
diff changeset
  1653
59557
ebd8ecacfba6 establish unique preferred fact names
haftmann
parents: 59555
diff changeset
  1654
hide_fact (open) comm_mult_left_mono comm_mult_strict_left_mono distrib
ebd8ecacfba6 establish unique preferred fact names
haftmann
parents: 59555
diff changeset
  1655
52435
6646bb548c6b migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
haftmann
parents: 51520
diff changeset
  1656
code_identifier
6646bb548c6b migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
haftmann
parents: 51520
diff changeset
  1657
  code_module Rings \<rightharpoonup> (SML) Arith and (OCaml) Arith and (Haskell) Arith
33364
2bd12592c5e8 tuned code setup
haftmann
parents: 33319
diff changeset
  1658
14265
95b42e69436c HOL: installation of Ring_and_Field as the basis for Naturals and Reals
paulson
parents:
diff changeset
  1659
end